diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index 3441ca351115d..22bda179c97ca 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -566,7 +566,7 @@ lemma mul_mem_sup {S T : subgroup G} {x y : G} (hx : x ∈ S) (hy : y ∈ T) : x (S ⊔ T).mul_mem (mem_sup_left hx) (mem_sup_right hy) @[to_additive] -lemma mem_supr_of_mem {ι : Type*} {S : ι → subgroup G} (i : ι) : +lemma mem_supr_of_mem {ι : Sort*} {S : ι → subgroup G} (i : ι) : ∀ {x : G}, x ∈ S i → x ∈ supr S := show S i ≤ supr S, from le_supr _ _ @@ -706,6 +706,11 @@ lemma closure_eq_bot_iff (G : Type*) [group G] (S : set G) : closure S = ⊥ ↔ S ⊆ {1} := by { rw [← le_bot_iff], exact closure_le _} +@[to_additive] +lemma supr_eq_closure {ι : Sort*} (p : ι → subgroup G) : + (⨆ i, p i) = closure (⋃ i, (p i : set G)) := +by simp_rw [closure_Union, closure_eq] + /-- The subgroup generated by an element of a group equals the set of integer number powers of the element. -/ @[to_additive /-"The `add_subgroup` generated by an element of an `add_group` equals the set of @@ -774,6 +779,41 @@ begin exact Hk_inv _ invmem } }, end +/-- An induction principle for elements of `⨆ i, S i`. +If `C` holds for `1` and all elements of `S i` for all `i`, and is preserved under multiplication, +then it holds for all elements of the supremum of `S`. -/ +@[elab_as_eliminator, to_additive /-" An induction principle for elements of `⨆ i, S i`. +If `C` holds for `0` and all elements of `S i` for all `i`, and is preserved under addition, +then it holds for all elements of the supremum of `S`. "-/] +lemma supr_induction {ι : Sort*} (S : ι → subgroup G) {C : G → Prop} {x : G} (hx : x ∈ ⨆ i, S i) + (hp : ∀ i (x ∈ S i), C x) + (h1 : C 1) + (hmul : ∀ x y, C x → C y → C (x * y)) : C x := +begin + rw supr_eq_closure at hx, + refine closure_induction'' hx (λ x hx, _) (λ x hx, _) h1 hmul, + { obtain ⟨i, hi⟩ := set.mem_Union.mp hx, + exact hp _ _ hi, }, + { obtain ⟨i, hi⟩ := set.mem_Union.mp hx, + exact hp _ _ (inv_mem _ hi), }, +end + +/-- A dependent version of `subgroup.supr_induction`. -/ +@[elab_as_eliminator, to_additive /-"A dependent version of `add_subgroup.supr_induction`. "-/] +lemma supr_induction' {ι : Sort*} (S : ι → subgroup G) {C : Π x, (x ∈ ⨆ i, S i) → Prop} + (hp : ∀ i (x ∈ S i), C x (mem_supr_of_mem i ‹_›)) + (h1 : C 1 (one_mem _)) + (hmul : ∀ x y hx hy, C x hx → C y hy → C (x * y) (mul_mem _ ‹_› ‹_›)) + {x : G} (hx : x ∈ ⨆ i, S i) : C x hx := +begin + refine exists.elim _ (λ (hx : x ∈ ⨆ i, S i) (hc : C x hx), hc), + refine supr_induction S hx (λ i x hx, _) _ (λ x y, _), + { exact ⟨_, hp _ _ hx⟩ }, + { exact ⟨_, h1⟩ }, + { rintro ⟨_, Cx⟩ ⟨_, Cy⟩, + refine ⟨_, hmul _ _ _ _ Cx Cy⟩ }, +end + @[to_additive] lemma mem_supr_of_directed {ι} [hι : nonempty ι] {K : ι → subgroup G} (hK : directed (≤) K) {x : G} : diff --git a/src/group_theory/submonoid/membership.lean b/src/group_theory/submonoid/membership.lean index 401aa8a29122c..ce8d5892f5e75 100644 --- a/src/group_theory/submonoid/membership.lean +++ b/src/group_theory/submonoid/membership.lean @@ -135,7 +135,7 @@ lemma mul_mem_sup {S T : submonoid M} {x y : M} (hx : x ∈ S) (hy : y ∈ T) : (S ⊔ T).mul_mem (mem_sup_left hx) (mem_sup_right hy) @[to_additive] -lemma mem_supr_of_mem {ι : Type*} {S : ι → submonoid M} (i : ι) : +lemma mem_supr_of_mem {ι : Sort*} {S : ι → submonoid M} (i : ι) : ∀ {x : M}, x ∈ S i → x ∈ supr S := show S i ≤ supr S, from le_supr _ _ @@ -144,6 +144,39 @@ lemma mem_Sup_of_mem {S : set (submonoid M)} {s : submonoid M} (hs : s ∈ S) : ∀ {x : M}, x ∈ s → x ∈ Sup S := show s ≤ Sup S, from le_Sup hs +/-- An induction principle for elements of `⨆ i, S i`. +If `C` holds for `1` and all elements of `S i` for all `i`, and is preserved under multiplication, +then it holds for all elements of the supremum of `S`. -/ +@[elab_as_eliminator, to_additive /-" An induction principle for elements of `⨆ i, S i`. +If `C` holds for `0` and all elements of `S i` for all `i`, and is preserved under addition, +then it holds for all elements of the supremum of `S`. "-/] +lemma supr_induction {ι : Sort*} (S : ι → submonoid M) {C : M → Prop} {x : M} (hx : x ∈ ⨆ i, S i) + (hp : ∀ i (x ∈ S i), C x) + (h1 : C 1) + (hmul : ∀ x y, C x → C y → C (x * y)) : C x := +begin + rw supr_eq_closure at hx, + refine closure_induction hx (λ x hx, _) h1 hmul, + obtain ⟨i, hi⟩ := set.mem_Union.mp hx, + exact hp _ _ hi, +end + +/-- A dependent version of `submonoid.supr_induction`. -/ +@[elab_as_eliminator, to_additive /-"A dependent version of `add_submonoid.supr_induction`. "-/] +lemma supr_induction' {ι : Sort*} (S : ι → submonoid M) {C : Π x, (x ∈ ⨆ i, S i) → Prop} + (hp : ∀ i (x ∈ S i), C x (mem_supr_of_mem i ‹_›)) + (h1 : C 1 (one_mem _)) + (hmul : ∀ x y hx hy, C x hx → C y hy → C (x * y) (mul_mem _ ‹_› ‹_›)) + {x : M} (hx : x ∈ ⨆ i, S i) : C x hx := +begin + refine exists.elim _ (λ (hx : x ∈ ⨆ i, S i) (hc : C x hx), hc), + refine supr_induction S hx (λ i x hx, _) _ (λ x y, _), + { exact ⟨_, hp _ _ hx⟩ }, + { exact ⟨_, h1⟩ }, + { rintro ⟨_, Cx⟩ ⟨_, Cy⟩, + refine ⟨_, hmul _ _ _ _ Cx Cy⟩ }, +end + end non_assoc end submonoid diff --git a/src/linear_algebra/basic.lean b/src/linear_algebra/basic.lean index f9077158864ab..ac9f4bb850daf 100644 --- a/src/linear_algebra/basic.lean +++ b/src/linear_algebra/basic.lean @@ -1225,6 +1225,35 @@ begin exact add_submonoid.add_mem _ hx hy, } } end +/-- An induction principle for elements of `⨆ i, p i`. +If `C` holds for `0` and all elements of `p i` for all `i`, and is preserved under addition, +then it holds for all elements of the supremum of `p`. -/ +@[elab_as_eliminator] +lemma supr_induction {ι : Sort*} (p : ι → submodule R M) {C : M → Prop} {x : M} (hx : x ∈ ⨆ i, p i) + (hp : ∀ i (x ∈ p i), C x) + (h0 : C 0) + (hadd : ∀ x y, C x → C y → C (x + y)) : C x := +begin + rw [←mem_to_add_submonoid, supr_to_add_submonoid] at hx, + exact add_submonoid.supr_induction _ hx hp h0 hadd, +end + +/-- A dependent version of `submodule.supr_induction`. -/ +@[elab_as_eliminator] +lemma supr_induction' {ι : Sort*} (p : ι → submodule R M) {C : Π x, (x ∈ ⨆ i, p i) → Prop} + (hp : ∀ i (x ∈ p i), C x (mem_supr_of_mem i ‹_›)) + (h0 : C 0 (zero_mem _)) + (hadd : ∀ x y hx hy, C x hx → C y hy → C (x + y) (add_mem _ ‹_› ‹_›)) + {x : M} (hx : x ∈ ⨆ i, p i) : C x hx := +begin + refine exists.elim _ (λ (hx : x ∈ ⨆ i, p i) (hc : C x hx), hc), + refine supr_induction p hx (λ i x hx, _) _ (λ x y, _), + { exact ⟨_, hp _ _ hx⟩ }, + { exact ⟨_, h0⟩ }, + { rintro ⟨_, Cx⟩ ⟨_, Cy⟩, + refine ⟨_, hadd _ _ _ _ Cx Cy⟩ }, +end + lemma span_singleton_le_iff_mem (m : M) (p : submodule R M) : (R ∙ m) ≤ p ↔ m ∈ p := by rw [span_le, singleton_subset_iff, set_like.mem_coe]