Skip to content

Commit 69ed7ba

Browse files
committed
chore: register TensorProduct with induction_eliminator (#14204)
Split from #12605
1 parent d465284 commit 69ed7ba

File tree

11 files changed

+26
-26
lines changed

11 files changed

+26
-26
lines changed

Mathlib/LinearAlgebra/DirectSum/Finsupp.lean

+3-3
Original file line numberDiff line numberDiff line change
@@ -101,7 +101,7 @@ lemma finsuppLeft_apply_tmul_apply (p : ι →₀ M) (n : N) (i : ι) :
101101

102102
theorem finsuppLeft_apply (t : (ι →₀ M) ⊗[R] N) (i : ι) :
103103
finsuppLeft R M N ι t i = rTensor N (Finsupp.lapply i) t := by
104-
induction t using TensorProduct.induction_on with
104+
induction t with
105105
| zero => simp
106106
| tmul f n => simp only [finsuppLeft_apply_tmul_apply, rTensor_tmul, Finsupp.lapply_apply]
107107
| add x y hx hy => simp [map_add, hx, hy]
@@ -136,7 +136,7 @@ lemma finsuppRight_apply_tmul_apply (m : M) (p : ι →₀ N) (i : ι) :
136136

137137
theorem finsuppRight_apply (t : M ⊗[R] (ι →₀ N)) (i : ι) :
138138
finsuppRight R M N ι t i = lTensor M (Finsupp.lapply i) t := by
139-
induction t using TensorProduct.induction_on with
139+
induction t with
140140
| zero => simp
141141
| tmul m f => simp [finsuppRight_apply_tmul_apply]
142142
| add x y hx hy => simp [map_add, hx, hy]
@@ -152,7 +152,7 @@ variable {S : Type*} [CommSemiring S] [Algebra R S]
152152

153153
lemma finsuppLeft_smul' (s : S) (t : (ι →₀ M) ⊗[R] N) :
154154
finsuppLeft R M N ι (s • t) = s • finsuppLeft R M N ι t := by
155-
induction t using TensorProduct.induction_on with
155+
induction t with
156156
| zero => simp
157157
| add x y hx hy => simp [hx, hy]
158158
| tmul p n => ext; simp [smul_tmul', finsuppLeft_apply_tmul_apply]

Mathlib/LinearAlgebra/TensorProduct/Basic.lean

+2-2
Original file line numberDiff line numberDiff line change
@@ -135,7 +135,7 @@ infixl:100 " ⊗ₜ " => tmul _
135135
notation:100 x " ⊗ₜ[" R "] " y:100 => tmul R x y
136136

137137
-- Porting note: make the arguments of induction_on explicit
138-
@[elab_as_elim]
138+
@[elab_as_elim, induction_eliminator]
139139
protected theorem induction_on {motive : M ⊗[R] N → Prop} (z : M ⊗[R] N)
140140
(zero : motive 0)
141141
(tmul : ∀ x y, motive <| x ⊗ₜ[R] y)
@@ -503,7 +503,7 @@ theorem map₂_mk_top_top_eq_top : Submodule.map₂ (mk R M N) ⊤ ⊤ = ⊤ :=
503503
theorem exists_eq_tmul_of_forall (x : TensorProduct R M N)
504504
(h : ∀ (m₁ m₂ : M) (n₁ n₂ : N), ∃ m n, m₁ ⊗ₜ n₁ + m₂ ⊗ₜ n₂ = m ⊗ₜ[R] n) :
505505
∃ m n, x = m ⊗ₜ n := by
506-
induction x using TensorProduct.induction_on with
506+
induction x with
507507
| zero =>
508508
use 0, 0
509509
rw [TensorProduct.zero_tmul]

Mathlib/LinearAlgebra/TensorProduct/Finiteness.lean

+2-2
Original file line numberDiff line numberDiff line change
@@ -51,7 +51,7 @@ namespace TensorProduct
5151
of `M × N`, such that `x` is equal to the sum of `m_i ⊗ₜ[R] n_i`. -/
5252
theorem exists_multiset (x : M ⊗[R] N) :
5353
∃ S : Multiset (M × N), x = (S.map fun i ↦ i.1 ⊗ₜ[R] i.2).sum := by
54-
induction x using TensorProduct.induction_on with
54+
induction x with
5555
| zero => exact ⟨0, by simp⟩
5656
| tmul x y => exact ⟨{(x, y)}, by simp⟩
5757
| add x y hx hy =>
@@ -64,7 +64,7 @@ of `M × N` such that each `m_i` is distinct (we represent it as an element of `
6464
such that `x` is equal to the sum of `m_i ⊗ₜ[R] n_i`. -/
6565
theorem exists_finsupp_left (x : M ⊗[R] N) :
6666
∃ S : M →₀ N, x = S.sum fun m n ↦ m ⊗ₜ[R] n := by
67-
induction x using TensorProduct.induction_on with
67+
induction x with
6868
| zero => exact ⟨0, by simp⟩
6969
| tmul x y => exact ⟨Finsupp.single x y, by simp⟩
7070
| add x y hx hy =>

Mathlib/LinearAlgebra/TensorProduct/RightExactness.lean

+8-8
Original file line numberDiff line numberDiff line change
@@ -111,7 +111,7 @@ variable (Q) {g}
111111
theorem LinearMap.lTensor_surjective (hg : Function.Surjective g) :
112112
Function.Surjective (lTensor Q g) := by
113113
intro z
114-
induction z using TensorProduct.induction_on with
114+
induction z with
115115
| zero => exact ⟨0, map_zero _⟩
116116
| tmul q p =>
117117
obtain ⟨n, rfl⟩ := hg p
@@ -136,7 +136,7 @@ theorem LinearMap.lTensor_range :
136136
theorem LinearMap.rTensor_surjective (hg : Function.Surjective g) :
137137
Function.Surjective (rTensor Q g) := by
138138
intro z
139-
induction z using TensorProduct.induction_on with
139+
induction z with
140140
| zero => exact ⟨0, map_zero _⟩
141141
| tmul p q =>
142142
obtain ⟨n, rfl⟩ := hg p
@@ -532,12 +532,12 @@ lemma Ideal.map_includeLeft_eq (I : Ideal A) :
532532
use x + y
533533
simp only [map_add]
534534
· rintro a x ⟨x, hx, rfl⟩
535-
induction a using TensorProduct.induction_on with
535+
induction a with
536536
| zero =>
537537
use 0
538538
simp only [map_zero, smul_eq_mul, zero_mul]
539539
| tmul a b =>
540-
induction x using TensorProduct.induction_on with
540+
induction x with
541541
| zero =>
542542
use 0
543543
simp only [map_zero, smul_eq_mul, mul_zero]
@@ -557,7 +557,7 @@ lemma Ideal.map_includeLeft_eq (I : Ideal A) :
557557
simp only [map_add, ha', add_smul, hb']
558558

559559
· rintro x ⟨y, rfl⟩
560-
induction y using TensorProduct.induction_on with
560+
induction y with
561561
| zero =>
562562
rw [map_zero]
563563
apply zero_mem
@@ -600,12 +600,12 @@ lemma Ideal.map_includeRight_eq (I : Ideal B) :
600600
use x + y
601601
simp only [map_add]
602602
· rintro a x ⟨x, hx, rfl⟩
603-
induction a using TensorProduct.induction_on with
603+
induction a with
604604
| zero =>
605605
use 0
606606
simp only [map_zero, smul_eq_mul, zero_mul]
607607
| tmul a b =>
608-
induction x using TensorProduct.induction_on with
608+
induction x with
609609
| zero =>
610610
use 0
611611
simp only [map_zero, smul_eq_mul, mul_zero]
@@ -625,7 +625,7 @@ lemma Ideal.map_includeRight_eq (I : Ideal B) :
625625
simp only [map_add, ha', add_smul, hb']
626626

627627
· rintro x ⟨y, rfl⟩
628-
induction y using TensorProduct.induction_on with
628+
induction y with
629629
| zero =>
630630
rw [map_zero]
631631
apply zero_mem

Mathlib/LinearAlgebra/TensorProduct/Submodule.lean

+1-1
Original file line numberDiff line numberDiff line change
@@ -98,7 +98,7 @@ theorem mulMap_eq_mul'_comp_mapIncl : mulMap M N = .mul' R S ∘ₗ TensorProduc
9898
theorem mulMap_range : LinearMap.range (mulMap M N) = M * N := by
9999
refine le_antisymm ?_ (mul_le.2 fun m hm n hn ↦ ⟨⟨m, hm⟩ ⊗ₜ[R] ⟨n, hn⟩, rfl⟩)
100100
rintro _ ⟨x, rfl⟩
101-
induction x using TensorProduct.induction_on with
101+
induction x with
102102
| zero => rw [_root_.map_zero]; exact zero_mem _
103103
| tmul a b => exact mul_mem_mul a.2 b.2
104104
| add a b ha hb => rw [_root_.map_add]; exact add_mem ha hb

Mathlib/RingTheory/Finiteness.lean

+2-2
Original file line numberDiff line numberDiff line change
@@ -422,7 +422,7 @@ theorem exists_fg_le_eq_rTensor_inclusion {R M N : Type*} [CommRing R] [AddCommG
422422
[AddCommGroup N] [Module R M] [Module R N] {I : Submodule R N} (x : I ⊗ M) :
423423
∃ (J : Submodule R N) (_ : J.FG) (hle : J ≤ I) (y : J ⊗ M),
424424
x = rTensor M (J.inclusion hle) y := by
425-
induction x using TensorProduct.induction_on with
425+
induction x with
426426
| zero => exact ⟨⊥, fg_bot, zero_le _, 0, rfl⟩
427427
| tmul i m => exact ⟨R ∙ i.val, fg_span_singleton i.val,
428428
(span_singleton_le_iff_mem _ _).mpr i.property,
@@ -733,7 +733,7 @@ instance Module.Finite.base_change [CommSemiring R] [Semiring A] [Algebra R A] [
733733
obtain ⟨s, hs⟩ := h.out
734734
refine ⟨⟨s.image (TensorProduct.mk R A M 1), eq_top_iff.mpr ?_⟩⟩
735735
rintro x -
736-
induction x using TensorProduct.induction_on with
736+
induction x with
737737
| zero => exact zero_mem _
738738
| tmul x y =>
739739
-- Porting note: new TC reminder

Mathlib/RingTheory/IsTensorProduct.lean

+1-1
Original file line numberDiff line numberDiff line change
@@ -117,7 +117,7 @@ theorem IsTensorProduct.inductionOn (h : IsTensorProduct f) {C : M → Prop} (m
117117
rw [← h.equiv.right_inv m]
118118
generalize h.equiv.invFun m = y
119119
change C (TensorProduct.lift f y)
120-
induction y using TensorProduct.induction_on with
120+
induction y with
121121
| zero => rwa [map_zero]
122122
| tmul _ _ =>
123123
rw [TensorProduct.lift.tmul]

Mathlib/RingTheory/Kaehler/Basic.lean

+1-1
Original file line numberDiff line numberDiff line change
@@ -761,7 +761,7 @@ lemma KaehlerDifferential.range_mapBaseChange :
761761
LinearMap.range (mapBaseChange R A B) = LinearMap.ker (map R A B B) := by
762762
apply le_antisymm
763763
· rintro _ ⟨x, rfl⟩
764-
induction' x using TensorProduct.induction_on with r s
764+
induction' x with r s
765765
· simp
766766
· obtain ⟨x, rfl⟩ := total_surjective _ _ s
767767
simp only [mapBaseChange_tmul, LinearMap.mem_ker, map_smul]

Mathlib/RingTheory/MatrixAlgebra.lean

+1-1
Original file line numberDiff line numberDiff line change
@@ -122,7 +122,7 @@ theorem right_inv (M : Matrix n n A) : (toFunAlgHom R A n) (invFun R A n M) = M
122122
#align matrix_equiv_tensor.right_inv MatrixEquivTensor.right_inv
123123

124124
theorem left_inv (M : A ⊗[R] Matrix n n R) : invFun R A n (toFunAlgHom R A n M) = M := by
125-
induction M using TensorProduct.induction_on with
125+
induction M with
126126
| zero => simp
127127
| tmul a m => simp
128128
| add x y hx hy =>

Mathlib/RingTheory/RingHom/Surjective.lean

+1-1
Original file line numberDiff line numberDiff line change
@@ -37,7 +37,7 @@ theorem surjective_stableUnderBaseChange : StableUnderBaseChange surjective := b
3737
refine StableUnderBaseChange.mk _ surjective_respectsIso ?_
3838
classical
3939
introv h x
40-
induction x using TensorProduct.induction_on with
40+
induction x with
4141
| zero => exact ⟨0, map_zero _⟩
4242
| tmul x y =>
4343
obtain ⟨y, rfl⟩ := h y; use y • x; dsimp

Mathlib/RingTheory/TensorProduct/Basic.lean

+4-4
Original file line numberDiff line numberDiff line change
@@ -258,9 +258,9 @@ instance (priority := 100) isScalarTower_right [Monoid S] [DistribMulAction S A]
258258
[IsScalarTower S A A] [SMulCommClass R S A] : IsScalarTower S (A ⊗[R] B) (A ⊗[R] B) where
259259
smul_assoc r x y := by
260260
change r • x * y = r • (x * y)
261-
induction y using TensorProduct.induction_on with
261+
induction y with
262262
| zero => simp [smul_zero]
263-
| tmul a b => induction x using TensorProduct.induction_on with
263+
| tmul a b => induction x with
264264
| zero => simp [smul_zero]
265265
| tmul a' b' =>
266266
dsimp
@@ -274,9 +274,9 @@ instance (priority := 100) sMulCommClass_right [Monoid S] [DistribMulAction S A]
274274
[SMulCommClass S A A] [SMulCommClass R S A] : SMulCommClass S (A ⊗[R] B) (A ⊗[R] B) where
275275
smul_comm r x y := by
276276
change r • (x * y) = x * r • y
277-
induction y using TensorProduct.induction_on with
277+
induction y with
278278
| zero => simp [smul_zero]
279-
| tmul a b => induction x using TensorProduct.induction_on with
279+
| tmul a b => induction x with
280280
| zero => simp [smul_zero]
281281
| tmul a' b' =>
282282
dsimp

0 commit comments

Comments
 (0)