diff --git a/scripts/nolints.txt b/scripts/nolints.txt index 60d36e81c535d..0da58c64c895a 100644 --- a/scripts/nolints.txt +++ b/scripts/nolints.txt @@ -345,9 +345,6 @@ apply_nolint stream.unfolds doc_blame -- data/stream/init.lean apply_nolint stream.is_bisimulation doc_blame --- geometry/euclidean/circumcenter.lean -apply_nolint affine_independent.exists_unique_dist_eq fintype_finite - -- group_theory/group_action/sub_mul_action.lean apply_nolint sub_mul_action.has_zero fails_quickly diff --git a/scripts/port_status.py b/scripts/port_status.py index d4456453058a1..42766e5c9809b 100755 --- a/scripts/port_status.py +++ b/scripts/port_status.py @@ -71,6 +71,11 @@ def mk_label(path: Path) -> str: for node in graph.nodes: if data[node].mathlib3_hash: verified[node] = data[node].mathlib3_hash + find_blobs_command = ['git', 'cat-file', '-t', data[node].mathlib3_hash] + hash_type = subprocess.check_output(find_blobs_command) + # the hash_type should be commits mostly, we are not interested in blobs + if b'blob\n' == hash_type: + break git_command = ['git', 'diff', '--quiet', f'--ignore-matching-lines={comment_git_re}', data[node].mathlib3_hash + "..HEAD", "--", "src" + os.sep + node.replace('.', os.sep) + ".lean"] diff --git a/src/algebra/direct_sum/ring.lean b/src/algebra/direct_sum/ring.lean index f3cf3b22b3b2d..8da08a69f5213 100644 --- a/src/algebra/direct_sum/ring.lean +++ b/src/algebra/direct_sum/ring.lean @@ -122,16 +122,6 @@ class gcomm_ring [add_comm_monoid ι] [Π i, add_comm_group (A i)] extends end defs -/-- A graded version of `semiring.to_module`. -/ -instance gsemiring.to_gmodule (A : ι → Type*) - [add_monoid ι] [Π (i : ι), add_comm_monoid (A i)] [gsemiring A] : - graded_monoid.gmodule A A := -{ smul_add := λ _ _, gsemiring.mul_add, - smul_zero := λ i j, gsemiring.mul_zero, - add_smul := λ i j, gsemiring.add_mul, - zero_smul := λ i j, gsemiring.zero_mul, - ..graded_monoid.gmonoid.to_gmul_action A } - lemma of_eq_of_graded_monoid_eq {A : ι → Type*} [Π (i : ι), add_comm_monoid (A i)] {i j : ι} {a : A i} {b : A j} (h : graded_monoid.mk i a = graded_monoid.mk j b) : direct_sum.of A i a = direct_sum.of A j b := diff --git a/src/algebra/gcd_monoid/finset.lean b/src/algebra/gcd_monoid/finset.lean index ddd10ea5a762b..b0857a36fe7ed 100644 --- a/src/algebra/gcd_monoid/finset.lean +++ b/src/algebra/gcd_monoid/finset.lean @@ -9,6 +9,9 @@ import algebra.gcd_monoid.multiset /-! # GCD and LCM operations on finsets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions - `finset.gcd` - the greatest common denominator of a `finset` of elements of a `gcd_monoid` diff --git a/src/algebra/graded_mul_action.lean b/src/algebra/graded_mul_action.lean index 31d9a0c4be31a..c70754ef199a2 100644 --- a/src/algebra/graded_mul_action.lean +++ b/src/algebra/graded_mul_action.lean @@ -104,12 +104,6 @@ class set_like.has_graded_smul {S R N M : Type*} [set_like S R] [set_like N M] [has_smul R M] [has_add ι] (A : ι → S) (B : ι → N) : Prop := (smul_mem : ∀ ⦃i j : ι⦄ {ai bj}, ai ∈ A i → bj ∈ B j → ai • bj ∈ B (i + j)) -lemma set_like.smul_mem_graded {S R N M : Type*} [set_like S R] [set_like N M] - [has_smul R M] [has_add ι] {A : ι → S} {B : ι → N} [set_like.has_graded_smul A B] - ⦃i j⦄ {gi gj} (hi : gi ∈ A i) (hj : gj ∈ B j) : - gi • gj ∈ B (i + j) := -set_like.has_graded_smul.smul_mem hi hj - instance set_like.ghas_smul {S R N M : Type*} [set_like S R] [set_like N M] [has_smul R M] [has_add ι] (A : ι → S) (B : ι → N) [set_like.has_graded_smul A B] : graded_monoid.ghas_smul (λ i, A i) (λ i, B i) := @@ -136,6 +130,6 @@ variables {S R N M : Type*} [set_like S R] [set_like N M] lemma set_like.is_homogeneous.graded_smul [has_add ι] [has_smul R M] {A : ι → S} {B : ι → N} [set_like.has_graded_smul A B] {a : R} {b : M} : set_like.is_homogeneous A a → set_like.is_homogeneous B b → set_like.is_homogeneous B (a • b) -| ⟨i, hi⟩ ⟨j, hj⟩ := ⟨i + j, set_like.smul_mem_graded hi hj⟩ +| ⟨i, hi⟩ ⟨j, hj⟩ := ⟨i + j, set_like.has_graded_smul.smul_mem hi hj⟩ end homogeneous_elements diff --git a/src/algebra/group/basic.lean b/src/algebra/group/basic.lean index 6740b509dae66..9d06c6de2b904 100644 --- a/src/algebra/group/basic.lean +++ b/src/algebra/group/basic.lean @@ -651,3 +651,47 @@ lemma bit1_sub [has_one M] (a b : M) : bit1 (a - b) = bit1 a - bit0 b := (congr_arg (+ (1 : M)) $ bit0_sub a b : _).trans $ sub_add_eq_add_sub _ _ _ end subtraction_comm_monoid + +section multiplicative + +variables [monoid β] (p r : α → α → Prop) [is_total α r] (f : α → α → β) + +@[to_additive additive_of_symmetric_of_is_total] +lemma multiplicative_of_symmetric_of_is_total + (hsymm : symmetric p) (hf_swap : ∀ {a b}, p a b → f a b * f b a = 1) + (hmul : ∀ {a b c}, r a b → r b c → p a b → p b c → p a c → f a c = f a b * f b c) + {a b c : α} (pab : p a b) (pbc : p b c) (pac : p a c) : f a c = f a b * f b c := +begin + suffices : ∀ {b c}, r b c → p a b → p b c → p a c → f a c = f a b * f b c, + { obtain rbc | rcb := total_of r b c, + { exact this rbc pab pbc pac }, + { rw [this rcb pac (hsymm pbc) pab, mul_assoc, hf_swap (hsymm pbc), mul_one] } }, + intros b c rbc pab pbc pac, + obtain rab | rba := total_of r a b, + { exact hmul rab rbc pab pbc pac }, + rw [← one_mul (f a c), ← hf_swap pab, mul_assoc], + obtain rac | rca := total_of r a c, + { rw [hmul rba rac (hsymm pab) pac pbc] }, + { rw [hmul rbc rca pbc (hsymm pac) (hsymm pab), mul_assoc, hf_swap (hsymm pac), mul_one] }, +end + +/-- If a binary function from a type equipped with a total relation `r` to a monoid is + anti-symmetric (i.e. satisfies `f a b * f b a = 1`), in order to show it is multiplicative + (i.e. satisfies `f a c = f a b * f b c`), we may assume `r a b` and `r b c` are satisfied. + We allow restricting to a subset specified by a predicate `p`. -/ +@[to_additive additive_of_is_total "If a binary function from a type equipped with a total relation + `r` to an additive monoid is anti-symmetric (i.e. satisfies `f a b + f b a = 0`), in order to show + it is additive (i.e. satisfies `f a c = f a b + f b c`), we may assume `r a b` and `r b c` + are satisfied. We allow restricting to a subset specified by a predicate `p`."] +lemma multiplicative_of_is_total (p : α → Prop) + (hswap : ∀ {a b}, p a → p b → f a b * f b a = 1) + (hmul : ∀ {a b c}, r a b → r b c → p a → p b → p c → f a c = f a b * f b c) + {a b c : α} (pa : p a) (pb : p b) (pc : p c) : f a c = f a b * f b c := +begin + apply multiplicative_of_symmetric_of_is_total (λ a b, p a ∧ p b) r f (λ _ _, and.swap), + { simp_rw and_imp, exact @hswap }, + { exact λ a b c rab rbc pab pbc pac, hmul rab rbc pab.1 pab.2 pac.2 }, + exacts [⟨pa, pb⟩, ⟨pb, pc⟩, ⟨pa, pc⟩], +end + +end multiplicative diff --git a/src/algebra/group/inj_surj.lean b/src/algebra/group/inj_surj.lean index d9c43e3513bdd..68f1acbe5576b 100644 --- a/src/algebra/group/inj_surj.lean +++ b/src/algebra/group/inj_surj.lean @@ -179,6 +179,17 @@ protected def comm_monoid [comm_monoid M₂] (f : M₁ → M₂) (hf : injective comm_monoid M₁ := { .. hf.comm_semigroup f mul, .. hf.monoid f one mul npow } +/-- A type endowed with `0`, `1` and `+` is an additive commutative monoid with one, if it admits an +injective map that preserves `0`, `1` and `+` to an additive commutative monoid with one. +See note [reducible non-instances]. -/ +@[reducible] +protected def add_comm_monoid_with_one {M₁} [has_zero M₁] [has_one M₁] [has_add M₁] [has_smul ℕ M₁] + [has_nat_cast M₁] [add_comm_monoid_with_one M₂] (f : M₁ → M₂) (hf : injective f) (zero : f 0 = 0) + (one : f 1 = 1) (add : ∀ x y, f (x + y) = f x + f y) (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) + (nat_cast : ∀ n : ℕ, f n = n) : + add_comm_monoid_with_one M₁ := +{ ..hf.add_monoid_with_one f zero one add nsmul nat_cast, ..hf.add_comm_monoid f zero add nsmul } + /-- A type endowed with `1` and `*` is a cancel commutative monoid, if it admits an injective map that preserves `1` and `*` to a cancel commutative monoid. See note [reducible non-instances]. -/ @@ -303,6 +314,21 @@ protected def comm_group [comm_group M₂] (f : M₁ → M₂) (hf : injective f comm_group M₁ := { .. hf.comm_monoid f one mul npow, .. hf.group f one mul inv div npow zpow } +/-- A type endowed with `0`, `1` and `+` is an additive commutative group with one, if it admits an +injective map that preserves `0`, `1` and `+` to an additive commutative group with one. +See note [reducible non-instances]. -/ +@[reducible] +protected def add_comm_group_with_one {M₁} [has_zero M₁] [has_one M₁] [has_add M₁] [has_smul ℕ M₁] + [has_neg M₁] [has_sub M₁] [has_smul ℤ M₁] [has_nat_cast M₁] [has_int_cast M₁] + [add_comm_group_with_one M₂] (f : M₁ → M₂) (hf : injective f) + (zero : f 0 = 0) (one : f 1 = 1) (add : ∀ x y, f (x + y) = f x + f y) + (neg : ∀ x, f (- x) = - f x) (sub : ∀ x y, f (x - y) = f x - f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (zsmul : ∀ x (n : ℤ), f (n • x) = n • f x) + (nat_cast : ∀ n : ℕ, f n = n) (int_cast : ∀ n : ℤ, f n = n) : + add_comm_group_with_one M₁ := +{ ..hf.add_group_with_one f zero one add neg sub nsmul zsmul nat_cast int_cast, + ..hf.add_comm_monoid f zero add nsmul } + end injective /-! @@ -395,6 +421,19 @@ protected def comm_monoid [comm_monoid M₁] (f : M₁ → M₂) (hf : surjectiv comm_monoid M₂ := { .. hf.comm_semigroup f mul, .. hf.monoid f one mul npow } +/-- A type endowed with `0`, `1` and `+` is an additive monoid with one, +if it admits a surjective map that preserves `0`, `1` and `*` from an additive monoid with one. +See note [reducible non-instances]. -/ +@[reducible] +protected def add_comm_monoid_with_one + {M₂} [has_zero M₂] [has_one M₂] [has_add M₂] [has_smul ℕ M₂] [has_nat_cast M₂] + [add_comm_monoid_with_one M₁] (f : M₁ → M₂) (hf : surjective f) + (zero : f 0 = 0) (one : f 1 = 1) (add : ∀ x y, f (x + y) = f x + f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) + (nat_cast : ∀ n : ℕ, f n = n) : + add_comm_monoid_with_one M₂ := +{ ..hf.add_monoid_with_one f zero one add nsmul nat_cast, ..hf.add_comm_monoid _ zero _ nsmul } + /-- A type has an involutive inversion if it admits a surjective map that preserves `⁻¹` to a type which has an involutive inversion. -/ @[reducible, to_additive "A type has an involutive negation if it admits a surjective map that @@ -446,6 +485,7 @@ protected def group [group M₁] (f : M₁ → M₂) (hf : surjective f) /-- A type endowed with `0`, `1`, `+` is an additive group with one, if it admits a surjective map that preserves `0`, `1`, and `+` to an additive group with one. See note [reducible non-instances]. -/ +@[reducible] protected def add_group_with_one {M₂} [has_zero M₂] [has_one M₂] [has_add M₂] [has_neg M₂] [has_sub M₂] [has_smul ℕ M₂] [has_smul ℤ M₂] [has_nat_cast M₂] [has_int_cast M₂] @@ -475,6 +515,22 @@ protected def comm_group [comm_group M₁] (f : M₁ → M₂) (hf : surjective comm_group M₂ := { .. hf.comm_monoid f one mul npow, .. hf.group f one mul inv div npow zpow } +/-- A type endowed with `0`, `1`, `+` is an additive commutative group with one, if it admits a +surjective map that preserves `0`, `1`, and `+` to an additive commutative group with one. +See note [reducible non-instances]. -/ +@[reducible] +protected def add_comm_group_with_one + {M₂} [has_zero M₂] [has_one M₂] [has_add M₂] [has_neg M₂] [has_sub M₂] + [has_smul ℕ M₂] [has_smul ℤ M₂] [has_nat_cast M₂] [has_int_cast M₂] + [add_comm_group_with_one M₁] (f : M₁ → M₂) (hf : surjective f) + (zero : f 0 = 0) (one : f 1 = 1) (add : ∀ x y, f (x + y) = f x + f y) + (neg : ∀ x, f (- x) = - f x) (sub : ∀ x y, f (x - y) = f x - f y) + (nsmul : ∀ x (n : ℕ), f (n • x) = n • f x) (zsmul : ∀ x (n : ℤ), f (n • x) = n • f x) + (nat_cast : ∀ n : ℕ, f n = n) (int_cast : ∀ n : ℤ, f n = n) : + add_comm_group_with_one M₂ := +{ ..hf.add_group_with_one f zero one add neg sub nsmul zsmul nat_cast int_cast, + ..hf.add_comm_monoid _ zero add nsmul } + end surjective end function diff --git a/src/algebra/hom/centroid.lean b/src/algebra/hom/centroid.lean index f074ac30ecd29..42578ad72ff4a 100644 --- a/src/algebra/hom/centroid.lean +++ b/src/algebra/hom/centroid.lean @@ -9,6 +9,9 @@ import algebra.hom.group_instances /-! # Centroid homomorphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Let `A` be a (non unital, non associative) algebra. The centroid of `A` is the set of linear maps `T` on `A` such that `T` commutes with left and right multiplication, that is to say, for all `a` and `b` in `A`, diff --git a/src/algebra/module/graded_module.lean b/src/algebra/module/graded_module.lean index 44f16be81fc54..085b9dbb85c0d 100644 --- a/src/algebra/module/graded_module.lean +++ b/src/algebra/module/graded_module.lean @@ -7,6 +7,7 @@ Authors: Jujian Zhang import ring_theory.graded_algebra.basic import algebra.graded_mul_action import algebra.direct_sum.decomposition +import algebra.module.big_operators /-! # Graded Module @@ -31,37 +32,49 @@ namespace direct_sum open graded_monoid /-- A graded version of `distrib_mul_action`. -/ -class gdistrib_mul_action [add_monoid ι] [gmonoid A] [Π i, add_monoid $ M i] - extends graded_monoid.gmul_action A M := +class gdistrib_mul_action [add_monoid ι] [gmonoid A] [Π i, add_monoid (M i)] + extends gmul_action A M := (smul_add {i j} (a : A i) (b c : M j) : smul a (b + c) = smul a b + smul a c) (smul_zero {i j} (a : A i) : smul a (0 : M j) = 0) /-- A graded version of `module`. -/ -class gmodule [add_monoid ι] [Π i, add_monoid $ A i] [Π i, add_monoid $ M i] - [graded_monoid.gmonoid A] extends gdistrib_mul_action A M := +class gmodule [add_monoid ι] [Π i, add_monoid (A i)] [Π i, add_monoid (M i)] + [gmonoid A] extends gdistrib_mul_action A M := (add_smul {i j} (a a' : A i) (b : M j) : smul (a + a') b = smul a b + smul a' b) (zero_smul {i j} (b : M j) : smul (0 : A i) b = 0) -variables [add_monoid ι] [Π (i : ι), add_comm_monoid (A i)] [Π i, add_comm_monoid $ M i] +/-- A graded version of `semiring.to_module`. -/ +instance gsemiring.to_gmodule [decidable_eq ι] [add_monoid ι] + [Π (i : ι), add_comm_monoid (A i)] [gsemiring A] : + gmodule A A := +{ smul_add := λ _ _, gsemiring.mul_add, + smul_zero := λ i j, gsemiring.mul_zero, + add_smul := λ i j, gsemiring.add_mul, + zero_smul := λ i j, gsemiring.zero_mul, + ..gmonoid.to_gmul_action A } + +variables [add_monoid ι] [Π (i : ι), add_comm_monoid (A i)] [Π i, add_comm_monoid (M i)] /-- The piecewise multiplication from the `has_mul` instance, as a bundled homomorphism. -/ -@[simps] def gsmul_hom [graded_monoid.gmonoid A] [direct_sum.gmodule A M] {i j} : +@[simps] def gsmul_hom [gmonoid A] [gmodule A M] {i j} : A i →+ M j →+ M (i + j) := { to_fun := λ a, - { to_fun := λ b, graded_monoid.ghas_smul.smul a b, - map_zero' := direct_sum.gdistrib_mul_action.smul_zero _, - map_add' := direct_sum.gdistrib_mul_action.smul_add _ }, - map_zero' := add_monoid_hom.ext $ λ a, direct_sum.gmodule.zero_smul a, - map_add' := λ a₁ a₂, add_monoid_hom.ext $ λ b, direct_sum.gmodule.add_smul _ _ _} + { to_fun := λ b, ghas_smul.smul a b, + map_zero' := gdistrib_mul_action.smul_zero _, + map_add' := gdistrib_mul_action.smul_add _ }, + map_zero' := add_monoid_hom.ext $ λ a, gmodule.zero_smul a, + map_add' := λ a₁ a₂, add_monoid_hom.ext $ λ b, gmodule.add_smul _ _ _} + +namespace gmodule /-- For graded monoid `A` and a graded module `M` over `A`. `gmodule.smul_add_monoid_hom` is the `⨁ᵢ Aᵢ`-scalar multiplication on `⨁ᵢ Mᵢ` induced by `gsmul_hom`. -/ -def gmodule.smul_add_monoid_hom - [decidable_eq ι] [graded_monoid.gmonoid A] [direct_sum.gmodule A M] : +def smul_add_monoid_hom + [decidable_eq ι] [gmonoid A] [gmodule A M] : (⨁ i, A i) →+ (⨁ i, M i) →+ (⨁ i, M i) := -direct_sum.to_add_monoid $ λ i, - add_monoid_hom.flip $ direct_sum.to_add_monoid $ λ j, add_monoid_hom.flip $ - (direct_sum.of M _).comp_hom.comp $ gsmul_hom A M +to_add_monoid $ λ i, add_monoid_hom.flip $ + to_add_monoid $ λ j, add_monoid_hom.flip $ + (of M _).comp_hom.comp $ gsmul_hom A M section @@ -69,34 +82,33 @@ open graded_monoid direct_sum gmodule instance [decidable_eq ι] [gmonoid A] [gmodule A M] : has_smul (⨁ i, A i) (⨁ i, M i) := -{ smul := λ x y, gmodule.smul_add_monoid_hom A M x y } +{ smul := λ x y, smul_add_monoid_hom A M x y } -@[simp] lemma gmodule.smul_def [decidable_eq ι] [gmonoid A] [gmodule A M] - (x : ⨁ i, A i) (y : ⨁ i, M i) : x • y = gmodule.smul_add_monoid_hom _ _ x y := rfl +@[simp] lemma smul_def [decidable_eq ι] [gmonoid A] [gmodule A M] + (x : ⨁ i, A i) (y : ⨁ i, M i) : x • y = smul_add_monoid_hom _ _ x y := rfl -@[simp] lemma gmodule.smul_add_monoid_hom_apply_of_of [decidable_eq ι] [gmonoid A] [gmodule A M] +@[simp] lemma smul_add_monoid_hom_apply_of_of [decidable_eq ι] [gmonoid A] [gmodule A M] {i j} (x : A i) (y : M j) : - gmodule.smul_add_monoid_hom A M (direct_sum.of A i x) (direct_sum.of M j y) = - direct_sum.of M (i + j) (graded_monoid.ghas_smul.smul x y) := -by simp [gmodule.smul_add_monoid_hom] + smul_add_monoid_hom A M (direct_sum.of A i x) (of M j y) = + of M (i + j) (ghas_smul.smul x y) := +by simp [smul_add_monoid_hom] -@[simp] lemma gmodule.of_smul_of [decidable_eq ι] [gmonoid A] [gmodule A M] +@[simp] lemma of_smul_of [decidable_eq ι] [gmonoid A] [gmodule A M] {i j} (x : A i) (y : M j) : - direct_sum.of A i x • direct_sum.of M j y = - direct_sum.of M (i + j) (graded_monoid.ghas_smul.smul x y) := -gmodule.smul_add_monoid_hom_apply_of_of _ _ _ _ + direct_sum.of A i x • of M j y = of M (i + j) (ghas_smul.smul x y) := +smul_add_monoid_hom_apply_of_of _ _ _ _ open add_monoid_hom -- Almost identical to the proof of `direct_sum.one_mul` private lemma one_smul [decidable_eq ι] [gmonoid A] [gmodule A M] (x : ⨁ i, M i) : (1 : ⨁ i, A i) • x = x := -suffices gmodule.smul_add_monoid_hom A M 1 = add_monoid_hom.id (⨁ i, M i), +suffices smul_add_monoid_hom A M 1 = add_monoid_hom.id (⨁ i, M i), from add_monoid_hom.congr_fun this x, begin apply direct_sum.add_hom_ext, intros i xi, unfold has_one.one, - rw gmodule.smul_add_monoid_hom_apply_of_of, + rw smul_add_monoid_hom_apply_of_of, exact direct_sum.of_eq_of_graded_monoid_eq (one_smul (graded_monoid A) $ graded_monoid.mk i xi), end @@ -106,31 +118,33 @@ private lemma mul_smul [decidable_eq ι] [gsemiring A] [gmodule A M] suffices (smul_add_monoid_hom A M).comp_hom.comp (direct_sum.mul_hom A) -- `λ a b c, (a * b) • c` as a bundled hom = (add_monoid_hom.comp_hom add_monoid_hom.flip_hom $ - (smul_add_monoid_hom A M).flip.comp_hom.comp $ gmodule.smul_add_monoid_hom A M).flip, + (smul_add_monoid_hom A M).flip.comp_hom.comp $ smul_add_monoid_hom A M).flip, -- `λ a b c, a • (b • c)` as a bundled hom from add_monoid_hom.congr_fun (add_monoid_hom.congr_fun (add_monoid_hom.congr_fun this a) b) c, begin ext ai ax bi bx ci cx : 6, dsimp only [coe_comp, function.comp_app, comp_hom_apply_apply, flip_apply, flip_hom_apply], - rw [gmodule.smul_add_monoid_hom_apply_of_of, gmodule.smul_add_monoid_hom_apply_of_of, - direct_sum.mul_hom_of_of, gmodule.smul_add_monoid_hom_apply_of_of], + rw [smul_add_monoid_hom_apply_of_of, smul_add_monoid_hom_apply_of_of, + direct_sum.mul_hom_of_of, smul_add_monoid_hom_apply_of_of], exact direct_sum.of_eq_of_graded_monoid_eq (mul_smul (graded_monoid.mk ai ax) (graded_monoid.mk bi bx) (graded_monoid.mk ci cx)), end /-- The `module` derived from `gmodule A M`. -/ -instance gmodule.module [decidable_eq ι] [gsemiring A] [gmodule A M] : +instance module [decidable_eq ι] [gsemiring A] [gmodule A M] : module (⨁ i, A i) (⨁ i, M i) := { smul := (•), one_smul := one_smul _ _, mul_smul := mul_smul _ _, - smul_add := λ r, (gmodule.smul_add_monoid_hom A M r).map_add, - smul_zero := λ r, (gmodule.smul_add_monoid_hom A M r).map_zero, - add_smul := λ r s x, by simp only [gmodule.smul_def, map_add, add_monoid_hom.add_apply], - zero_smul := λ x, by simp only [gmodule.smul_def, map_zero, add_monoid_hom.zero_apply] } + smul_add := λ r, (smul_add_monoid_hom A M r).map_add, + smul_zero := λ r, (smul_add_monoid_hom A M r).map_zero, + add_smul := λ r s x, by simp only [smul_def, map_add, add_monoid_hom.add_apply], + zero_smul := λ x, by simp only [smul_def, map_zero, add_monoid_hom.zero_apply] } end +end gmodule + end direct_sum end @@ -188,8 +202,7 @@ turns `⨁ i, 𝓜 i` into an `A`-module def is_module [decidable_eq ι] [graded_ring 𝓐] : module A (⨁ i, 𝓜 i) := { smul := λ a b, direct_sum.decompose 𝓐 a • b, - ..module.comp_hom _ $ - (direct_sum.decompose_ring_equiv 𝓐 : A ≃+* ⨁ i, 𝓐 i).to_ring_hom } + .. module.comp_hom _ (direct_sum.decompose_ring_equiv 𝓐 : A ≃+* ⨁ i, 𝓐 i).to_ring_hom } local attribute [instance] graded_module.is_module @@ -207,12 +220,11 @@ def linear_equiv [decidable_eq ι] [graded_ring 𝓐] finset.sum_smul, finset.sum_congr rfl (λ i hi, _)], rw [ring_hom.id_apply, ← direct_sum.sum_support_decompose 𝓜 y, map_sum, finset.smul_sum, map_sum, finset.smul_sum, finset.sum_congr rfl (λ j hj, _)], - unfold has_smul.smul, - simp only [direct_sum.decompose_add_equiv_apply, direct_sum.decompose_coe, + simp only [(•), direct_sum.decompose_add_equiv_apply, direct_sum.decompose_coe, direct_sum.gmodule.smul_add_monoid_hom_apply_of_of], convert direct_sum.decompose_coe 𝓜 _, refl, end, - .. (direct_sum.decompose_add_equiv 𝓜) } + .. direct_sum.decompose_add_equiv 𝓜 } end graded_module diff --git a/src/algebra/module/submodule/basic.lean b/src/algebra/module/submodule/basic.lean index 1ffe24a81fd9e..c476f739d5f20 100644 --- a/src/algebra/module/submodule/basic.lean +++ b/src/algebra/module/submodule/basic.lean @@ -6,6 +6,8 @@ Authors: Nathaniel Thomas, Jeremy Avigad, Johannes Hölzl, Mario Carneiro import algebra.module.linear_map import algebra.module.equiv import group_theory.group_action.sub_mul_action +import group_theory.submonoid.membership + /-! # Submodules of a module diff --git a/src/algebra/periodic.lean b/src/algebra/periodic.lean index 442c4cafb585b..097404f758d10 100644 --- a/src/algebra/periodic.lean +++ b/src/algebra/periodic.lean @@ -3,11 +3,14 @@ Copyright (c) 2021 Benjamin Davidson. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Benjamin Davidson -/ +import algebra.big_operators.basic import algebra.field.opposite import algebra.module.basic import algebra.order.archimedean import data.int.parity import group_theory.coset +import group_theory.subgroup.zpowers +import group_theory.submonoid.membership /-! # Periodicity diff --git a/src/algebra/triv_sq_zero_ext.lean b/src/algebra/triv_sq_zero_ext.lean index d04a32e95c5aa..fe789d9ff6354 100644 --- a/src/algebra/triv_sq_zero_ext.lean +++ b/src/algebra/triv_sq_zero_ext.lean @@ -1,7 +1,7 @@ /- Copyright (c) 2020 Kenny Lau. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Authors: Kenny Lau +Authors: Kenny Lau, Eric Wieser -/ import algebra.algebra.basic @@ -76,12 +76,16 @@ section variables (M) @[simp] lemma fst_inl [has_zero M] (r : R) : (inl r : tsze R M).fst = r := rfl @[simp] lemma snd_inl [has_zero M] (r : R) : (inl r : tsze R M).snd = 0 := rfl +@[simp] lemma fst_comp_inl [has_zero M] : fst ∘ (inl : R → tsze R M) = id := rfl +@[simp] lemma snd_comp_inl [has_zero M] : snd ∘ (inl : R → tsze R M) = 0 := rfl end section variables (R) @[simp] lemma fst_inr [has_zero R] (m : M) : (inr m : tsze R M).fst = 0 := rfl @[simp] lemma snd_inr [has_zero R] (m : M) : (inr m : tsze R M).snd = m := rfl +@[simp] lemma fst_comp_inr [has_zero R] : fst ∘ (inr : M → tsze R M) = 0 := rfl +@[simp] lemma snd_comp_inr [has_zero R] : snd ∘ (inr : M → tsze R M) = id := rfl end lemma inl_injective [has_zero M] : function.injective (inl : R → tsze R M) := @@ -109,6 +113,9 @@ prod.has_zero instance [has_add R] [has_add M] : has_add (tsze R M) := prod.has_add +instance [has_sub R] [has_sub M] : has_sub (tsze R M) := +prod.has_sub + instance [has_neg R] [has_neg M] : has_neg (tsze R M) := prod.has_neg @@ -170,6 +177,11 @@ prod.module @[simp] lemma fst_neg [has_neg R] [has_neg M] (x : tsze R M) : (-x).fst = -x.fst := rfl @[simp] lemma snd_neg [has_neg R] [has_neg M] (x : tsze R M) : (-x).snd = -x.snd := rfl +@[simp] lemma fst_sub [has_sub R] [has_sub M] (x₁ x₂ : tsze R M) : + (x₁ - x₂).fst = x₁.fst - x₂.fst := rfl +@[simp] lemma snd_sub [has_sub R] [has_sub M] (x₁ x₂ : tsze R M) : + (x₁ - x₂).snd = x₁.snd - x₂.snd := rfl + @[simp] lemma fst_smul [has_smul S R] [has_smul S M] (s : S) (x : tsze R M) : (s • x).fst = s • x.fst := rfl @[simp] lemma snd_smul [has_smul S R] [has_smul S M] (s : S) (x : tsze R M) : @@ -184,10 +196,14 @@ variables (M) (inl (r₁ + r₂) : tsze R M) = inl r₁ + inl r₂ := ext rfl (add_zero 0).symm -@[simp] lemma inl_neg [has_neg R] [add_group M] (r : R) : +@[simp] lemma inl_neg [has_neg R] [sub_neg_zero_monoid M] (r : R) : (inl (-r) : tsze R M) = -inl r := ext rfl neg_zero.symm +@[simp] lemma inl_sub [has_sub R] [sub_neg_zero_monoid M] (r₁ r₂ : R) : + (inl (r₁ - r₂) : tsze R M) = inl r₁ - inl r₂ := +ext rfl (sub_zero _).symm + @[simp] lemma inl_smul [monoid S] [add_monoid M] [has_smul S R] [distrib_mul_action S M] (s : S) (r : R) : (inl (s • r) : tsze R M) = s • inl r := ext rfl (smul_zero s).symm @@ -203,10 +219,14 @@ variables (R) (inr (m₁ + m₂) : tsze R M) = inr m₁ + inr m₂ := ext (add_zero 0).symm rfl -@[simp] lemma inr_neg [add_group R] [has_neg M] (m : M) : +@[simp] lemma inr_neg [sub_neg_zero_monoid R] [has_neg M] (m : M) : (inr (-m) : tsze R M) = -inr m := ext neg_zero.symm rfl +@[simp] lemma inr_sub [sub_neg_zero_monoid R] [has_sub M] (m₁ m₂ : M) : + (inr (m₁ - m₂) : tsze R M) = inr m₁ - inr m₂ := +ext (sub_zero _).symm rfl + @[simp] lemma inr_smul [has_zero R] [has_zero S] [smul_with_zero S R] [has_smul S M] (r : S) (m : M) : (inr (r • m) : tsze R M) = r • inr m := ext (smul_zero _).symm rfl @@ -307,12 +327,33 @@ instance [monoid R] [add_monoid M] [distrib_mul_action R M] : mul_one_class (tsz .. triv_sq_zero_ext.has_mul } instance [add_monoid_with_one R] [add_monoid M] : add_monoid_with_one (tsze R M) := -{ nat_cast := λ n, (n, 0), +{ nat_cast := λ n, inl n, nat_cast_zero := by simp [nat.cast], nat_cast_succ := λ _, by ext; simp [nat.cast], .. triv_sq_zero_ext.add_monoid, .. triv_sq_zero_ext.has_one } +@[simp] lemma fst_nat_cast [add_monoid_with_one R] [add_monoid M] (n : ℕ) : + (n : tsze R M).fst = n := rfl +@[simp] lemma snd_nat_cast [add_monoid_with_one R] [add_monoid M] (n : ℕ) : + (n : tsze R M).snd = 0 := rfl +@[simp] lemma inl_nat_cast [add_monoid_with_one R] [add_monoid M] (n : ℕ) : + (inl n : tsze R M) = n := rfl + +instance [add_group_with_one R] [add_group M] : add_group_with_one (tsze R M) := +{ int_cast := λ z, inl z, + int_cast_of_nat := λ n, ext (int.cast_coe_nat _) rfl, + int_cast_neg_succ_of_nat := λ n, ext (int.cast_neg_succ_of_nat _) neg_zero.symm, + .. triv_sq_zero_ext.add_group, + .. triv_sq_zero_ext.add_monoid_with_one } + +@[simp] lemma fst_int_cast [add_group_with_one R] [add_group M] (z : ℤ) : + (z : tsze R M).fst = z := rfl +@[simp] lemma snd_int_cast [add_group_with_one R] [add_group M] (z : ℤ) : + (z : tsze R M).snd = 0 := rfl +@[simp] lemma inl_int_cast [add_group_with_one R] [add_group M] (z : ℤ) : + (inl z : tsze R M) = z := rfl + instance [semiring R] [add_comm_monoid M] [module R M] : non_assoc_semiring (tsze R M) := { zero_mul := λ x, ext (zero_mul x.1) $ show (0 : R) • x.2 + x.1 • 0 = 0, by rw [zero_smul, zero_add, smul_zero], @@ -330,11 +371,40 @@ instance [semiring R] [add_comm_monoid M] [module R M] : non_assoc_semiring (tsz .. triv_sq_zero_ext.mul_one_class, .. triv_sq_zero_ext.add_comm_monoid } +instance [ring R] [add_comm_group M] [module R M] : non_assoc_ring (tsze R M) := +{ .. triv_sq_zero_ext.add_group_with_one, + .. triv_sq_zero_ext.non_assoc_semiring } + +instance [comm_monoid R] [add_monoid M] [distrib_mul_action R M] : has_pow (tsze R M) ℕ := +⟨λ x n, ⟨x.fst^n, n • x.fst ^ n.pred • x.snd⟩⟩ + +@[simp] lemma fst_pow [comm_monoid R] [add_monoid M] [distrib_mul_action R M] + (x : tsze R M) (n : ℕ) : + fst (x ^ n) = x.fst ^ n := rfl + +@[simp] lemma snd_pow [comm_monoid R] [add_monoid M] [distrib_mul_action R M] + (x : tsze R M) (n : ℕ) : + snd (x ^ n) = n • x.fst ^ n.pred • x.snd := rfl + +@[simp] lemma inl_pow [comm_monoid R] [add_monoid M] [distrib_mul_action R M] + (r : R) (n : ℕ) : + (inl r ^ n : tsze R M) = inl (r ^ n) := +ext rfl $ by simp + instance [comm_monoid R] [add_monoid M] [distrib_mul_action R M] : monoid (tsze R M) := { mul_assoc := λ x y z, ext (mul_assoc x.1 y.1 z.1) $ show (x.1 * y.1) • z.2 + z.1 • (x.1 • y.2 + y.1 • x.2) = x.1 • (y.1 • z.2 + z.1 • y.2) + (y.1 * z.1) • x.2, by simp_rw [smul_add, ← mul_smul, add_assoc, mul_comm], + npow := λ n x, x ^ n, + npow_zero' := λ x, ext (pow_zero x.fst) (zero_smul _ _), + npow_succ' := λ n x, ext (pow_succ _ _) begin + dsimp, + rw [smul_comm (_ : R) n (_ : M), smul_smul, succ_nsmul'], + cases n, + { simp_rw [zero_smul ]}, + { rw [nat.pred_succ, pow_succ] } + end, .. triv_sq_zero_ext.mul_one_class } instance [comm_monoid R] [add_comm_monoid M] [distrib_mul_action R M] : comm_monoid (tsze R M) := @@ -346,6 +416,10 @@ instance [comm_semiring R] [add_comm_monoid M] [module R M] : comm_semiring (tsz { .. triv_sq_zero_ext.comm_monoid, .. triv_sq_zero_ext.non_assoc_semiring } +instance [comm_ring R] [add_comm_group M] [module R M] : comm_ring (tsze R M) := +{ .. triv_sq_zero_ext.non_assoc_ring, + .. triv_sq_zero_ext.comm_semiring } + variables (R M) /-- The canonical inclusion of rings `R → triv_sq_zero_ext R M`. -/ @@ -365,7 +439,8 @@ variables [comm_semiring S] [comm_semiring R] [add_comm_monoid M] variables [algebra S R] [module S M] [module R M] [is_scalar_tower S R M] instance algebra' : algebra S (tsze R M) := -{ commutes' := λ r x, mul_comm _ _, +{ smul := (•), + commutes' := λ r x, mul_comm _ _, smul_def' := λ r x, ext (algebra.smul_def _ _) $ show r • x.2 = algebra_map S R r • x.2 + x.1 • 0, by rw [smul_zero, add_zero, algebra_map_smul], .. (triv_sq_zero_ext.inl_hom R M).comp (algebra_map S R) } diff --git a/src/analysis/bounded_variation.lean b/src/analysis/bounded_variation.lean index bb435c8f434af..aa78a0f8eda05 100644 --- a/src/analysis/bounded_variation.lean +++ b/src/analysis/bounded_variation.lean @@ -6,6 +6,7 @@ Authors: Sébastien Gouëzel import measure_theory.measure.lebesgue import analysis.calculus.monotone import data.set.function +import algebra.group.basic /-! # Functions of bounded variation @@ -21,6 +22,7 @@ almost everywhere. * `has_bounded_variation_on f s` registers that the variation of `f` on `s` is finite. * `has_locally_bounded_variation f s` registers that `f` has finite variation on any compact subinterval of `s`. +* `variation_on_from_to f s a b` is the signed variation of `f` on `s ∩ Icc a b`, converted to `ℝ`. * `evariation_on.Icc_add_Icc` states that the variation of `f` on `[a, c]` is the sum of its variations on `[a, b]` and `[b, c]`. @@ -686,60 +688,126 @@ lemma monotone_on.has_locally_bounded_variation_on {f : α → ℝ} {s : set α} has_locally_bounded_variation_on f s := λ a b as bs, ((hf.evariation_on_le as bs).trans_lt ennreal.of_real_lt_top).ne +/-- +The **signed** variation of `f` on the interval `Icc a b` intersected with the set `s`, +squashed to a real (therefore only really meaningful if the variation is finite) +-/ +noncomputable def variation_on_from_to (f : α → E) (s : set α) (a b : α) : ℝ := +if a ≤ b then (evariation_on f (s ∩ Icc a b)).to_real else + - (evariation_on f (s ∩ Icc b a)).to_real + +namespace variation_on_from_to + +variables (f : α → E) (s : set α) + +@[protected] +lemma self (a : α) : variation_on_from_to f s a a = 0 := +begin + dsimp only [variation_on_from_to], + rw [if_pos le_rfl, Icc_self, evariation_on.subsingleton, ennreal.zero_to_real], + exact λ x hx y hy, hx.2.trans hy.2.symm, +end + +@[protected] +lemma nonneg_of_le {a b : α} (h : a ≤ b) : 0 ≤ variation_on_from_to f s a b := +by simp only [variation_on_from_to, if_pos h, ennreal.to_real_nonneg] + +@[protected] +lemma eq_neg_swap (a b : α) : + variation_on_from_to f s a b = - variation_on_from_to f s b a := +begin + rcases lt_trichotomy a b with ab|rfl|ba, + { simp only [variation_on_from_to, if_pos ab.le, if_neg ab.not_le, neg_neg], }, + { simp only [self, neg_zero], }, + { simp only [variation_on_from_to, if_pos ba.le, if_neg ba.not_le, neg_neg], }, +end + +@[protected] +lemma nonpos_of_ge {a b : α} (h : b ≤ a) : variation_on_from_to f s a b ≤ 0 := +begin + rw eq_neg_swap, + exact neg_nonpos_of_nonneg (nonneg_of_le f s h), +end + +@[protected] +lemma eq_of_le {a b : α} (h : a ≤ b) : + variation_on_from_to f s a b = (evariation_on f (s ∩ Icc a b)).to_real := if_pos h + +@[protected] +lemma eq_of_ge {a b : α} (h : b ≤ a) : + variation_on_from_to f s a b = - (evariation_on f (s ∩ Icc b a)).to_real := +by rw [eq_neg_swap, neg_inj, eq_of_le f s h] + +@[protected] +lemma add {f : α → E} {s : set α} (hf : has_locally_bounded_variation_on f s) + {a b c : α} (ha : a ∈ s) (hb : b ∈ s) (hc : c ∈ s) : + variation_on_from_to f s a b + variation_on_from_to f s b c = variation_on_from_to f s a c := +begin + symmetry, + refine additive_of_is_total (≤) (variation_on_from_to f s) (∈s) _ _ ha hb hc, + { rintro x y xs ys, + simp only [eq_neg_swap f s y x, subtype.coe_mk, add_right_neg, forall_true_left], }, + { rintro x y z xy yz xs ys zs, + rw [eq_of_le f s xy, eq_of_le f s yz, eq_of_le f s (xy.trans yz), + ←ennreal.to_real_add (hf x y xs ys) (hf y z ys zs), + evariation_on.Icc_add_Icc f xy yz ys], }, +end + +variables {f} {s} + +@[protected] +lemma monotone_on (hf : has_locally_bounded_variation_on f s) + {a : α} (as : a ∈ s) : monotone_on (variation_on_from_to f s a) s := +begin + rintro b bs c cs bc, + rw ←add hf as bs cs, + exact le_add_of_nonneg_right (nonneg_of_le f s bc), +end + +@[protected] +lemma antitone_on (hf : has_locally_bounded_variation_on f s) + {b : α} (bs : b ∈ s) : antitone_on (λ a, variation_on_from_to f s a b) s := +begin + rintro a as c cs ac, + dsimp only, + rw ←add hf as cs bs, + exact le_add_of_nonneg_left (nonneg_of_le f s ac), +end + +@[protected] +lemma sub_self_monotone_on {f : α → ℝ} {s : set α} + (hf : has_locally_bounded_variation_on f s) {a : α} (as : a ∈ s) : + monotone_on (variation_on_from_to f s a - f) s := +begin + rintro b bs c cs bc, + rw [pi.sub_apply, pi.sub_apply, le_sub_iff_add_le, add_comm_sub, ← le_sub_iff_add_le'], + calc f c - f b + ≤ |f c - f b| : le_abs_self _ + ... = dist (f b) (f c) : by rw [dist_comm, real.dist_eq] + ... ≤ variation_on_from_to f s b c : by + { rw [eq_of_le f s bc, dist_edist], + apply ennreal.to_real_mono (hf b c bs cs), + apply evariation_on.edist_le f, + exacts [⟨bs, le_rfl, bc⟩, ⟨cs, bc, le_rfl⟩] } + ... = variation_on_from_to f s a c - variation_on_from_to f s a b : + by rw [←add hf as bs cs, add_sub_cancel'] +end + +end variation_on_from_to + /-- If a real valued function has bounded variation on a set, then it is a difference of monotone functions there. -/ lemma has_locally_bounded_variation_on.exists_monotone_on_sub_monotone_on {f : α → ℝ} {s : set α} (h : has_locally_bounded_variation_on f s) : ∃ (p q : α → ℝ), monotone_on p s ∧ monotone_on q s ∧ f = p - q := begin - rcases eq_empty_or_nonempty s with rfl|hs, + rcases eq_empty_or_nonempty s with rfl|⟨c, cs⟩, { exact ⟨f, 0, subsingleton_empty.monotone_on _, subsingleton_empty.monotone_on _, - by simp only [tsub_zero]⟩ }, - rcases hs with ⟨c, cs⟩, - let p := λ x, if c ≤ x then (evariation_on f (s ∩ Icc c x)).to_real - else -(evariation_on f (s ∩ Icc x c)).to_real, - have hp : monotone_on p s, - { assume x xs y ys hxy, - dsimp only [p], - split_ifs with hcx hcy hcy, - { have : evariation_on f (s ∩ Icc c x) + evariation_on f (s ∩ Icc x y) - = evariation_on f (s ∩ Icc c y), from evariation_on.Icc_add_Icc f hcx hxy xs, - rw [← this, ennreal.to_real_add (h c x cs xs) (h x y xs ys)], - exact le_add_of_le_of_nonneg le_rfl ennreal.to_real_nonneg }, - { exact (lt_irrefl _ ((not_le.1 hcy).trans_le (hcx.trans hxy))).elim }, - { exact (neg_nonpos.2 ennreal.to_real_nonneg).trans ennreal.to_real_nonneg }, - { simp only [neg_le_neg_iff], - have : evariation_on f (s ∩ Icc x y) + evariation_on f (s ∩ Icc y c) - = evariation_on f (s ∩ Icc x c), from evariation_on.Icc_add_Icc f hxy (not_le.1 hcy).le ys, - rw [← this, ennreal.to_real_add (h x y xs ys) (h y c ys cs), add_comm], - exact le_add_of_le_of_nonneg le_rfl ennreal.to_real_nonneg } }, - have hq : monotone_on (λ x, p x - f x) s, - { assume x xs y ys hxy, - dsimp only [p], - split_ifs with hcx hcy hcy, - { have : evariation_on f (s ∩ Icc c x) + evariation_on f (s ∩ Icc x y) - = evariation_on f (s ∩ Icc c y), from evariation_on.Icc_add_Icc f hcx hxy xs, - rw [← this, ennreal.to_real_add (h c x cs xs) (h x y xs ys)], - suffices : f y - f x ≤ (evariation_on f (s ∩ Icc x y)).to_real, by linarith, - exact (h x y xs ys).sub_le ⟨ys, hxy, le_rfl⟩ ⟨xs, le_rfl, hxy⟩ }, - { exact (lt_irrefl _ ((not_le.1 hcy).trans_le (hcx.trans hxy))).elim }, - { suffices : f y - f x ≤ (evariation_on f (s ∩ Icc x c)).to_real - + (evariation_on f (s ∩ Icc c y)).to_real, by linarith, - rw [← ennreal.to_real_add (h x c xs cs) (h c y cs ys), - evariation_on.Icc_add_Icc f (not_le.1 hcx).le hcy cs], - exact (h x y xs ys).sub_le ⟨ys, hxy, le_rfl⟩ ⟨xs, le_rfl, hxy⟩ }, - { have : evariation_on f (s ∩ Icc x y) + evariation_on f (s ∩ Icc y c) - = evariation_on f (s ∩ Icc x c), from evariation_on.Icc_add_Icc f hxy (not_le.1 hcy).le ys, - rw [← this, ennreal.to_real_add (h x y xs ys) (h y c ys cs)], - suffices : f y - f x ≤ (evariation_on f (s ∩ Icc x y)).to_real, by linarith, - exact (h x y xs ys).sub_le ⟨ys, hxy, le_rfl⟩ ⟨xs, le_rfl, hxy⟩ } }, - refine ⟨p, λ x, p x - f x, hp, hq, _⟩, - ext x, - dsimp, - abel, + (sub_zero f).symm⟩ }, + { exact ⟨_, _, variation_on_from_to.monotone_on h cs, + variation_on_from_to.sub_self_monotone_on h cs, (sub_sub_cancel _ _).symm⟩ }, end - /-! ## Lipschitz functions and bounded variation -/ lemma lipschitz_on_with.comp_evariation_on_le {f : E → F} {C : ℝ≥0} {t : set E} @@ -885,4 +953,3 @@ lemma lipschitz_with.ae_differentiable_at {C : ℝ≥0} {f : ℝ → V} (h : lipschitz_with C f) : ∀ᵐ x, differentiable_at ℝ f x := (h.has_locally_bounded_variation_on univ).ae_differentiable_at - diff --git a/src/analysis/calculus/conformal/normed_space.lean b/src/analysis/calculus/conformal/normed_space.lean index 5fdcd41193683..03ad42c84f295 100644 --- a/src/analysis/calculus/conformal/normed_space.lean +++ b/src/analysis/calculus/conformal/normed_space.lean @@ -26,7 +26,7 @@ if it is real differentiable at that point and its differential `is_conformal_li In `analysis.calculus.conformal.inner_product`: * `conformal_at_iff`: an equivalent definition of the conformality of a map -In `geometry.euclidean.basic`: +In `geometry.euclidean.angle.unoriented.conformal`: * `conformal_at.preserves_angle`: if a map is conformal at `x`, then its differential preserves all angles at `x` diff --git a/src/analysis/calculus/mean_value.lean b/src/analysis/calculus/mean_value.lean index f64fd61e91c71..636a38d50cb4e 100644 --- a/src/analysis/calculus/mean_value.lean +++ b/src/analysis/calculus/mean_value.lean @@ -5,7 +5,7 @@ Authors: Sébastien Gouëzel, Yury Kudryashov -/ import analysis.calculus.local_extr import analysis.convex.slope -import analysis.convex.topology +import analysis.convex.normed import data.complex.is_R_or_C /-! diff --git a/src/analysis/complex/cauchy_integral.lean b/src/analysis/complex/cauchy_integral.lean index 2cbc77d447e99..b4c6d6e96f5ee 100644 --- a/src/analysis/complex/cauchy_integral.lean +++ b/src/analysis/complex/cauchy_integral.lean @@ -176,7 +176,10 @@ begin set F : (ℝ × ℝ) → E := f ∘ e, set F' : (ℝ × ℝ) → (ℝ × ℝ) →L[ℝ] E := λ p, (f' (e p)).comp (e : (ℝ × ℝ) →L[ℝ] ℂ), have hF' : ∀ p : ℝ × ℝ, (-(I • F' p)) (1, 0) + F' p (0, 1) = -(I • f' (e p) 1 - f' (e p) I), - { rintro ⟨x, y⟩, simp [F', he₁, he₂, ← sub_eq_neg_add], }, + { rintro ⟨x, y⟩, + simp only [continuous_linear_map.neg_apply, continuous_linear_map.smul_apply, F', + continuous_linear_map.comp_apply, continuous_linear_equiv.coe_coe, he₁, he₂, + neg_add_eq_sub, neg_sub], }, set R : set (ℝ × ℝ) := [z.re, w.re] ×ˢ [w.im, z.im], set t : set (ℝ × ℝ) := e ⁻¹' s, rw [uIcc_comm z.im] at Hc Hi, rw [min_comm z.im, max_comm z.im] at Hd, diff --git a/src/analysis/complex/upper_half_plane/topology.lean b/src/analysis/complex/upper_half_plane/topology.lean index a47570b6e516b..65fba4c6df834 100644 --- a/src/analysis/complex/upper_half_plane/topology.lean +++ b/src/analysis/complex/upper_half_plane/topology.lean @@ -5,7 +5,7 @@ Authors: Yury G. Kudryashov -/ import analysis.complex.upper_half_plane.basic import analysis.convex.contractible -import analysis.convex.topology +import analysis.convex.normed import analysis.convex.complex import analysis.complex.re_im_topology import topology.homotopy.contractible diff --git a/src/analysis/convex/normed.lean b/src/analysis/convex/normed.lean new file mode 100644 index 0000000000000..b5ac0d553abd5 --- /dev/null +++ b/src/analysis/convex/normed.lean @@ -0,0 +1,143 @@ +/- +Copyright (c) 2020 Yury Kudryashov. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Alexander Bentkamp, Yury Kudryashov +-/ +import analysis.convex.jensen +import analysis.convex.topology +import analysis.normed.group.pointwise +import analysis.normed_space.ray + +/-! +# Topological and metric properties of convex sets in normed spaces + +We prove the following facts: + +* `convex_on_norm`, `convex_on_dist` : norm and distance to a fixed point is convex on any convex + set; +* `convex_on_univ_norm`, `convex_on_univ_dist` : norm and distance to a fixed point is convex on + the whole space; +* `convex_hull_ediam`, `convex_hull_diam` : convex hull of a set has the same (e)metric diameter + as the original set; +* `bounded_convex_hull` : convex hull of a set is bounded if and only if the original set + is bounded. +* `bounded_std_simplex`, `is_closed_std_simplex`, `compact_std_simplex`: topological properties + of the standard simplex. +-/ + +variables {ι : Type*} {E : Type*} + +open metric set +open_locale pointwise convex + +variables [seminormed_add_comm_group E] [normed_space ℝ E] {s t : set E} + +/-- The norm on a real normed space is convex on any convex set. See also `seminorm.convex_on` +and `convex_on_univ_norm`. -/ +lemma convex_on_norm (hs : convex ℝ s) : convex_on ℝ s norm := +⟨hs, λ x hx y hy a b ha hb hab, + calc ‖a • x + b • y‖ ≤ ‖a • x‖ + ‖b • y‖ : norm_add_le _ _ + ... = a * ‖x‖ + b * ‖y‖ + : by rw [norm_smul, norm_smul, real.norm_of_nonneg ha, real.norm_of_nonneg hb]⟩ + +/-- The norm on a real normed space is convex on the whole space. See also `seminorm.convex_on` +and `convex_on_norm`. -/ +lemma convex_on_univ_norm : convex_on ℝ univ (norm : E → ℝ) := convex_on_norm convex_univ + +lemma convex_on_dist (z : E) (hs : convex ℝ s) : convex_on ℝ s (λ z', dist z' z) := +by simpa [dist_eq_norm, preimage_preimage] + using (convex_on_norm (hs.translate (-z))).comp_affine_map + (affine_map.id ℝ E - affine_map.const ℝ E z) + +lemma convex_on_univ_dist (z : E) : convex_on ℝ univ (λz', dist z' z) := +convex_on_dist z convex_univ + +lemma convex_ball (a : E) (r : ℝ) : convex ℝ (metric.ball a r) := +by simpa only [metric.ball, sep_univ] using (convex_on_univ_dist a).convex_lt r + +lemma convex_closed_ball (a : E) (r : ℝ) : convex ℝ (metric.closed_ball a r) := +by simpa only [metric.closed_ball, sep_univ] using (convex_on_univ_dist a).convex_le r + +lemma convex.thickening (hs : convex ℝ s) (δ : ℝ) : convex ℝ (thickening δ s) := +by { rw ←add_ball_zero, exact hs.add (convex_ball 0 _) } + +lemma convex.cthickening (hs : convex ℝ s) (δ : ℝ) : convex ℝ (cthickening δ s) := +begin + obtain hδ | hδ := le_total 0 δ, + { rw cthickening_eq_Inter_thickening hδ, + exact convex_Inter₂ (λ _ _, hs.thickening _) }, + { rw cthickening_of_nonpos hδ, + exact hs.closure } +end + +/-- Given a point `x` in the convex hull of `s` and a point `y`, there exists a point +of `s` at distance at least `dist x y` from `y`. -/ +lemma convex_hull_exists_dist_ge {s : set E} {x : E} (hx : x ∈ convex_hull ℝ s) (y : E) : + ∃ x' ∈ s, dist x y ≤ dist x' y := +(convex_on_dist y (convex_convex_hull ℝ _)).exists_ge_of_mem_convex_hull hx + +/-- Given a point `x` in the convex hull of `s` and a point `y` in the convex hull of `t`, +there exist points `x' ∈ s` and `y' ∈ t` at distance at least `dist x y`. -/ +lemma convex_hull_exists_dist_ge2 {s t : set E} {x y : E} + (hx : x ∈ convex_hull ℝ s) (hy : y ∈ convex_hull ℝ t) : + ∃ (x' ∈ s) (y' ∈ t), dist x y ≤ dist x' y' := +begin + rcases convex_hull_exists_dist_ge hx y with ⟨x', hx', Hx'⟩, + rcases convex_hull_exists_dist_ge hy x' with ⟨y', hy', Hy'⟩, + use [x', hx', y', hy'], + exact le_trans Hx' (dist_comm y x' ▸ dist_comm y' x' ▸ Hy') +end + +/-- Emetric diameter of the convex hull of a set `s` equals the emetric diameter of `s. -/ +@[simp] lemma convex_hull_ediam (s : set E) : + emetric.diam (convex_hull ℝ s) = emetric.diam s := +begin + refine (emetric.diam_le $ λ x hx y hy, _).antisymm (emetric.diam_mono $ subset_convex_hull ℝ s), + rcases convex_hull_exists_dist_ge2 hx hy with ⟨x', hx', y', hy', H⟩, + rw edist_dist, + apply le_trans (ennreal.of_real_le_of_real H), + rw ← edist_dist, + exact emetric.edist_le_diam_of_mem hx' hy' +end + +/-- Diameter of the convex hull of a set `s` equals the emetric diameter of `s. -/ +@[simp] lemma convex_hull_diam (s : set E) : + metric.diam (convex_hull ℝ s) = metric.diam s := +by simp only [metric.diam, convex_hull_ediam] + +/-- Convex hull of `s` is bounded if and only if `s` is bounded. -/ +@[simp] lemma bounded_convex_hull {s : set E} : + metric.bounded (convex_hull ℝ s) ↔ metric.bounded s := +by simp only [metric.bounded_iff_ediam_ne_top, convex_hull_ediam] + +@[priority 100] +instance normed_space.path_connected : path_connected_space E := +topological_add_group.path_connected + +@[priority 100] +instance normed_space.loc_path_connected : loc_path_connected_space E := +loc_path_connected_of_bases (λ x, metric.nhds_basis_ball) + (λ x r r_pos, (convex_ball x r).is_path_connected $ by simp [r_pos]) + +lemma dist_add_dist_of_mem_segment {x y z : E} (h : y ∈ [x -[ℝ] z]) : + dist x y + dist y z = dist x z := +begin + simp only [dist_eq_norm, mem_segment_iff_same_ray] at *, + simpa only [sub_add_sub_cancel', norm_sub_rev] using h.norm_add.symm +end + +/-- The set of vectors in the same ray as `x` is connected. -/ +lemma is_connected_set_of_same_ray (x : E) : is_connected {y | same_ray ℝ x y} := +begin + by_cases hx : x = 0, { simpa [hx] using is_connected_univ }, + simp_rw ←exists_nonneg_left_iff_same_ray hx, + exact is_connected_Ici.image _ ((continuous_id.smul continuous_const).continuous_on) +end + +/-- The set of nonzero vectors in the same ray as the nonzero vector `x` is connected. -/ +lemma is_connected_set_of_same_ray_and_ne_zero {x : E} (hx : x ≠ 0) : + is_connected {y | same_ray ℝ x y ∧ y ≠ 0} := +begin + simp_rw ←exists_pos_left_iff_same_ray_and_ne_zero hx, + exact is_connected_Ioi.image _ ((continuous_id.smul continuous_const).continuous_on) +end diff --git a/src/analysis/convex/side.lean b/src/analysis/convex/side.lean index 827cd0dff4c6b..a6941b8d4cc43 100644 --- a/src/analysis/convex/side.lean +++ b/src/analysis/convex/side.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Myers -/ import analysis.convex.between -import analysis.convex.topology +import analysis.convex.normed import analysis.normed.group.add_torsor /-! diff --git a/src/analysis/convex/strict_convex_space.lean b/src/analysis/convex/strict_convex_space.lean index 39d0309c9d23d..751345c147d10 100644 --- a/src/analysis/convex/strict_convex_space.lean +++ b/src/analysis/convex/strict_convex_space.lean @@ -3,8 +3,8 @@ Copyright (c) 2022 Yaël Dillies. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yaël Dillies, Yury Kudryashov -/ +import analysis.convex.normed import analysis.convex.strict -import analysis.convex.topology import analysis.normed.order.basic import analysis.normed_space.add_torsor import analysis.normed_space.pointwise diff --git a/src/analysis/convex/topology.lean b/src/analysis/convex/topology.lean index 078d9b74569c5..f36f01234f2f7 100644 --- a/src/analysis/convex/topology.lean +++ b/src/analysis/convex/topology.lean @@ -3,35 +3,25 @@ Copyright (c) 2020 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Alexander Bentkamp, Yury Kudryashov -/ -import analysis.convex.jensen +import analysis.convex.combination import analysis.convex.strict -import analysis.normed.group.pointwise -import topology.algebra.module.finite_dimension -import analysis.normed_space.ray import topology.path_connected import topology.algebra.affine +import topology.algebra.module.basic /-! -# Topological and metric properties of convex sets +# Topological properties of convex sets We prove the following facts: * `convex.interior` : interior of a convex set is convex; * `convex.closure` : closure of a convex set is convex; * `set.finite.compact_convex_hull` : convex hull of a finite set is compact; -* `set.finite.is_closed_convex_hull` : convex hull of a finite set is closed; -* `convex_on_norm`, `convex_on_dist` : norm and distance to a fixed point is convex on any convex - set; -* `convex_on_univ_norm`, `convex_on_univ_dist` : norm and distance to a fixed point is convex on - the whole space; -* `convex_hull_ediam`, `convex_hull_diam` : convex hull of a set has the same (e)metric diameter - as the original set; -* `bounded_convex_hull` : convex hull of a set is bounded if and only if the original set - is bounded. -* `bounded_std_simplex`, `is_closed_std_simplex`, `compact_std_simplex`: topological properties - of the standard simplex; +* `set.finite.is_closed_convex_hull` : convex hull of a finite set is closed. -/ +assert_not_exists has_norm + variables {ι : Type*} {E : Type*} open metric set @@ -53,11 +43,10 @@ lemma std_simplex_subset_closed_ball : std_simplex ℝ ι ⊆ metric.closed_ball 0 1 := begin assume f hf, - rw [metric.mem_closed_ball, dist_zero_right], - refine (nnreal.coe_one ▸ nnreal.coe_le_coe.2 $ finset.sup_le $ λ x hx, _), - change |f x| ≤ 1, - rw [abs_of_nonneg $ hf.1 x], - exact (mem_Icc_of_mem_std_simplex hf x).2 + rw [metric.mem_closed_ball, dist_pi_le_iff zero_le_one], + intros x, + rw [pi.zero_apply, real.dist_0_eq_abs, abs_of_nonneg $ hf.1 x], + exact (mem_Icc_of_mem_std_simplex hf x).2, end variable (ι) @@ -329,120 +318,3 @@ protected lemma topological_add_group.path_connected : path_connected_space E := path_connected_space_iff_univ.mpr $ convex_univ.is_path_connected ⟨(0 : E), trivial⟩ end has_continuous_smul - -/-! ### Normed vector space -/ - -section normed_space -variables [seminormed_add_comm_group E] [normed_space ℝ E] {s t : set E} - -/-- The norm on a real normed space is convex on any convex set. See also `seminorm.convex_on` -and `convex_on_univ_norm`. -/ -lemma convex_on_norm (hs : convex ℝ s) : convex_on ℝ s norm := -⟨hs, λ x hx y hy a b ha hb hab, - calc ‖a • x + b • y‖ ≤ ‖a • x‖ + ‖b • y‖ : norm_add_le _ _ - ... = a * ‖x‖ + b * ‖y‖ - : by rw [norm_smul, norm_smul, real.norm_of_nonneg ha, real.norm_of_nonneg hb]⟩ - -/-- The norm on a real normed space is convex on the whole space. See also `seminorm.convex_on` -and `convex_on_norm`. -/ -lemma convex_on_univ_norm : convex_on ℝ univ (norm : E → ℝ) := convex_on_norm convex_univ - -lemma convex_on_dist (z : E) (hs : convex ℝ s) : convex_on ℝ s (λ z', dist z' z) := -by simpa [dist_eq_norm, preimage_preimage] - using (convex_on_norm (hs.translate (-z))).comp_affine_map - (affine_map.id ℝ E - affine_map.const ℝ E z) - -lemma convex_on_univ_dist (z : E) : convex_on ℝ univ (λz', dist z' z) := -convex_on_dist z convex_univ - -lemma convex_ball (a : E) (r : ℝ) : convex ℝ (metric.ball a r) := -by simpa only [metric.ball, sep_univ] using (convex_on_univ_dist a).convex_lt r - -lemma convex_closed_ball (a : E) (r : ℝ) : convex ℝ (metric.closed_ball a r) := -by simpa only [metric.closed_ball, sep_univ] using (convex_on_univ_dist a).convex_le r - -lemma convex.thickening (hs : convex ℝ s) (δ : ℝ) : convex ℝ (thickening δ s) := -by { rw ←add_ball_zero, exact hs.add (convex_ball 0 _) } - -lemma convex.cthickening (hs : convex ℝ s) (δ : ℝ) : convex ℝ (cthickening δ s) := -begin - obtain hδ | hδ := le_total 0 δ, - { rw cthickening_eq_Inter_thickening hδ, - exact convex_Inter₂ (λ _ _, hs.thickening _) }, - { rw cthickening_of_nonpos hδ, - exact hs.closure } -end - -/-- Given a point `x` in the convex hull of `s` and a point `y`, there exists a point -of `s` at distance at least `dist x y` from `y`. -/ -lemma convex_hull_exists_dist_ge {s : set E} {x : E} (hx : x ∈ convex_hull ℝ s) (y : E) : - ∃ x' ∈ s, dist x y ≤ dist x' y := -(convex_on_dist y (convex_convex_hull ℝ _)).exists_ge_of_mem_convex_hull hx - -/-- Given a point `x` in the convex hull of `s` and a point `y` in the convex hull of `t`, -there exist points `x' ∈ s` and `y' ∈ t` at distance at least `dist x y`. -/ -lemma convex_hull_exists_dist_ge2 {s t : set E} {x y : E} - (hx : x ∈ convex_hull ℝ s) (hy : y ∈ convex_hull ℝ t) : - ∃ (x' ∈ s) (y' ∈ t), dist x y ≤ dist x' y' := -begin - rcases convex_hull_exists_dist_ge hx y with ⟨x', hx', Hx'⟩, - rcases convex_hull_exists_dist_ge hy x' with ⟨y', hy', Hy'⟩, - use [x', hx', y', hy'], - exact le_trans Hx' (dist_comm y x' ▸ dist_comm y' x' ▸ Hy') -end - -/-- Emetric diameter of the convex hull of a set `s` equals the emetric diameter of `s. -/ -@[simp] lemma convex_hull_ediam (s : set E) : - emetric.diam (convex_hull ℝ s) = emetric.diam s := -begin - refine (emetric.diam_le $ λ x hx y hy, _).antisymm (emetric.diam_mono $ subset_convex_hull ℝ s), - rcases convex_hull_exists_dist_ge2 hx hy with ⟨x', hx', y', hy', H⟩, - rw edist_dist, - apply le_trans (ennreal.of_real_le_of_real H), - rw ← edist_dist, - exact emetric.edist_le_diam_of_mem hx' hy' -end - -/-- Diameter of the convex hull of a set `s` equals the emetric diameter of `s. -/ -@[simp] lemma convex_hull_diam (s : set E) : - metric.diam (convex_hull ℝ s) = metric.diam s := -by simp only [metric.diam, convex_hull_ediam] - -/-- Convex hull of `s` is bounded if and only if `s` is bounded. -/ -@[simp] lemma bounded_convex_hull {s : set E} : - metric.bounded (convex_hull ℝ s) ↔ metric.bounded s := -by simp only [metric.bounded_iff_ediam_ne_top, convex_hull_ediam] - -@[priority 100] -instance normed_space.path_connected : path_connected_space E := -topological_add_group.path_connected - -@[priority 100] -instance normed_space.loc_path_connected : loc_path_connected_space E := -loc_path_connected_of_bases (λ x, metric.nhds_basis_ball) - (λ x r r_pos, (convex_ball x r).is_path_connected $ by simp [r_pos]) - -lemma dist_add_dist_of_mem_segment {x y z : E} (h : y ∈ [x -[ℝ] z]) : - dist x y + dist y z = dist x z := -begin - simp only [dist_eq_norm, mem_segment_iff_same_ray] at *, - simpa only [sub_add_sub_cancel', norm_sub_rev] using h.norm_add.symm -end - -/-- The set of vectors in the same ray as `x` is connected. -/ -lemma is_connected_set_of_same_ray (x : E) : is_connected {y | same_ray ℝ x y} := -begin - by_cases hx : x = 0, { simpa [hx] using is_connected_univ }, - simp_rw ←exists_nonneg_left_iff_same_ray hx, - exact is_connected_Ici.image _ ((continuous_id.smul continuous_const).continuous_on) -end - -/-- The set of nonzero vectors in the same ray as the nonzero vector `x` is connected. -/ -lemma is_connected_set_of_same_ray_and_ne_zero {x : E} (hx : x ≠ 0) : - is_connected {y | same_ray ℝ x y ∧ y ≠ 0} := -begin - simp_rw ←exists_pos_left_iff_same_ray_and_ne_zero hx, - exact is_connected_Ioi.image _ ((continuous_id.smul continuous_const).continuous_on) -end - -end normed_space diff --git a/src/analysis/normed/field/basic.lean b/src/analysis/normed/field/basic.lean index 8b14076599c1b..9850a7146b863 100644 --- a/src/analysis/normed/field/basic.lean +++ b/src/analysis/normed/field/basic.lean @@ -6,7 +6,6 @@ Authors: Patrick Massot, Johannes Hölzl import algebra.algebra.subalgebra.basic import analysis.normed.group.basic import topology.instances.ennreal -import topology.instances.rat /-! # Normed fields @@ -676,62 +675,21 @@ lemma normed_add_comm_group.tendsto_at_top' [nonempty α] [semilattice_sup α] [ (at_top_basis_Ioi.tendsto_iff metric.nhds_basis_ball).trans (by simp [dist_eq_norm]) instance : normed_comm_ring ℤ := -{ norm := λ n, ‖(n : ℝ)‖, - norm_mul := λ m n, le_of_eq $ by simp only [norm, int.cast_mul, abs_mul], - dist_eq := λ m n, by simp only [int.dist_eq, norm, int.cast_sub], - mul_comm := mul_comm } - -@[norm_cast] lemma int.norm_cast_real (m : ℤ) : ‖(m : ℝ)‖ = ‖m‖ := rfl - -lemma int.norm_eq_abs (n : ℤ) : ‖n‖ = |n| := rfl - -@[simp] lemma int.norm_coe_nat (n : ℕ) : ‖(n : ℤ)‖ = n := by simp [int.norm_eq_abs] - -lemma nnreal.coe_nat_abs (n : ℤ) : (n.nat_abs : ℝ≥0) = ‖n‖₊ := -nnreal.eq $ calc ((n.nat_abs : ℝ≥0) : ℝ) - = (n.nat_abs : ℤ) : by simp only [int.cast_coe_nat, nnreal.coe_nat_cast] - ... = |n| : by simp only [int.coe_nat_abs, int.cast_abs] - ... = ‖n‖ : rfl - -lemma int.abs_le_floor_nnreal_iff (z : ℤ) (c : ℝ≥0) : |z| ≤ ⌊c⌋₊ ↔ ‖z‖₊ ≤ c := -begin - rw [int.abs_eq_nat_abs, int.coe_nat_le, nat.le_floor_iff (zero_le c)], - congr', - exact nnreal.coe_nat_abs z, -end +{ norm_mul := λ m n, le_of_eq $ by simp only [norm, int.cast_mul, abs_mul], + mul_comm := mul_comm, + .. int.normed_add_comm_group } instance : norm_one_class ℤ := ⟨by simp [← int.norm_cast_real]⟩ instance : normed_field ℚ := -{ norm := λ r, ‖(r : ℝ)‖, - norm_mul' := λ r₁ r₂, by simp only [norm, rat.cast_mul, abs_mul], - dist_eq := λ r₁ r₂, by simp only [rat.dist_eq, norm, rat.cast_sub] } +{ norm_mul' := λ r₁ r₂, by simp only [norm, rat.cast_mul, abs_mul], + .. rat.normed_add_comm_group } instance : densely_normed_field ℚ := { lt_norm_lt := λ r₁ r₂ h₀ hr, let ⟨q, h⟩ := exists_rat_btwn hr in ⟨q, by { unfold norm, rwa abs_of_pos (h₀.trans_lt h.1) } ⟩ } -@[norm_cast, simp] lemma rat.norm_cast_real (r : ℚ) : ‖(r : ℝ)‖ = ‖r‖ := rfl - -@[norm_cast, simp] lemma int.norm_cast_rat (m : ℤ) : ‖(m : ℚ)‖ = ‖m‖ := -by rw [← rat.norm_cast_real, ← int.norm_cast_real]; congr' 1; norm_cast - --- Now that we've installed the norm on `ℤ`, --- we can state some lemmas about `zsmul`. -section -variables [seminormed_comm_group α] - -@[to_additive norm_zsmul_le] -lemma norm_zpow_le_mul_norm (n : ℤ) (a : α) : ‖a^n‖ ≤ ‖n‖ * ‖a‖ := -by rcases n.eq_coe_or_neg with ⟨n, rfl | rfl⟩; simpa using norm_pow_le_mul_norm n a - -@[to_additive nnnorm_zsmul_le] -lemma nnnorm_zpow_le_mul_norm (n : ℤ) (a : α) : ‖a^n‖₊ ≤ ‖n‖₊ * ‖a‖₊ := -by simpa only [← nnreal.coe_le_coe, nnreal.coe_mul] using norm_zpow_le_mul_norm n a - -end - section ring_hom_isometric variables {R₁ : Type*} {R₂ : Type*} {R₃ : Type*} diff --git a/src/analysis/normed/group/basic.lean b/src/analysis/normed/group/basic.lean index f82dcde4fcc56..32bdec02e5262 100644 --- a/src/analysis/normed/group/basic.lean +++ b/src/analysis/normed/group/basic.lean @@ -6,6 +6,7 @@ Authors: Patrick Massot, Johannes Hölzl, Yaël Dillies import analysis.normed.group.seminorm import order.liminf_limsup import topology.algebra.uniform_group +import topology.instances.rat import topology.metric_space.algebra import topology.metric_space.isometry import topology.sequences @@ -1271,6 +1272,61 @@ end end real +namespace int + +instance : normed_add_comm_group ℤ := +{ norm := λ n, ‖(n : ℝ)‖, + dist_eq := λ m n, by simp only [int.dist_eq, norm, int.cast_sub] } + +@[norm_cast] lemma norm_cast_real (m : ℤ) : ‖(m : ℝ)‖ = ‖m‖ := rfl + +lemma norm_eq_abs (n : ℤ) : ‖n‖ = |n| := rfl + +@[simp] lemma norm_coe_nat (n : ℕ) : ‖(n : ℤ)‖ = n := by simp [int.norm_eq_abs] + +lemma _root_.nnreal.coe_nat_abs (n : ℤ) : (n.nat_abs : ℝ≥0) = ‖n‖₊ := +nnreal.eq $ calc ((n.nat_abs : ℝ≥0) : ℝ) + = (n.nat_abs : ℤ) : by simp only [int.cast_coe_nat, nnreal.coe_nat_cast] + ... = |n| : by simp only [int.coe_nat_abs, int.cast_abs] + ... = ‖n‖ : rfl + +lemma abs_le_floor_nnreal_iff (z : ℤ) (c : ℝ≥0) : |z| ≤ ⌊c⌋₊ ↔ ‖z‖₊ ≤ c := +begin + rw [int.abs_eq_nat_abs, int.coe_nat_le, nat.le_floor_iff (zero_le c)], + congr', + exact nnreal.coe_nat_abs z, +end + +end int + +namespace rat + +instance : normed_add_comm_group ℚ := +{ norm := λ r, ‖(r : ℝ)‖, + dist_eq := λ r₁ r₂, by simp only [rat.dist_eq, norm, rat.cast_sub] } + +@[norm_cast, simp] lemma norm_cast_real (r : ℚ) : ‖(r : ℝ)‖ = ‖r‖ := rfl + +@[norm_cast, simp] lemma _root_.int.norm_cast_rat (m : ℤ) : ‖(m : ℚ)‖ = ‖m‖ := +by rw [← rat.norm_cast_real, ← int.norm_cast_real]; congr' 1; norm_cast + +end rat + +-- Now that we've installed the norm on `ℤ`, +-- we can state some lemmas about `zsmul`. +section +variables [seminormed_comm_group α] + +@[to_additive norm_zsmul_le] +lemma norm_zpow_le_mul_norm (n : ℤ) (a : α) : ‖a^n‖ ≤ ‖n‖ * ‖a‖ := +by rcases n.eq_coe_or_neg with ⟨n, rfl | rfl⟩; simpa using norm_pow_le_mul_norm n a + +@[to_additive nnnorm_zsmul_le] +lemma nnnorm_zpow_le_mul_norm (n : ℤ) (a : α) : ‖a^n‖₊ ≤ ‖n‖₊ * ‖a‖₊ := +by simpa only [← nnreal.coe_le_coe, nnreal.coe_mul] using norm_zpow_le_mul_norm n a + +end + namespace lipschitz_with variables [pseudo_emetric_space α] {K Kf Kg : ℝ≥0} {f g : α → E} diff --git a/src/analysis/normed/group/controlled_closure.lean b/src/analysis/normed/group/controlled_closure.lean new file mode 100644 index 0000000000000..535ac9b93a592 --- /dev/null +++ b/src/analysis/normed/group/controlled_closure.lean @@ -0,0 +1,121 @@ +/- +Copyright (c) 2021 Patrick Massot. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Patrick Massot +-/ +import analysis.normed.group.hom +import analysis.specific_limits.normed + +/-! # Extending a backward bound on a normed group homomorphism from a dense set + +Possible TODO (from the PR's review, https://github.com/leanprover-community/mathlib/pull/8498 ): +"This feels a lot like the second step in the proof of the Banach open mapping theorem +(`exists_preimage_norm_le`) ... wonder if it would be possible to refactor it using one of [the +lemmas in this file]." +-/ + +open filter finset +open_locale topological_space big_operators + +variables {G : Type*} [normed_add_comm_group G] [complete_space G] +variables {H : Type*} [normed_add_comm_group H] + +/-- Given `f : normed_add_group_hom G H` for some complete `G` and a subgroup `K` of `H`, if every +element `x` of `K` has a preimage under `f` whose norm is at most `C*‖x‖` then the same holds for +elements of the (topological) closure of `K` with constant `C+ε` instead of `C`, for any +positive `ε`. +-/ +lemma controlled_closure_of_complete {f : normed_add_group_hom G H} {K : add_subgroup H} + {C ε : ℝ} (hC : 0 < C) (hε : 0 < ε) (hyp : f.surjective_on_with K C) : + f.surjective_on_with K.topological_closure (C + ε) := +begin + rintros (h : H) (h_in : h ∈ K.topological_closure), + /- We first get rid of the easy case where `h = 0`.-/ + by_cases hyp_h : h = 0, + { rw hyp_h, + use 0, + simp }, + /- The desired preimage will be constructed as the sum of a series. Convergence of + the series will be guaranteed by completeness of `G`. We first write `h` as the sum + of a sequence `v` of elements of `K` which starts close to `h` and then quickly goes to zero. + The sequence `b` below quantifies this. -/ + set b : ℕ → ℝ := λ i, (1/2)^i*(ε*‖h‖/2)/C, + have b_pos : ∀ i, 0 < b i, + { intro i, + field_simp [b, hC], + exact div_pos (mul_pos hε (norm_pos_iff.mpr hyp_h)) + (mul_pos (by norm_num : (0 : ℝ) < 2^i*2) hC) }, + obtain ⟨v : ℕ → H, lim_v : tendsto (λ (n : ℕ), ∑ k in range (n + 1), v k) at_top (𝓝 h), + v_in : ∀ n, v n ∈ K, hv₀ : ‖v 0 - h‖ < b 0, hv : ∀ n > 0, ‖v n‖ < b n⟩ := + controlled_sum_of_mem_closure h_in b_pos, + /- The controlled surjectivity assumption on `f` allows to build preimages `u n` for all + elements `v n` of the `v` sequence.-/ + have : ∀ n, ∃ m' : G, f m' = v n ∧ ‖m'‖ ≤ C * ‖v n‖ := λ (n : ℕ), hyp (v n) (v_in n), + choose u hu hnorm_u using this, + /- The desired series `s` is then obtained by summing `u`. We then check our choice of + `b` ensures `s` is Cauchy. -/ + set s : ℕ → G := λ n, ∑ k in range (n+1), u k, + have : cauchy_seq s, + { apply normed_add_comm_group.cauchy_series_of_le_geometric'' (by norm_num) one_half_lt_one, + rintro n (hn : n ≥ 1), + calc ‖u n‖ ≤ C*‖v n‖ : hnorm_u n + ... ≤ C * b n : mul_le_mul_of_nonneg_left (hv _ $ nat.succ_le_iff.mp hn).le hC.le + ... = (1/2)^n * (ε * ‖h‖/2) : by simp [b, mul_div_cancel' _ hC.ne.symm] + ... = (ε * ‖h‖/2) * (1/2)^n : mul_comm _ _ }, + /- We now show that the limit `g` of `s` is the desired preimage. -/ + obtain ⟨g : G, hg⟩ := cauchy_seq_tendsto_of_complete this, + refine ⟨g, _, _⟩, + { /- We indeed get a preimage. First note: -/ + have : f ∘ s = λ n, ∑ k in range (n + 1), v k, + { ext n, + simp [map_sum, hu] }, + /- In the above equality, the left-hand-side converges to `f g` by continuity of `f` and + definition of `g` while the right-hand-side converges to `h` by construction of `v` so + `g` is indeed a preimage of `h`. -/ + rw ← this at lim_v, + exact tendsto_nhds_unique ((f.continuous.tendsto g).comp hg) lim_v }, + { /- Then we need to estimate the norm of `g`, using our careful choice of `b`. -/ + suffices : ∀ n, ‖s n‖ ≤ (C + ε) * ‖h‖, + from le_of_tendsto' (continuous_norm.continuous_at.tendsto.comp hg) this, + intros n, + have hnorm₀ : ‖u 0‖ ≤ C*b 0 + C*‖h‖, + { have := calc + ‖v 0‖ ≤ ‖h‖ + ‖v 0 - h‖ : norm_le_insert' _ _ + ... ≤ ‖h‖ + b 0 : by apply add_le_add_left hv₀.le, + calc ‖u 0‖ ≤ C*‖v 0‖ : hnorm_u 0 + ... ≤ C*(‖h‖ + b 0) : mul_le_mul_of_nonneg_left this hC.le + ... = C * b 0 + C * ‖h‖ : by rw [add_comm, mul_add] }, + have : ∑ k in range (n + 1), C * b k ≤ ε * ‖h‖ := calc + ∑ k in range (n + 1), C * b k = (∑ k in range (n + 1), (1 / 2) ^ k) * (ε * ‖h‖ / 2) : + by simp only [b, mul_div_cancel' _ hC.ne.symm, ← sum_mul] + ... ≤ 2 * (ε * ‖h‖ / 2) : mul_le_mul_of_nonneg_right (sum_geometric_two_le _) + (by nlinarith [hε, norm_nonneg h]) + ... = ε * ‖h‖ : mul_div_cancel' _ two_ne_zero, + calc ‖s n‖ ≤ ∑ k in range (n+1), ‖u k‖ : norm_sum_le _ _ + ... = ∑ k in range n, ‖u (k + 1)‖ + ‖u 0‖ : sum_range_succ' _ _ + ... ≤ ∑ k in range n, C*‖v (k + 1)‖ + ‖u 0‖ : add_le_add_right (sum_le_sum (λ _ _, hnorm_u _)) _ + ... ≤ ∑ k in range n, C*b (k+1) + (C*b 0 + C*‖h‖) : + add_le_add (sum_le_sum (λ k _, mul_le_mul_of_nonneg_left (hv _ k.succ_pos).le hC.le)) hnorm₀ + ... = ∑ k in range (n+1), C*b k + C*‖h‖ : by rw [← add_assoc, sum_range_succ'] + ... ≤ (C+ε)*‖h‖ : by { rw [add_comm, add_mul], apply add_le_add_left this } } +end + +/-- Given `f : normed_add_group_hom G H` for some complete `G`, if every element `x` of the image of +an isometric immersion `j : normed_add_group_hom K H` has a preimage under `f` whose norm is at most +`C*‖x‖` then the same holds for elements of the (topological) closure of this image with constant +`C+ε` instead of `C`, for any positive `ε`. +This is useful in particular if `j` is the inclusion of a normed group into its completion +(in this case the closure is the full target group). +-/ +lemma controlled_closure_range_of_complete {f : normed_add_group_hom G H} + {K : Type*} [seminormed_add_comm_group K] {j : normed_add_group_hom K H} (hj : ∀ x, ‖j x‖ = ‖x‖) + {C ε : ℝ} (hC : 0 < C) (hε : 0 < ε) (hyp : ∀ k, ∃ g, f g = j k ∧ ‖g‖ ≤ C*‖k‖) : + f.surjective_on_with j.range.topological_closure (C + ε) := +begin + replace hyp : ∀ h ∈ j.range, ∃ g, f g = h ∧ ‖g‖ ≤ C*‖h‖, + { intros h h_in, + rcases (j.mem_range _).mp h_in with ⟨k, rfl⟩, + rw hj, + exact hyp k }, + exact controlled_closure_of_complete hC hε hyp +end diff --git a/src/analysis/normed/group/hom.lean b/src/analysis/normed/group/hom.lean index 62bd8e1f499b3..4954266502108 100644 --- a/src/analysis/normed/group/hom.lean +++ b/src/analysis/normed/group/hom.lean @@ -3,8 +3,7 @@ Copyright (c) 2021 Johan Commelin. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin -/ - -import analysis.specific_limits.normed +import analysis.normed.group.basic /-! # Normed groups homomorphisms @@ -755,110 +754,3 @@ norm_lift_le _ _ _ hφ end equalizer end normed_add_group_hom - -section controlled_closure -open filter finset -open_locale topological_space -variables {G : Type*} [normed_add_comm_group G] [complete_space G] -variables {H : Type*} [normed_add_comm_group H] - -/-- Given `f : normed_add_group_hom G H` for some complete `G` and a subgroup `K` of `H`, if every -element `x` of `K` has a preimage under `f` whose norm is at most `C*‖x‖` then the same holds for -elements of the (topological) closure of `K` with constant `C+ε` instead of `C`, for any -positive `ε`. --/ -lemma controlled_closure_of_complete {f : normed_add_group_hom G H} {K : add_subgroup H} - {C ε : ℝ} (hC : 0 < C) (hε : 0 < ε) (hyp : f.surjective_on_with K C) : - f.surjective_on_with K.topological_closure (C + ε) := -begin - rintros (h : H) (h_in : h ∈ K.topological_closure), - /- We first get rid of the easy case where `h = 0`.-/ - by_cases hyp_h : h = 0, - { rw hyp_h, - use 0, - simp }, - /- The desired preimage will be constructed as the sum of a series. Convergence of - the series will be guaranteed by completeness of `G`. We first write `h` as the sum - of a sequence `v` of elements of `K` which starts close to `h` and then quickly goes to zero. - The sequence `b` below quantifies this. -/ - set b : ℕ → ℝ := λ i, (1/2)^i*(ε*‖h‖/2)/C, - have b_pos : ∀ i, 0 < b i, - { intro i, - field_simp [b, hC], - exact div_pos (mul_pos hε (norm_pos_iff.mpr hyp_h)) - (mul_pos (by norm_num : (0 : ℝ) < 2^i*2) hC) }, - obtain ⟨v : ℕ → H, lim_v : tendsto (λ (n : ℕ), ∑ k in range (n + 1), v k) at_top (𝓝 h), - v_in : ∀ n, v n ∈ K, hv₀ : ‖v 0 - h‖ < b 0, hv : ∀ n > 0, ‖v n‖ < b n⟩ := - controlled_sum_of_mem_closure h_in b_pos, - /- The controlled surjectivity assumption on `f` allows to build preimages `u n` for all - elements `v n` of the `v` sequence.-/ - have : ∀ n, ∃ m' : G, f m' = v n ∧ ‖m'‖ ≤ C * ‖v n‖ := λ (n : ℕ), hyp (v n) (v_in n), - choose u hu hnorm_u using this, - /- The desired series `s` is then obtained by summing `u`. We then check our choice of - `b` ensures `s` is Cauchy. -/ - set s : ℕ → G := λ n, ∑ k in range (n+1), u k, - have : cauchy_seq s, - { apply normed_add_comm_group.cauchy_series_of_le_geometric'' (by norm_num) one_half_lt_one, - rintro n (hn : n ≥ 1), - calc ‖u n‖ ≤ C*‖v n‖ : hnorm_u n - ... ≤ C * b n : mul_le_mul_of_nonneg_left (hv _ $ nat.succ_le_iff.mp hn).le hC.le - ... = (1/2)^n * (ε * ‖h‖/2) : by simp [b, mul_div_cancel' _ hC.ne.symm] - ... = (ε * ‖h‖/2) * (1/2)^n : mul_comm _ _ }, - /- We now show that the limit `g` of `s` is the desired preimage. -/ - obtain ⟨g : G, hg⟩ := cauchy_seq_tendsto_of_complete this, - refine ⟨g, _, _⟩, - { /- We indeed get a preimage. First note: -/ - have : f ∘ s = λ n, ∑ k in range (n + 1), v k, - { ext n, - simp [map_sum, hu] }, - /- In the above equality, the left-hand-side converges to `f g` by continuity of `f` and - definition of `g` while the right-hand-side converges to `h` by construction of `v` so - `g` is indeed a preimage of `h`. -/ - rw ← this at lim_v, - exact tendsto_nhds_unique ((f.continuous.tendsto g).comp hg) lim_v }, - { /- Then we need to estimate the norm of `g`, using our careful choice of `b`. -/ - suffices : ∀ n, ‖s n‖ ≤ (C + ε) * ‖h‖, - from le_of_tendsto' (continuous_norm.continuous_at.tendsto.comp hg) this, - intros n, - have hnorm₀ : ‖u 0‖ ≤ C*b 0 + C*‖h‖, - { have := calc - ‖v 0‖ ≤ ‖h‖ + ‖v 0 - h‖ : norm_le_insert' _ _ - ... ≤ ‖h‖ + b 0 : by apply add_le_add_left hv₀.le, - calc ‖u 0‖ ≤ C*‖v 0‖ : hnorm_u 0 - ... ≤ C*(‖h‖ + b 0) : mul_le_mul_of_nonneg_left this hC.le - ... = C * b 0 + C * ‖h‖ : by rw [add_comm, mul_add] }, - have : ∑ k in range (n + 1), C * b k ≤ ε * ‖h‖ := calc - ∑ k in range (n + 1), C * b k = (∑ k in range (n + 1), (1 / 2) ^ k) * (ε * ‖h‖ / 2) : - by simp only [b, mul_div_cancel' _ hC.ne.symm, ← sum_mul] - ... ≤ 2 * (ε * ‖h‖ / 2) : mul_le_mul_of_nonneg_right (sum_geometric_two_le _) - (by nlinarith [hε, norm_nonneg h]) - ... = ε * ‖h‖ : mul_div_cancel' _ two_ne_zero, - calc ‖s n‖ ≤ ∑ k in range (n+1), ‖u k‖ : norm_sum_le _ _ - ... = ∑ k in range n, ‖u (k + 1)‖ + ‖u 0‖ : sum_range_succ' _ _ - ... ≤ ∑ k in range n, C*‖v (k + 1)‖ + ‖u 0‖ : add_le_add_right (sum_le_sum (λ _ _, hnorm_u _)) _ - ... ≤ ∑ k in range n, C*b (k+1) + (C*b 0 + C*‖h‖) : - add_le_add (sum_le_sum (λ k _, mul_le_mul_of_nonneg_left (hv _ k.succ_pos).le hC.le)) hnorm₀ - ... = ∑ k in range (n+1), C*b k + C*‖h‖ : by rw [← add_assoc, sum_range_succ'] - ... ≤ (C+ε)*‖h‖ : by { rw [add_comm, add_mul], apply add_le_add_left this } } -end - -/-- Given `f : normed_add_group_hom G H` for some complete `G`, if every element `x` of the image of -an isometric immersion `j : normed_add_group_hom K H` has a preimage under `f` whose norm is at most -`C*‖x‖` then the same holds for elements of the (topological) closure of this image with constant -`C+ε` instead of `C`, for any positive `ε`. -This is useful in particular if `j` is the inclusion of a normed group into its completion -(in this case the closure is the full target group). --/ -lemma controlled_closure_range_of_complete {f : normed_add_group_hom G H} - {K : Type*} [seminormed_add_comm_group K] {j : normed_add_group_hom K H} (hj : ∀ x, ‖j x‖ = ‖x‖) - {C ε : ℝ} (hC : 0 < C) (hε : 0 < ε) (hyp : ∀ k, ∃ g, f g = j k ∧ ‖g‖ ≤ C*‖k‖) : - f.surjective_on_with j.range.topological_closure (C + ε) := -begin - replace hyp : ∀ h ∈ j.range, ∃ g, f g = h ∧ ‖g‖ ≤ C*‖h‖, - { intros h h_in, - rcases (j.mem_range _).mp h_in with ⟨k, rfl⟩, - rw hj, - exact hyp k }, - exact controlled_closure_of_complete hC hε hyp -end -end controlled_closure diff --git a/src/analysis/normed/group/quotient.lean b/src/analysis/normed/group/quotient.lean index 836809a7fec30..963dda9704818 100644 --- a/src/analysis/normed/group/quotient.lean +++ b/src/analysis/normed/group/quotient.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Patrick Massot. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Patrick Massot, Riccardo Brasca -/ +import analysis.normed_space.basic import analysis.normed.group.hom /-! diff --git a/src/analysis/normed_space/continuous_linear_map.lean b/src/analysis/normed_space/continuous_linear_map.lean index bb3bc4ae41afb..50c0db7872b3f 100644 --- a/src/analysis/normed_space/continuous_linear_map.lean +++ b/src/analysis/normed_space/continuous_linear_map.lean @@ -175,7 +175,7 @@ linear_equiv.to_continuous_linear_equiv_of_bounds f a a⁻¹ end seminormed -/- ## The span of a single vector -/ +/-! ## The span of a single vector -/ section seminormed diff --git a/src/analysis/special_functions/bernstein.lean b/src/analysis/special_functions/bernstein.lean index 67df5d5c9e191..a76e8f9207c3e 100644 --- a/src/analysis/special_functions/bernstein.lean +++ b/src/analysis/special_functions/bernstein.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Scott Morrison -/ import algebra.order.field.basic +import analysis.specific_limits.basic import ring_theory.polynomial.bernstein import topology.continuous_function.polynomial import topology.continuous_function.compact diff --git a/src/analysis/special_functions/exp.lean b/src/analysis/special_functions/exp.lean index 7960020f35ea6..553a2a3e86118 100644 --- a/src/analysis/special_functions/exp.lean +++ b/src/analysis/special_functions/exp.lean @@ -3,9 +3,10 @@ Copyright (c) 2018 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes, Abhimanyu Pallavi Sudhir, Jean Lo, Calle Sönne -/ +import analysis.asymptotics.theta import analysis.complex.basic +import analysis.specific_limits.normed import data.complex.exponential -import analysis.asymptotics.theta /-! # Complex and real exponential diff --git a/src/analysis/special_functions/gamma.lean b/src/analysis/special_functions/gamma.lean index e0d175e767c5b..d6775ee17112c 100644 --- a/src/analysis/special_functions/gamma.lean +++ b/src/analysis/special_functions/gamma.lean @@ -17,7 +17,33 @@ integral `Γ(s) = ∫ x in Ioi 0, exp (-x) * x ^ (s - 1)` in the range where thi We show that this integral satisfies `Γ(1) = 1` and `Γ(s + 1) = s * Γ(s)`; hence we can define `Γ(s)` for all `s` as the unique function satisfying this recurrence and agreeing with Euler's integral in the convergence range. In the complex case we also prove that the resulting function is -holomorphic on `ℂ` away from the points `{-n : n ∈ ℤ}`. +holomorphic on `ℂ` away from the points `{-n : n ∈ ℕ}`. + +## Main statements (real case) + +* `real.Gamma` : the `Γ` function (of a real variable). +* `real.Gamma_eq_integral` : for `0 < s`, `Γ(s)` agrees with Euler's integral + `∫ (x:ℝ) in Ioi 0, exp (-x) * x ^ (s - 1)` +* `real.Gamma_add_one` : for all `s : ℝ` with `s ≠ 0`, we have `Γ(s + 1) = s Γ(s)`. +* `real.Gamma_nat_eq_factorial` : for all `n : ℕ` we have `Γ (n + 1) = n!`. +* `real.differentiable_at_Gamma` : `Γ` is real-differentiable at all `s : ℝ` with + `s ∉ {-n : n ∈ ℕ}`. +* `real.Gamma_ne_zero`: for all `s : ℝ` with `s ∉ {-n : n ∈ ℕ}` we have `Γ s ≠ 0`. +* `real.tendsto_log_Gamma`: for all `0 < s`, the limit as `n → ∞` of the sequence + `n ↦ s log n + log n! - (log s + ... + log (s + n))` is `log Γ(s)`. +* `real.convex_on_log_Gamma` : `log ∘ Γ` is convex on `Ioi 0`. +* `real.eq_Gamma_of_log_convex` : the Bohr-Mollerup theorem, which states that the `Γ` function is + the unique log-convex, positive-valued function on `Ioi 0` satisfying the functional equation + and having `Γ 1 = 1`. + +## Main statements (complex case) + +* `complex.Gamma` : the `Γ` function (of a complex variable). +* `complex.Gamma_eq_integral` : for `0 < re s`, `Γ(s)` agrees with Euler's integral. +* `complex.Gamma_add_one` : for all `s : ℂ` with `s ≠ 0`, we have `Γ(s + 1) = s Γ(s)`. +* `complex.Gamma_nat_eq_factorial` : for all `n : ℕ` we have `Γ (n + 1) = n!`. +* `complex.differentiable_at_Gamma` : `Γ` is complex-differentiable at all `s : ℂ` with + `s ∉ {-n : n ∈ ℕ}`. ## Tags @@ -26,7 +52,7 @@ Gamma noncomputable theory open filter interval_integral set real measure_theory asymptotics -open_locale topological_space +open_locale nat topological_space ennreal big_operators lemma integral_exp_neg_Ioi : ∫ (x : ℝ) in Ioi 0, exp (-x) = 1 := begin @@ -52,13 +78,7 @@ begin exact (tendsto_exp_mul_div_rpow_at_top s (1 / 2) one_half_pos).inv_tendsto_at_top, end -/-- Euler's integral for the `Γ` function (of a real variable `s`), defined as -`∫ x in Ioi 0, exp (-x) * x ^ (s - 1)`. - -See `Gamma_integral_convergent` for a proof of the convergence of the integral for `0 < s`. -/ -def Gamma_integral (s : ℝ) : ℝ := ∫ x in Ioi (0:ℝ), exp (-x) * x ^ (s - 1) - -/-- The integral defining the `Γ` function converges for positive real `s`. -/ +/-- The Euler integral for the `Γ` function converges for positive real `s`. -/ lemma Gamma_integral_convergent {s : ℝ} (h : 0 < s) : integrable_on (λ x:ℝ, exp (-x) * x ^ (s - 1)) (Ioi 0) := begin @@ -74,9 +94,6 @@ begin exact or.inl ((zero_lt_one : (0 : ℝ) < 1).trans_le hx).ne' } end -lemma Gamma_integral_one : Gamma_integral 1 = 1 := -by simpa only [Gamma_integral, sub_self, rpow_zero, mul_one] using integral_exp_neg_Ioi - end real namespace complex @@ -116,9 +133,9 @@ See `complex.Gamma_integral_convergent` for a proof of the convergence of the in def Gamma_integral (s : ℂ) : ℂ := ∫ x in Ioi (0:ℝ), ↑(-x).exp * ↑x ^ (s - 1) lemma Gamma_integral_of_real (s : ℝ) : - Gamma_integral ↑s = ↑(s.Gamma_integral) := + Gamma_integral ↑s = ↑(∫ x:ℝ in Ioi 0, real.exp (-x) * x ^ (s - 1)) := begin - rw [real.Gamma_integral, ←_root_.integral_of_real], + rw [Gamma_integral, ←_root_.integral_of_real], refine set_integral_congr measurable_set_Ioi _, intros x hx, dsimp only, rw [of_real_mul, of_real_cpow (mem_Ioi.mp hx).le], @@ -126,10 +143,8 @@ begin end lemma Gamma_integral_one : Gamma_integral 1 = 1 := -begin - rw [←of_real_one, Gamma_integral_of_real, of_real_inj], - exact real.Gamma_integral_one, -end +by simpa only [←of_real_one, Gamma_integral_of_real, of_real_inj, sub_self, + rpow_zero, mul_one] using integral_exp_neg_Ioi end complex @@ -327,16 +342,18 @@ begin field_simp, ring, end -theorem Gamma_eq_integral (s : ℂ) (hs : 0 < s.re) : Gamma s = Gamma_integral s := +theorem Gamma_eq_integral {s : ℂ} (hs : 0 < s.re) : Gamma s = Gamma_integral s := Gamma_eq_Gamma_aux s 0 (by { norm_cast, linarith }) -theorem Gamma_nat_eq_factorial (n : ℕ) : Gamma (n+1) = nat.factorial n := +lemma Gamma_one : Gamma 1 = 1 := +by { rw Gamma_eq_integral, simpa using Gamma_integral_one, simp } + +theorem Gamma_nat_eq_factorial (n : ℕ) : Gamma (n+1) = n! := begin induction n with n hn, - { rw [nat.cast_zero, zero_add], rw Gamma_eq_integral, - simpa using Gamma_integral_one, simp,}, - rw (Gamma_add_one n.succ $ nat.cast_ne_zero.mpr $ nat.succ_ne_zero n), - { simp only [nat.cast_succ, nat.factorial_succ, nat.cast_mul], congr, exact hn }, + { simpa using Gamma_one }, + { rw (Gamma_add_one n.succ $ nat.cast_ne_zero.mpr $ nat.succ_ne_zero n), + simp only [nat.cast_succ, nat.factorial_succ, nat.cast_mul], congr, exact hn }, end end Gamma_def @@ -519,3 +536,396 @@ end end complex end Gamma_has_deriv + +namespace real + +/-- The `Γ` function (of a real variable `s`). -/ +@[pp_nodot] def Gamma (s : ℝ) : ℝ := (complex.Gamma s).re + +lemma Gamma_eq_integral {s : ℝ} (hs : 0 < s) : Gamma s = ∫ x in Ioi 0, exp (-x) * x ^ (s - 1) := +begin + rw [Gamma, complex.Gamma_eq_integral (by rwa complex.of_real_re : 0 < complex.re s)], + dsimp only [complex.Gamma_integral], + simp_rw [←complex.of_real_one, ←complex.of_real_sub], + suffices : ∫ (x : ℝ) in Ioi 0, ↑(exp (-x)) * (x : ℂ) ^ ((s - 1 : ℝ) : ℂ) = + ∫ (x : ℝ) in Ioi 0, ((exp (-x) * x ^ (s - 1) : ℝ) : ℂ), + { rw [this, _root_.integral_of_real, complex.of_real_re], }, + refine set_integral_congr measurable_set_Ioi (λ x hx, _), + push_cast, + rw complex.of_real_cpow (le_of_lt hx), + push_cast, +end + +lemma Gamma_add_one {s : ℝ} (hs : s ≠ 0) : Gamma (s + 1) = s * Gamma s := +begin + simp_rw Gamma, + rw [complex.of_real_add, complex.of_real_one, complex.Gamma_add_one, complex.of_real_mul_re], + rwa complex.of_real_ne_zero, +end + +lemma Gamma_one : Gamma 1 = 1 := +by rw [Gamma, complex.of_real_one, complex.Gamma_one, complex.one_re] + +theorem Gamma_nat_eq_factorial (n : ℕ) : Gamma (n + 1) = n! := +by rw [Gamma, complex.of_real_add, complex.of_real_nat_cast, complex.of_real_one, + complex.Gamma_nat_eq_factorial, ←complex.of_real_nat_cast, complex.of_real_re] + +lemma Gamma_pos_of_pos {s : ℝ} (hs : 0 < s) : 0 < Gamma s := +begin + rw Gamma_eq_integral hs, + have : function.support (λ (x : ℝ), exp (-x) * x ^ (s - 1)) ∩ Ioi 0 = Ioi 0, + { rw inter_eq_right_iff_subset, + intros x hx, + rw function.mem_support, + exact mul_ne_zero (exp_pos _).ne' (rpow_pos_of_pos hx _).ne' }, + rw set_integral_pos_iff_support_of_nonneg_ae, + { rw [this, volume_Ioi, ←ennreal.of_real_zero], + exact ennreal.of_real_lt_top }, + { refine eventually_of_mem (self_mem_ae_restrict measurable_set_Ioi) _, + exact λ x hx, (mul_pos (exp_pos _) (rpow_pos_of_pos hx _)).le }, + { exact Gamma_integral_convergent hs }, +end + +lemma Gamma_ne_zero {s : ℝ} (hs : ∀ m:ℕ, s + m ≠ 0) : Gamma s ≠ 0 := +begin + suffices : ∀ {n : ℕ}, (-(n:ℝ) < s) → Gamma s ≠ 0, + { apply this, + swap, use (⌊-s⌋₊ + 1), + rw [neg_lt, nat.cast_add, nat.cast_one], + exact nat.lt_floor_add_one _ }, + intro n, + induction n generalizing s, + { intro hs, + refine (Gamma_pos_of_pos _).ne', + rwa [nat.cast_zero, neg_zero] at hs }, + { intro hs', + have : Gamma (s + 1) ≠ 0, + { apply n_ih, + { intro m, + convert hs (1 + m) using 1, + push_cast, + ring, }, + { rw [nat.succ_eq_add_one, nat.cast_add, nat.cast_one, neg_add] at hs', + linarith } }, + rw [Gamma_add_one, mul_ne_zero_iff] at this, + { exact this.2 }, + { simpa using hs 0 } }, +end + +lemma differentiable_at_Gamma {s : ℝ} (hs : ∀ m:ℕ, s + m ≠ 0) : differentiable_at ℝ Gamma s := +begin + apply has_deriv_at.differentiable_at, + apply has_deriv_at.real_of_complex, + apply differentiable_at.has_deriv_at, + apply complex.differentiable_at_Gamma, + simp_rw [←complex.of_real_nat_cast, ←complex.of_real_add, complex.of_real_ne_zero], + exact hs, +end + +/-- Log-convexity of the Gamma function on the positive reals (stated in multiplicative form), +proved using the Hölder inequality applied to Euler's integral. -/ +lemma Gamma_mul_add_mul_le_rpow_Gamma_mul_rpow_Gamma {s t a b : ℝ} + (hs : 0 < s) (ht : 0 < t) (ha : 0 < a) (hb : 0 < b) (hab : a + b = 1) : + Gamma (a * s + b * t) ≤ Gamma s ^ a * Gamma t ^ b := +begin + -- We will apply Hölder's inequality, for the conjugate exponents `p = 1 / a` + -- and `q = 1 / b`, to the functions `f a s` and `f b t`, where `f` is as follows: + let f : ℝ → ℝ → ℝ → ℝ := λ c u x, exp (-c * x) * x ^ (c * (u - 1)), + have e : is_conjugate_exponent (1 / a) (1 / b) := real.is_conjugate_exponent_one_div ha hb hab, + have hab' : b = 1 - a := by linarith, + have hst : 0 < a * s + b * t := add_pos (mul_pos ha hs) (mul_pos hb ht), + -- some properties of f: + have posf : ∀ (c u x : ℝ), x ∈ Ioi (0:ℝ) → 0 ≤ f c u x := + λ c u x hx, mul_nonneg (exp_pos _).le (rpow_pos_of_pos hx _).le, + have posf' : ∀ (c u : ℝ), ∀ᵐ (x : ℝ) ∂volume.restrict (Ioi 0), 0 ≤ f c u x := + λ c u, (ae_restrict_iff' measurable_set_Ioi).mpr (ae_of_all _ (posf c u)), + have fpow : ∀ {c x : ℝ} (hc : 0 < c) (u : ℝ) (hx : 0 < x), + exp (-x) * x ^ (u - 1) = f c u x ^ (1 / c), + { intros c x hc u hx, + dsimp only [f], + rw [mul_rpow (exp_pos _).le ((rpow_nonneg_of_nonneg hx.le) _), ←exp_mul, ←rpow_mul hx.le], + congr' 2; + { field_simp [hc.ne'], ring } }, + -- show `f c u` is in `ℒp` for `p = 1/c`: + have f_mem_Lp : ∀ {c u : ℝ} (hc : 0 < c) (hu : 0 < u), + mem_ℒp (f c u) (ennreal.of_real (1 / c)) (volume.restrict (Ioi 0)), + { intros c u hc hu, + have A : ennreal.of_real (1 / c) ≠ 0, + by rwa [ne.def, ennreal.of_real_eq_zero, not_le, one_div_pos], + have B : ennreal.of_real (1 / c) ≠ ∞, from ennreal.of_real_ne_top, + rw [←mem_ℒp_norm_rpow_iff _ A B, ennreal.to_real_of_real (one_div_nonneg.mpr hc.le), + ennreal.div_self A B, mem_ℒp_one_iff_integrable], + { apply integrable.congr (Gamma_integral_convergent hu), + refine eventually_eq_of_mem (self_mem_ae_restrict measurable_set_Ioi) (λ x hx, _), + dsimp only, + rw fpow hc u hx, + congr' 1, + exact (norm_of_nonneg (posf _ _ x hx)).symm }, + { refine continuous_on.ae_strongly_measurable _ measurable_set_Ioi, + refine (continuous.continuous_on _).mul (continuous_at.continuous_on (λ x hx, _)), + { exact continuous_exp.comp (continuous_const.mul continuous_id'), }, + { exact continuous_at_rpow_const _ _ (or.inl (ne_of_lt hx).symm), } } }, + -- now apply Hölder: + rw [Gamma_eq_integral hs, Gamma_eq_integral ht, Gamma_eq_integral hst], + convert measure_theory.integral_mul_le_Lp_mul_Lq_of_nonneg e (posf' a s) (posf' b t) + (f_mem_Lp ha hs) (f_mem_Lp hb ht) using 1, + { refine set_integral_congr measurable_set_Ioi (λ x hx, _), + dsimp only [f], + have A : exp (-x) = exp (-a * x) * exp (-b * x), + { rw [←exp_add, ←add_mul, ←neg_add, hab, neg_one_mul] }, + have B : x ^ (a * s + b * t - 1) = (x ^ (a * (s - 1))) * (x ^ (b * (t - 1))), + { rw [←rpow_add hx, hab'], congr' 1, ring }, + rw [A, B], + ring }, + { rw [one_div_one_div, one_div_one_div], + congr' 2; + exact set_integral_congr measurable_set_Ioi (λ x hx, fpow (by assumption) _ hx) }, +end + +lemma convex_on_log_Gamma : convex_on ℝ (Ioi 0) (log ∘ Gamma) := +begin + refine convex_on_iff_forall_pos.mpr ⟨convex_Ioi _, λ x hx y hy a b ha hb hab, _⟩, + have : b = 1 - a := by linarith, subst this, + simp_rw [function.comp_app, smul_eq_mul], + rw [←log_rpow (Gamma_pos_of_pos hy), ←log_rpow (Gamma_pos_of_pos hx), + ←log_mul + ((rpow_pos_of_pos (Gamma_pos_of_pos hx) _).ne') (rpow_pos_of_pos (Gamma_pos_of_pos hy) _).ne', + log_le_log + (Gamma_pos_of_pos (add_pos (mul_pos ha hx) (mul_pos hb hy))) + (mul_pos + (rpow_pos_of_pos (Gamma_pos_of_pos hx) _) (rpow_pos_of_pos (Gamma_pos_of_pos hy) _))], + exact Gamma_mul_add_mul_le_rpow_Gamma_mul_rpow_Gamma hx hy ha hb hab, +end + +section bohr_mollerup + +/-! ## The Euler limit formula and the Bohr-Mollerup theorem + +In this section we prove two interrelated statements about the `Γ` function on the positive reals: + +* the Euler limit formula `real.tendsto_log_gamma_seq`, stating that the sequence + `x * log n + log n! - ∑ (m : ℕ) in finset.range (n + 1), log (x + m)` + tends to `log Γ(x)` as `n → ∞`. +* the Bohr-Mollerup theorem (`real.eq_Gamma_of_log_convex`) which states that `Γ` is the unique + *log-convex*, positive-real-valued function on the positive reals satisfying + `f (x + 1) = x f x` and `f 1 = 1`. + +To do this, we prove that any function satisfying the hypotheses of the Bohr--Mollerup theorem must +agree with the limit in the Gauss formula, so there is at most one such function. Then we show that +`Γ` satisfies these conditions. +-/ + +/-- The function `n ↦ x log n + log n! - (log x + ... + log (x + n))`, which we will show tends to +`log (Gamma x)` as `n → ∞`. -/ +def log_gamma_seq (x : ℝ) (n : ℕ) : ℝ := +x * log n + log n! - ∑ (m : ℕ) in finset.range (n + 1), log (x + m) + +/-! The following are auxiliary lemmas for the Bohr-Mollerup theorem, which are +placed in a separate namespace `bohr_mollerup` to avoid clutter. -/ +namespace bohr_mollerup + +variables {f : ℝ → ℝ} {x : ℝ} {n : ℕ} + +lemma f_nat_eq (hf_feq : ∀ {y:ℝ}, 0 < y → f (y + 1) = f y + log y) (hn : n ≠ 0) : + f n = f 1 + log (n - 1)! := +begin + refine nat.le_induction (by simp) (λ m hm IH, _) n (nat.one_le_iff_ne_zero.2 hn), + have A : 0 < (m : ℝ), from nat.cast_pos.2 hm, + simp only [hf_feq A, nat.cast_add, algebra_map.coe_one, nat.add_succ_sub_one, add_zero], + rw [IH, add_assoc, ← log_mul (nat.cast_ne_zero.mpr (nat.factorial_ne_zero _)) A.ne', + ← nat.cast_mul], + conv_rhs { rw [← nat.succ_pred_eq_of_pos hm, nat.factorial_succ, mul_comm] }, + congr, + exact (nat.succ_pred_eq_of_pos hm).symm +end + +lemma f_add_nat_eq (hf_feq : ∀ {y:ℝ}, 0 < y → f (y + 1) = f y + log y) (hx : 0 < x) (n : ℕ) : + f (x + n) = f x + ∑ (m : ℕ) in finset.range n, log (x + m) := +begin + induction n with n hn, + { simp }, + { have : x + n.succ = (x + n) + 1, + { push_cast, ring }, + rw [this, hf_feq, hn], + rw [finset.range_succ, finset.sum_insert (finset.not_mem_range_self)], + abel, + linarith [(nat.cast_nonneg n : 0 ≤ (n:ℝ))] }, +end + +/-- Linear upper bound for `f (x + n)` on unit interval -/ +lemma f_add_nat_le + (hf_conv : convex_on ℝ (Ioi 0) f) (hf_feq : ∀ {y:ℝ}, 0 < y → f (y + 1) = f y + log y) + (hn : n ≠ 0) (hx : 0 < x) (hx' : x ≤ 1) : + f (n + x) ≤ f n + x * log n := +begin + have hn': 0 < (n:ℝ) := nat.cast_pos.mpr (nat.pos_of_ne_zero hn), + have : f n + x * log n = (1 - x) * f n + x * f (n + 1), + { rw [hf_feq hn'], ring, }, + rw [this, (by ring : (n:ℝ) + x = (1 - x) * n + x * (n + 1))], + simpa only [smul_eq_mul] using hf_conv.2 hn' (by linarith : 0 < (n + 1 : ℝ)) + (by linarith : 0 ≤ 1 - x) hx.le (by linarith), +end + +/-- Linear lower bound for `f (x + n)` on unit interval -/ +lemma f_add_nat_ge + (hf_conv : convex_on ℝ (Ioi 0) f) (hf_feq : ∀ {y:ℝ}, 0 < y → f (y + 1) = f y + log y) + (hn : 2 ≤ n) (hx : 0 < x) : + f n + x * log (n - 1) ≤ f (n + x) := +begin + have npos : 0 < (n:ℝ) - 1, + { rw [←nat.cast_one, sub_pos, nat.cast_lt], linarith, }, + have c := (convex_on_iff_slope_mono_adjacent.mp $ hf_conv).2 + npos (by linarith : 0 < (n:ℝ) + x) (by linarith : (n:ℝ) - 1 < (n:ℝ)) (by linarith), + rw [add_sub_cancel', sub_sub_cancel, div_one] at c, + have : f (↑n - 1) = f n - log (↑n - 1), + { nth_rewrite_rhs 0 (by ring : (n:ℝ) = (↑n - 1) + 1), + rw [hf_feq npos, add_sub_cancel] }, + rwa [this, le_div_iff hx, sub_sub_cancel, le_sub_iff_add_le, mul_comm _ x, add_comm] at c, +end + +lemma log_gamma_seq_add_one (x : ℝ) (n : ℕ) : + log_gamma_seq (x + 1) n = log_gamma_seq x (n + 1) + log x - (x + 1) * (log (n + 1) - log n) := +begin + dsimp only [nat.factorial_succ, log_gamma_seq], + conv_rhs { rw [finset.sum_range_succ', nat.cast_zero, add_zero], }, + rw [nat.cast_mul, log_mul], rotate, + { rw nat.cast_ne_zero, exact nat.succ_ne_zero n }, + { rw nat.cast_ne_zero, exact nat.factorial_ne_zero n, }, + have : ∑ (m : ℕ) in finset.range (n + 1), log (x + 1 + ↑m) = + ∑ (k : ℕ) in finset.range (n + 1), log (x + ↑(k + 1)), + { refine finset.sum_congr (by refl) (λ m hm, _), + congr' 1, + push_cast, + abel }, + rw [←this, nat.cast_add_one n], + ring, +end + +lemma le_log_gamma_seq + (hf_conv : convex_on ℝ (Ioi 0) f) (hf_feq : ∀ {y:ℝ}, 0 < y → f (y + 1) = f y + log y) + (hx : 0 < x) (hx' : x ≤ 1) (n : ℕ) : + f x ≤ f 1 + x * log (n + 1) - x * log n + log_gamma_seq x n := +begin + rw [log_gamma_seq, ←add_sub_assoc, le_sub_iff_add_le, ←f_add_nat_eq @hf_feq hx, add_comm x], + refine (f_add_nat_le hf_conv @hf_feq (nat.add_one_ne_zero n) hx hx').trans (le_of_eq _), + rw [f_nat_eq @hf_feq (by linarith : n + 1 ≠ 0), nat.add_sub_cancel, nat.cast_add_one], + ring, +end + +lemma ge_log_gamma_seq + (hf_conv : convex_on ℝ (Ioi 0) f) (hf_feq : ∀ {y:ℝ}, 0 < y → f (y + 1) = f y + log y) + (hx : 0 < x) (hn : n ≠ 0) : + f 1 + log_gamma_seq x n ≤ f x := +begin + dsimp [log_gamma_seq], + rw [←add_sub_assoc, sub_le_iff_le_add, ←f_add_nat_eq @hf_feq hx, add_comm x _], + refine le_trans (le_of_eq _) (f_add_nat_ge hf_conv @hf_feq _ hx), + { rw [f_nat_eq @hf_feq, nat.add_sub_cancel, nat.cast_add_one, add_sub_cancel], + { ring }, + { exact nat.succ_ne_zero _} }, + { apply nat.succ_le_succ, + linarith [nat.pos_of_ne_zero hn] }, +end + +lemma tendsto_log_gamma_seq_of_le_one + (hf_conv : convex_on ℝ (Ioi 0) f) (hf_feq : ∀ {y:ℝ}, 0 < y → f (y + 1) = f y + log y) + (hx : 0 < x) (hx' : x ≤ 1) : + tendsto (log_gamma_seq x) at_top (𝓝 $ f x - f 1) := +begin + refine tendsto_of_tendsto_of_tendsto_of_le_of_le' _ tendsto_const_nhds _ _, + show ∀ᶠ (n : ℕ) in at_top, log_gamma_seq x n ≤ f x - f 1, + { refine eventually.mp (eventually_ne_at_top 0) (eventually_of_forall (λ n hn, _)), + exact le_sub_iff_add_le'.mpr (ge_log_gamma_seq hf_conv @hf_feq hx hn) }, + show ∀ᶠ (n : ℕ) in at_top, f x - f 1 - x * (log (n + 1) - log n) ≤ log_gamma_seq x n, + { refine eventually_of_forall (λ n, _), + rw [sub_le_iff_le_add', sub_le_iff_le_add'], + convert le_log_gamma_seq hf_conv @hf_feq hx hx' n using 1, + ring }, + { have : f x - f 1 = (f x - f 1) - x * 0 := by ring, + nth_rewrite 0 this, + exact tendsto.sub tendsto_const_nhds (tendsto_log_nat_add_one_sub_log.const_mul _), } +end + +lemma tendsto_log_gamma_seq + (hf_conv : convex_on ℝ (Ioi 0) f) (hf_feq : ∀ {y:ℝ}, 0 < y → f (y + 1) = f y + log y) + (hx : 0 < x) : + tendsto (log_gamma_seq x) at_top (𝓝 $ f x - f 1) := +begin + suffices : ∀ (m : ℕ), ↑m < x → x ≤ m + 1 → + tendsto (log_gamma_seq x) at_top (𝓝 $ f x - f 1), + { refine this (⌈x - 1⌉₊) _ _, + { rcases lt_or_le x 1, + { rwa [nat.ceil_eq_zero.mpr (by linarith : x - 1 ≤ 0), nat.cast_zero] }, + { convert nat.ceil_lt_add_one (by linarith : 0 ≤ x - 1), + abel } }, + { rw ←sub_le_iff_le_add, exact nat.le_ceil _}, }, + intro m, + induction m with m hm generalizing x, + { rw [nat.cast_zero, zero_add], + exact λ _ hx', tendsto_log_gamma_seq_of_le_one hf_conv @hf_feq hx hx' }, + { intros hy hy', + rw [nat.cast_succ, ←sub_le_iff_le_add] at hy', + rw [nat.cast_succ, ←lt_sub_iff_add_lt] at hy, + specialize hm ((nat.cast_nonneg _).trans_lt hy) hy hy', + -- now massage gauss_product n (x - 1) into gauss_product (n - 1) x + have : ∀ᶠ (n:ℕ) in at_top, log_gamma_seq (x - 1) n = log_gamma_seq x (n - 1) + + x * (log (↑(n - 1) + 1) - log ↑(n - 1)) - log (x - 1), + { refine eventually.mp (eventually_ge_at_top 1) (eventually_of_forall (λ n hn, _)), + have := log_gamma_seq_add_one (x - 1) (n - 1), + rw [sub_add_cancel, nat.sub_add_cancel hn] at this, + rw this, + ring }, + replace hm := ((tendsto.congr' this hm).add + (tendsto_const_nhds : tendsto (λ _, log (x - 1)) _ _)).comp (tendsto_add_at_top_nat 1), + have : + (λ (x_1 : ℕ), (λ (n : ℕ), log_gamma_seq x (n - 1) + + x * (log (↑(n - 1) + 1) - log ↑(n - 1)) - log (x - 1)) x_1 + + (λ (b : ℕ), log (x - 1)) x_1) ∘ (λ (a : ℕ), a + 1) = + λ n, log_gamma_seq x n + x * (log (↑n + 1) - log ↑n), + { ext1 n, + dsimp only [function.comp_app], + rw [sub_add_cancel, nat.add_sub_cancel] }, + rw this at hm, + convert hm.sub (tendsto_log_nat_add_one_sub_log.const_mul x) using 2, + { ext1 n, ring }, + { have := hf_feq ((nat.cast_nonneg m).trans_lt hy), + rw sub_add_cancel at this, + rw this, + ring } }, +end + +end bohr_mollerup + +lemma tendsto_log_Gamma {x : ℝ} (hx : 0 < x) : + tendsto (log_gamma_seq x) at_top (𝓝 $ log (Gamma x)) := +begin + have : log (Gamma x) = (log ∘ Gamma) x - (log ∘ Gamma) 1, + { simp_rw [function.comp_app, Gamma_one, log_one, sub_zero] }, + rw this, + refine bohr_mollerup.tendsto_log_gamma_seq convex_on_log_Gamma (λ y hy, _) hx, + rw [function.comp_app, Gamma_add_one hy.ne', log_mul hy.ne' (Gamma_pos_of_pos hy).ne', add_comm], +end + +/-- The **Bohr-Mollerup theorem**: the Gamma function is the *unique* log-convex, positive-valued +function on the positive reals which satisfies `f 1 = 1` and `f (x + 1) = x * f x` for all `x`. -/ +lemma eq_Gamma_of_log_convex {f : ℝ → ℝ} + (hf_conv : convex_on ℝ (Ioi 0) (log ∘ f)) + (hf_feq : ∀ {y:ℝ}, 0 < y → f (y + 1) = y * f y) + (hf_pos : ∀ {y:ℝ}, 0 < y → 0 < f y) + (hf_one : f 1 = 1) : + eq_on f Gamma (Ioi (0:ℝ)) := +begin + suffices : eq_on (log ∘ f) (log ∘ Gamma) (Ioi (0:ℝ)), + from λ x hx, log_inj_on_pos (hf_pos hx) (Gamma_pos_of_pos hx) (this hx), + intros x hx, + have e1 := bohr_mollerup.tendsto_log_gamma_seq hf_conv _ hx, + { rw [function.comp_app log f 1, hf_one, log_one, sub_zero] at e1, + exact tendsto_nhds_unique e1 (tendsto_log_Gamma hx) }, + { intros y hy, + rw [function.comp_app, hf_feq hy, log_mul hy.ne' (hf_pos hy).ne'], + ring } +end + +end bohr_mollerup + +end real diff --git a/src/analysis/special_functions/gaussian.lean b/src/analysis/special_functions/gaussian.lean index ef4792ec70e2d..8053436bc7b1a 100644 --- a/src/analysis/special_functions/gaussian.lean +++ b/src/analysis/special_functions/gaussian.lean @@ -320,7 +320,7 @@ begin have hh2 : (1 / 2 : ℂ).re = 1 / 2, { convert of_real_re (1 / 2 : ℝ) }, replace hh2 : 0 < (1 / 2 : ℂ).re := by { rw hh2, exact one_half_pos, }, - rw [Gamma_eq_integral _ hh2, hh, Gamma_integral_of_real, of_real_inj, real.Gamma_integral], + rw [Gamma_eq_integral hh2, hh, Gamma_integral_of_real, of_real_inj], -- now do change-of-variables rw ←integral_comp_rpow_Ioi_of_pos zero_lt_two, have : eq_on (λ x:ℝ, (2 * x^((2:ℝ) - 1)) • (real.exp (-x^(2:ℝ)) * (x^(2:ℝ)) ^ (1 / (2:ℝ) - 1))) diff --git a/src/analysis/special_functions/integrals.lean b/src/analysis/special_functions/integrals.lean index efd34ff0461cc..cedeb1c491154 100644 --- a/src/analysis/special_functions/integrals.lean +++ b/src/analysis/special_functions/integrals.lean @@ -400,6 +400,60 @@ end lemma integral_one_div_one_add_sq : ∫ x : ℝ in a..b, 1 / (1 + x^2) = arctan b - arctan a := by simp only [one_div, integral_inv_one_add_sq] +section rpow_cpow +open complex + +lemma integral_mul_cpow_one_add_sq {t : ℂ} (ht : t ≠ -1) : + ∫ x : ℝ in a..b, (x:ℂ) * (1 + x ^ 2) ^ t = + (1 + b ^ 2) ^ (t + 1) / (2 * (t + 1)) - (1 + a ^ 2) ^ (t + 1) / (2 * (t + 1)) := +begin + have : t + 1 ≠ 0 := by { contrapose! ht, rwa add_eq_zero_iff_eq_neg at ht }, + apply integral_eq_sub_of_has_deriv_at, + { intros x hx, + have f : has_deriv_at (λ y:ℂ, 1 + y ^ 2) (2 * x) x, + { convert (has_deriv_at_pow 2 (x:ℂ)).const_add 1, { norm_cast }, { simp } }, + have g : ∀ {z : ℂ}, (0 < z.re) → has_deriv_at (λ z, z ^ (t + 1) / (2 * (t + 1))) (z ^ t / 2) z, + { intros z hz, + have : z ≠ 0 := by { contrapose! hz, rw [hz, zero_re], }, + convert (has_deriv_at.cpow_const (has_deriv_at_id _) (or.inl hz)).div_const + (2 * (t + 1)) using 1, + field_simp, + ring }, + convert (has_deriv_at.comp ↑x (g _) f).comp_of_real using 1, + { field_simp, ring }, + { rw [add_re, one_re, ←of_real_pow, of_real_re], + exact (add_pos_of_pos_of_nonneg zero_lt_one (sq_nonneg x)) } }, + { apply continuous.interval_integrable, + refine continuous_of_real.mul _, + apply continuous.cpow, + { exact continuous_const.add (continuous_of_real.pow 2)}, + { exact continuous_const }, + { intro a, + rw [add_re, one_re, ←of_real_pow, of_real_re], + exact or.inl (add_pos_of_pos_of_nonneg zero_lt_one (sq_nonneg a)) } } +end + +lemma integral_mul_rpow_one_add_sq {t : ℝ} (ht : t ≠ -1) : + ∫ x : ℝ in a..b, x * (1 + x ^ 2) ^ t = + (1 + b ^ 2) ^ (t + 1) / (2 * (t + 1)) - (1 + a ^ 2) ^ (t + 1) / (2 * (t + 1)) := +begin + have : ∀ (x s : ℝ), (((1 + x ^ 2) ^ s : ℝ) : ℂ) = (1 + (x : ℂ) ^ 2) ^ ↑s, + { intros x s, + rw [of_real_cpow, of_real_add, of_real_pow, of_real_one], + exact add_nonneg zero_le_one (sq_nonneg x), }, + rw ←of_real_inj, + convert integral_mul_cpow_one_add_sq (_ : (t:ℂ) ≠ -1), + { rw ←interval_integral.integral_of_real, + congr' with x:1, + rw [of_real_mul, this x t] }, + { simp_rw [of_real_sub, of_real_div, this a (t + 1), this b (t + 1)], + push_cast }, + { rw [←of_real_one, ←of_real_neg, ne.def, of_real_inj], + exact ht }, +end + +end rpow_cpow + /-! ### Integral of `sin x ^ n` -/ lemma integral_sin_pow_aux : diff --git a/src/analysis/special_functions/log/basic.lean b/src/analysis/special_functions/log/basic.lean index a4ffa13124f6f..6667ebe175a9b 100644 --- a/src/analysis/special_functions/log/basic.lean +++ b/src/analysis/special_functions/log/basic.lean @@ -325,3 +325,30 @@ lemma continuous_on.log (hf : continuous_on f s) (h₀ : ∀ x ∈ s, f x ≠ 0) λ x hx, (hf x hx).log (h₀ x hx) end continuity + + +section tendsto_comp_add_sub + +open filter +namespace real + +lemma tendsto_log_comp_add_sub_log (y : ℝ) : + tendsto (λ x:ℝ, log (x + y) - log x) at_top (𝓝 0) := +begin + refine tendsto.congr' (_ : ∀ᶠ (x : ℝ) in at_top, log (1 + y / x) = _) _, + { refine eventually.mp ((eventually_ne_at_top 0).and (eventually_gt_at_top (-y))) + (eventually_of_forall (λ x hx, _)), + rw ← log_div _ hx.1, + { congr' 1, + field_simp [hx.1] }, + { linarith [hx.2] } }, + { suffices : tendsto (λ (x : ℝ), log (1 + y / x)) at_top (𝓝 (log (1 + 0))), by simpa, + refine tendsto.log _ (by simp), + exact tendsto_const_nhds.add (tendsto_const_nhds.div_at_top tendsto_id) }, +end + +lemma tendsto_log_nat_add_one_sub_log : tendsto (λ (k : ℕ), log (k + 1) - log k) at_top (𝓝 0) := +(tendsto_log_comp_add_sub_log 1).comp tendsto_coe_nat_at_top_at_top + +end real +end tendsto_comp_add_sub diff --git a/src/analysis/special_functions/pow.lean b/src/analysis/special_functions/pow.lean index 4fd66dc7d3314..ddcfb486c74dd 100644 --- a/src/analysis/special_functions/pow.lean +++ b/src/analysis/special_functions/pow.lean @@ -21,6 +21,7 @@ We also prove basic properties of these functions. noncomputable theory open_locale classical real topological_space nnreal ennreal filter big_operators asymptotics +open_locale complex_conjugate open filter finset set namespace complex @@ -493,6 +494,42 @@ begin { exact abs_cpow_eq_rpow_re_of_pos hlt y } end +lemma inv_cpow_eq_ite (x : ℂ) (n : ℂ) : + x⁻¹ ^ n = if x.arg = π then conj (x ^ conj n)⁻¹ else (x ^ n)⁻¹ := +begin + simp_rw [complex.cpow_def, log_inv_eq_ite, inv_eq_zero, map_eq_zero, ite_mul, neg_mul, + is_R_or_C.conj_inv, apply_ite conj, apply_ite exp, apply_ite has_inv.inv, map_zero, map_one, + exp_neg, inv_one, inv_zero, ←exp_conj, map_mul, conj_conj], + split_ifs with hx hn ha ha; refl, +end + +lemma inv_cpow (x : ℂ) (n : ℂ) (hx : x.arg ≠ π) : x⁻¹ ^ n = (x ^ n)⁻¹ := +by rw [inv_cpow_eq_ite, if_neg hx] + +/-- `complex.inv_cpow_eq_ite` with the `ite` on the other side. -/ +lemma inv_cpow_eq_ite' (x : ℂ) (n : ℂ) : + (x ^ n)⁻¹ = if x.arg = π then conj (x⁻¹ ^ conj n) else x⁻¹ ^ n := +begin + rw [inv_cpow_eq_ite, apply_ite conj, conj_conj, conj_conj], + split_ifs, + { refl }, + { rw inv_cpow _ _ h } +end + +lemma conj_cpow_eq_ite (x : ℂ) (n : ℂ) : + conj x ^ n = if x.arg = π then x ^ n else conj (x ^ conj n) := +begin + simp_rw [cpow_def, map_eq_zero, apply_ite conj, map_one, map_zero, ←exp_conj, map_mul, + conj_conj, log_conj_eq_ite], + split_ifs with hcx hn hx; refl +end + +lemma conj_cpow (x : ℂ) (n : ℂ) (hx : x.arg ≠ π) : conj x ^ n = conj (x ^ conj n) := +by rw [conj_cpow_eq_ite, if_neg hx] + +lemma cpow_conj (x : ℂ) (n : ℂ) (hx : x.arg ≠ π) : x ^ conj n = conj (conj x ^ n) := +by rw [conj_cpow _ _ hx, conj_conj] + end complex namespace real diff --git a/src/combinatorics/set_family/compression/down.lean b/src/combinatorics/set_family/compression/down.lean index 94d553f354236..2445f2eb75c10 100644 --- a/src/combinatorics/set_family/compression/down.lean +++ b/src/combinatorics/set_family/compression/down.lean @@ -8,6 +8,9 @@ import data.finset.card /-! # Down-compressions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines down-compression. Down-compressing `𝒜 : finset (finset α)` along `a : α` means removing `a` from the elements of `𝒜`, diff --git a/src/computability/primrec.lean b/src/computability/primrec.lean index 669d748677efa..e6b4673132d2f 100644 --- a/src/computability/primrec.lean +++ b/src/computability/primrec.lean @@ -3,6 +3,7 @@ Copyright (c) 2018 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ +import logic.equiv.array import logic.equiv.list import logic.function.iterate diff --git a/src/control/equiv_functor/instances.lean b/src/control/equiv_functor/instances.lean index f88b4ca3a9b37..793d7e910ae04 100644 --- a/src/control/equiv_functor/instances.lean +++ b/src/control/equiv_functor/instances.lean @@ -9,6 +9,9 @@ import control.equiv_functor /-! # `equiv_functor` instances +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We derive some `equiv_functor` instances, to enable `equiv_rw` to rewrite under these functions. -/ diff --git a/src/data/array/lemmas.lean b/src/data/array/lemmas.lean index 15745bf153213..2f0f22eff241c 100644 --- a/src/data/array/lemmas.lean +++ b/src/data/array/lemmas.lean @@ -270,32 +270,3 @@ read_foreach end map₂ end array - -namespace equiv - -/-- The natural equivalence between length-`n` heterogeneous arrays -and dependent functions from `fin n`. -/ -def d_array_equiv_fin {n : ℕ} (α : fin n → Type*) : d_array n α ≃ (Π i, α i) := -⟨d_array.read, d_array.mk, λ ⟨f⟩, rfl, λ f, rfl⟩ - -/-- The natural equivalence between length-`n` arrays and functions from `fin n`. -/ -def array_equiv_fin (n : ℕ) (α : Type*) : array n α ≃ (fin n → α) := -d_array_equiv_fin _ - -/-- The natural equivalence between length-`n` vectors and length-`n` arrays. -/ -def vector_equiv_array (α : Type*) (n : ℕ) : vector α n ≃ array n α := -(vector_equiv_fin _ _).trans (array_equiv_fin _ _).symm - -end equiv - -namespace array -open function -variable {n : ℕ} - -instance : traversable (array n) := -@equiv.traversable (flip vector n) _ (λ α, equiv.vector_equiv_array α n) _ - -instance : is_lawful_traversable (array n) := -@equiv.is_lawful_traversable (flip vector n) _ (λ α, equiv.vector_equiv_array α n) _ _ - -end array diff --git a/src/data/fin/tuple/basic.lean b/src/data/fin/tuple/basic.lean index b388460072fc0..d2b4cb189fac6 100644 --- a/src/data/fin/tuple/basic.lean +++ b/src/data/fin/tuple/basic.lean @@ -10,6 +10,9 @@ import data.set.intervals.basic /-! # Operation on tuples +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + We interpret maps `Π i : fin n, α i` as `n`-tuples of elements of possibly varying type `α i`, `(α 0, …, α (n-1))`. A particular case is `fin n → α` of elements with all the same type. In this case when `α i` is a constant map, then tuples are isomorphic (but not definitionally equal) diff --git a/src/data/fin_enum.lean b/src/data/fin_enum.lean index 380f6a35cfebf..d8119788da8e0 100644 --- a/src/data/fin_enum.lean +++ b/src/data/fin_enum.lean @@ -8,6 +8,9 @@ import data.fintype.basic import data.list.prod_sigma /-! +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Type class for finitely enumerable types. The property is stronger than `fintype` in that it assigns each element a rank in a finite enumeration. diff --git a/src/data/finset/basic.lean b/src/data/finset/basic.lean index 8e804a6b8e8ae..28c672b401f2e 100644 --- a/src/data/finset/basic.lean +++ b/src/data/finset/basic.lean @@ -463,6 +463,9 @@ theorem not_mem_singleton {a b : α} : a ∉ ({b} : finset α) ↔ a ≠ b := no theorem mem_singleton_self (a : α) : a ∈ ({a} : finset α) := or.inl rfl +@[simp] lemma val_eq_singleton_iff {a : α} {s : finset α} : s.val = {a} ↔ s = {a} := +by { rw ←val_inj, refl } + lemma singleton_injective : injective (singleton : α → finset α) := λ a b h, mem_singleton.1 (h ▸ mem_singleton_self _) @@ -2215,6 +2218,11 @@ lemma coe_to_list (s : finset α) : (s.to_list : multiset α) = s.val := s.val.c @[simp] lemma to_list_to_finset [decidable_eq α] (s : finset α) : s.to_list.to_finset = s := by { ext, simp } +@[simp] lemma to_list_eq_singleton_iff {a : α} {s : finset α} : s.to_list = [a] ↔ s = {a} := +by rw [to_list, to_list_eq_singleton_iff, val_eq_singleton_iff] + +@[simp] lemma to_list_singleton : ∀ a, ({a} : finset α).to_list = [a] := to_list_singleton + lemma exists_list_nodup_eq [decidable_eq α] (s : finset α) : ∃ (l : list α), l.nodup ∧ l.to_finset = s := ⟨s.to_list, s.nodup_to_list, s.to_list_to_finset⟩ diff --git a/src/data/finset/lattice.lean b/src/data/finset/lattice.lean index 321f110d5ab60..0a2be525db223 100644 --- a/src/data/finset/lattice.lean +++ b/src/data/finset/lattice.lean @@ -11,6 +11,9 @@ import order.complete_lattice /-! # Lattice operations on finsets + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ variables {α β γ ι : Type*} diff --git a/src/data/finset/n_ary.lean b/src/data/finset/n_ary.lean index 7b23ce0dd5208..34cef362469ca 100644 --- a/src/data/finset/n_ary.lean +++ b/src/data/finset/n_ary.lean @@ -8,6 +8,9 @@ import data.finset.prod /-! # N-ary images of finsets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines `finset.image₂`, the binary image of finsets. This is the finset version of `set.image2`. This is mostly useful to define pointwise operations. diff --git a/src/data/finset/nat_antidiagonal.lean b/src/data/finset/nat_antidiagonal.lean index f5037e150821b..89b428effedb0 100644 --- a/src/data/finset/nat_antidiagonal.lean +++ b/src/data/finset/nat_antidiagonal.lean @@ -9,6 +9,9 @@ import data.multiset.nat_antidiagonal /-! # Antidiagonals in ℕ × ℕ as finsets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines the antidiagonals of ℕ × ℕ as finsets: the `n`-th antidiagonal is the finset of pairs `(i, j)` such that `i + j = n`. This is useful for polynomial multiplication and more generally for sums going from `0` to `n`. diff --git a/src/data/finset/pairwise.lean b/src/data/finset/pairwise.lean index 846fd4d0ac2df..d71306e29b6c1 100644 --- a/src/data/finset/pairwise.lean +++ b/src/data/finset/pairwise.lean @@ -8,6 +8,9 @@ import data.finset.lattice /-! # Relations holding pairwise on finite sets +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file we prove a few results about the interaction of `set.pairwise_disjoint` and `finset`, as well as the interaction of `list.pairwise disjoint` and the condition of `disjoint` on `list.to_finset`, in `set` form. diff --git a/src/data/finset/powerset.lean b/src/data/finset/powerset.lean index 0dd929a6a0e97..adeeb63e6f1c4 100644 --- a/src/data/finset/powerset.lean +++ b/src/data/finset/powerset.lean @@ -8,6 +8,9 @@ import data.multiset.powerset /-! # The powerset of a finset + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ namespace finset diff --git a/src/data/finset/sigma.lean b/src/data/finset/sigma.lean index 6c463ce72a0fb..e000ed62e556e 100644 --- a/src/data/finset/sigma.lean +++ b/src/data/finset/sigma.lean @@ -9,6 +9,9 @@ import data.set.sigma /-! # Finite sets in a sigma type +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines a few `finset` constructions on `Σ i, α i`. ## Main declarations diff --git a/src/data/finsupp/basic.lean b/src/data/finsupp/basic.lean index ae25e8f65ba2e..ad52601377153 100644 --- a/src/data/finsupp/basic.lean +++ b/src/data/finsupp/basic.lean @@ -45,7 +45,7 @@ This file is a `noncomputable theory` and uses classical logic throughout. noncomputable theory open finset function -open_locale classical big_operators +open_locale big_operators variables {α β γ ι M M' N P G H R S : Type*} @@ -84,12 +84,16 @@ lemma apply_eq_of_mem_graph {a : α} {m : M} {f : α →₀ M} (h : (a, m) ∈ f @[simp] lemma not_mem_graph_snd_zero (a : α) (f : α →₀ M) : (a, (0 : M)) ∉ f.graph := λ h, (mem_graph_iff.1 h).2.irrefl -@[simp] lemma image_fst_graph (f : α →₀ M) : f.graph.image prod.fst = f.support := -by simp only [graph, map_eq_image, image_image, embedding.coe_fn_mk, (∘), image_id'] +@[simp] lemma image_fst_graph [decidable_eq α] (f : α →₀ M) : f.graph.image prod.fst = f.support := +begin + classical, + simp only [graph, map_eq_image, image_image, embedding.coe_fn_mk, (∘), image_id'], +end lemma graph_injective (α M) [has_zero M] : injective (@graph α M _) := begin intros f g h, + classical, have hsup : f.support = g.support, by rw [← image_fst_graph, h, image_fst_graph], refine ext_iff'.2 ⟨hsup, λ x hx, apply_eq_of_mem_graph $ h.symm ▸ _⟩, exact mk_mem_graph _ (hsup ▸ hx) @@ -114,11 +118,15 @@ end { apply nodup_to_list } end⟩ -@[simp] lemma to_alist_keys_to_finset (f : α →₀ M) : f.to_alist.keys.to_finset = f.support := +@[simp] lemma to_alist_keys_to_finset [decidable_eq α] (f : α →₀ M) : + f.to_alist.keys.to_finset = f.support := by { ext, simp [to_alist, alist.mem_keys, alist.keys, list.keys] } @[simp] lemma mem_to_alist {f : α →₀ M} {x : α} : x ∈ f.to_alist ↔ f x ≠ 0 := -by rw [alist.mem_keys, ←list.mem_to_finset, to_alist_keys_to_finset, mem_support_iff] +begin + classical, + rw [alist.mem_keys, ←list.mem_to_finset, to_alist_keys_to_finset, mem_support_iff] +end end graph @@ -135,35 +143,44 @@ open list /-- Converts an association list into a finitely supported function via `alist.lookup`, sending absent keys to zero. -/ -@[simps] def lookup_finsupp (l : alist (λ x : α, M)) : α →₀ M := -{ support := (l.1.filter $ λ x, sigma.snd x ≠ 0).keys.to_finset, - to_fun := λ a, (l.lookup a).get_or_else 0, +def lookup_finsupp (l : alist (λ x : α, M)) : α →₀ M := +{ support := by haveI := classical.dec_eq α; haveI := classical.dec_eq M; exact + (l.1.filter $ λ x, sigma.snd x ≠ 0).keys.to_finset, + to_fun := λ a, by haveI := classical.dec_eq α; exact (l.lookup a).get_or_else 0, mem_support_to_fun := λ a, begin + classical, simp_rw [mem_to_finset, list.mem_keys, list.mem_filter, ←mem_lookup_iff], cases lookup a l; simp end } -alias lookup_finsupp_to_fun ← lookup_finsupp_apply +@[simp] lemma lookup_finsupp_apply [decidable_eq α] (l : alist (λ x : α, M)) (a : α) : + l.lookup_finsupp a = (l.lookup a).get_or_else 0 := +by convert rfl + +@[simp] lemma lookup_finsupp_support [decidable_eq α] [decidable_eq M] (l : alist (λ x : α, M)) : + l.lookup_finsupp.support = (l.1.filter $ λ x, sigma.snd x ≠ 0).keys.to_finset := +by convert rfl -lemma lookup_finsupp_eq_iff_of_ne_zero {l : alist (λ x : α, M)} {a : α} {x : M} (hx : x ≠ 0) : +lemma lookup_finsupp_eq_iff_of_ne_zero [decidable_eq α] + {l : alist (λ x : α, M)} {a : α} {x : M} (hx : x ≠ 0) : l.lookup_finsupp a = x ↔ x ∈ l.lookup a := -by { rw lookup_finsupp_to_fun, cases lookup a l with m; simp [hx.symm] } +by { rw lookup_finsupp_apply, cases lookup a l with m; simp [hx.symm] } -lemma lookup_finsupp_eq_zero_iff {l : alist (λ x : α, M)} {a : α} : +lemma lookup_finsupp_eq_zero_iff [decidable_eq α] {l : alist (λ x : α, M)} {a : α} : l.lookup_finsupp a = 0 ↔ a ∉ l ∨ (0 : M) ∈ l.lookup a := -by { rw [lookup_finsupp_to_fun, ←lookup_eq_none], cases lookup a l with m; simp } +by { rw [lookup_finsupp_apply, ←lookup_eq_none], cases lookup a l with m; simp } @[simp] lemma empty_lookup_finsupp : lookup_finsupp (∅ : alist (λ x : α, M)) = 0 := -by { ext, simp } +by { classical, ext, simp } -@[simp] lemma insert_lookup_finsupp (l : alist (λ x : α, M)) (a : α) (m : M) : +@[simp] lemma insert_lookup_finsupp [decidable_eq α] (l : alist (λ x : α, M)) (a : α) (m : M) : (l.insert a m).lookup_finsupp = l.lookup_finsupp.update a m := by { ext b, by_cases h : b = a; simp [h] } @[simp] lemma singleton_lookup_finsupp (a : α) (m : M) : (singleton a m).lookup_finsupp = finsupp.single a m := -by simp [←alist.insert_empty] +by { classical, simp [←alist.insert_empty] } @[simp] lemma _root_.finsupp.to_alist_lookup_finsupp (f : α →₀ M) : f.to_alist.lookup_finsupp = f := begin @@ -370,7 +387,11 @@ lemma equiv_map_domain_trans' (f : α ≃ β) (g : β ≃ γ) : @[simp] lemma equiv_map_domain_single (f : α ≃ β) (a : α) (b : M) : equiv_map_domain f (single a b) = single (f a) b := -by ext x; simp only [single_apply, equiv.apply_eq_iff_eq_symm_apply, equiv_map_domain_apply]; congr +begin + classical, + ext x, + simp only [single_apply, equiv.apply_eq_iff_eq_symm_apply, equiv_map_domain_apply], +end @[simp] lemma equiv_map_domain_zero {f : α ≃ β} : equiv_map_domain f (0 : α →₀ M) = (0 : β →₀ M) := by ext x; simp only [equiv_map_domain_apply, coe_zero, pi.zero_apply] @@ -537,6 +558,7 @@ lemma map_domain_apply' (S : set α) {f : α → β} (x : α →₀ M) (hS : (x.support : set α) ⊆ S) (hf : set.inj_on f S) {a : α} (ha : a ∈ S) : map_domain f x (f a) = x a := begin + classical, rw [map_domain, sum_apply, sum], simp_rw single_apply, by_cases hax : a ∈ x.support, @@ -646,6 +668,7 @@ lemma map_domain_inj_on (S : set α) {f : α → β} begin intros v₁ hv₁ v₂ hv₂ eq, ext a, + classical, by_cases h : a ∈ v₁.support ∪ v₂.support, { rw [← map_domain_apply' S _ hv₁ hf _, ← map_domain_apply' S _ hv₂ hf _, eq]; { apply set.union_subset hv₁ hv₂, @@ -796,7 +819,7 @@ by { ext, simp, } @[simp] lemma some_single_some [has_zero M] (a : α) (m : M) : (single (option.some a) m : option α →₀ M).some = single a m := -by { ext b, simp [single_apply], } +by { classical, ext b, simp [single_apply], } @[to_additive] lemma prod_option_index [add_comm_monoid M] [comm_monoid N] @@ -831,8 +854,8 @@ variables [has_zero M] (p : α → Prop) (f : α →₀ M) /-- `filter p f` is the finitely supported function that is `f a` if `p a` is true and 0 otherwise. -/ def filter (p : α → Prop) (f : α →₀ M) : α →₀ M := -{ to_fun := λ a, if p a then f a else 0, - support := f.support.filter (λ a, p a), +{ to_fun := λ a, by haveI := classical.dec_pred p; exact if p a then f a else 0, + support := by haveI := classical.dec_pred p; exact f.support.filter (λ a, p a), mem_support_to_fun := λ a, by split_ifs; { simp only [h, mem_filter, mem_support_iff], tauto } } lemma filter_apply (a : α) [D : decidable (p a)] : f.filter p a = if p a then f a else 0 := @@ -849,16 +872,16 @@ by simp only [fun_like.ext_iff, filter_eq_indicator, set.indicator_apply_eq_self not_imp_comm] @[simp] lemma filter_apply_pos {a : α} (h : p a) : f.filter p a = f a := -if_pos h +by { classical, convert if_pos h } @[simp] lemma filter_apply_neg {a : α} (h : ¬ p a) : f.filter p a = 0 := -if_neg h +by { classical, convert if_neg h } @[simp] lemma support_filter [D : decidable_pred p] : (f.filter p).support = f.support.filter p := by rw subsingleton.elim D; refl lemma filter_zero : (0 : α →₀ M).filter p = 0 := -by rw [← support_eq_empty, support_filter, support_zero, finset.filter_empty] +by { classical, rw [← support_eq_empty, support_filter, support_zero, finset.filter_empty] } @[simp] lemma filter_single_of_pos {a : α} {b : M} (h : p a) : (single a b).filter p = single a b := @@ -870,6 +893,7 @@ by rw [← support_eq_empty, support_filter, support_zero, finset.filter_empty] @[to_additive] lemma prod_filter_index [comm_monoid N] (g : α → M → N) : (f.filter p).prod g = ∏ x in (f.filter p).support, g x (f x) := begin + classical, refine finset.prod_congr rfl (λ x hx, _), rw [support_filter, finset.mem_filter] at hx, rw [filter_apply_pos _ _ hx.2] @@ -877,7 +901,10 @@ end @[simp, to_additive] lemma prod_filter_mul_prod_filter_not [comm_monoid N] (g : α → M → N) : (f.filter p).prod g * (f.filter (λ a, ¬ p a)).prod g = f.prod g := -by simp_rw [prod_filter_index, support_filter, prod_filter_mul_prod_filter_not, finsupp.prod] +begin + classical, + simp_rw [prod_filter_index, support_filter, prod_filter_mul_prod_filter_not, finsupp.prod] +end @[simp, to_additive] lemma prod_div_prod_filter [comm_group G] (g : α → M → G) : f.prod g / (f.filter p).prod g = (f.filter (λ a, ¬p a)).prod g := @@ -897,11 +924,12 @@ section frange variables [has_zero M] /-- `frange f` is the image of `f` on the support of `f`. -/ -def frange (f : α →₀ M) : finset M := finset.image f f.support +def frange (f : α →₀ M) : finset M := +by haveI := classical.dec_eq M; exact finset.image f f.support theorem mem_frange {f : α →₀ M} {y : M} : y ∈ f.frange ↔ y ≠ 0 ∧ ∃ x, f x = y := -finset.mem_image.trans +by classical; exact finset.mem_image.trans ⟨λ ⟨x, hx1, hx2⟩, ⟨hx2 ▸ mem_support_iff.1 hx1, x, hx2⟩, λ ⟨hy, x, hx⟩, ⟨x, mem_support_iff.2 (hx.symm ▸ hy), hx⟩⟩ @@ -909,8 +937,13 @@ theorem zero_not_mem_frange {f : α →₀ M} : (0:M) ∉ f.frange := λ H, (mem_frange.1 H).1 rfl theorem frange_single {x : α} {y : M} : frange (single x y) ⊆ {y} := -λ r hr, let ⟨t, ht1, ht2⟩ := mem_frange.1 hr in ht2 ▸ - (by rw single_apply at ht2 ⊢; split_ifs at ht2 ⊢; [exact finset.mem_singleton_self _, cc]) +λ r hr, let ⟨t, ht1, ht2⟩ := mem_frange.1 hr in ht2 ▸ begin + classical, + rw single_apply at ht2 ⊢, + split_ifs at ht2 ⊢, + { exact finset.mem_singleton_self _ }, + { exact (t ht2.symm).elim } +end end frange @@ -925,7 +958,9 @@ variables [has_zero M] {p : α → Prop} /-- `subtype_domain p f` is the restriction of the finitely supported function `f` to subtype `p`. -/ def subtype_domain (p : α → Prop) (f : α →₀ M) : (subtype p →₀ M) := -⟨f.support.subtype p, f ∘ coe, λ a, by simp only [mem_subtype, mem_support_iff]⟩ +{ support := by haveI := classical.dec_pred p; exact f.support.subtype p, + to_fun := f ∘ coe, + mem_support_to_fun := λ a, by simp only [mem_subtype, mem_support_iff] } @[simp] lemma support_subtype_domain [D : decidable_pred p] {f : α →₀ M} : (subtype_domain p f).support = f.support.subtype p := @@ -940,19 +975,23 @@ rfl lemma subtype_domain_eq_zero_iff' {f : α →₀ M} : f.subtype_domain p = 0 ↔ ∀ x, p x → f x = 0 := -by simp_rw [← support_eq_empty, support_subtype_domain, subtype_eq_empty, not_mem_support_iff] +begin + classical, + simp_rw [← support_eq_empty, support_subtype_domain, subtype_eq_empty, not_mem_support_iff] +end lemma subtype_domain_eq_zero_iff {f : α →₀ M} (hf : ∀ x ∈ f.support , p x) : f.subtype_domain p = 0 ↔ f = 0 := subtype_domain_eq_zero_iff'.trans ⟨λ H, ext $ λ x, - if hx : p x then H x hx else not_mem_support_iff.1 $ mt (hf x) hx, λ H x _, by simp [H]⟩ + by classical; exact + if hx : p x then H x hx else not_mem_support_iff.1 $ mt (hf x) hx, λ H x _, by simp [H]⟩ @[to_additive] lemma prod_subtype_domain_index [comm_monoid N] {v : α →₀ M} {h : α → M → N} (hp : ∀x∈v.support, p x) : (v.subtype_domain p).prod (λa b, h a b) = v.prod h := prod_bij (λp _, p.val) - (λ _, mem_subtype.1) + (λ _, by classical; exact mem_subtype.1) (λ _ _, rfl) (λ _ _ _ _, subtype.eq) (λ b hb, ⟨⟨b, hp b hb⟩, mem_subtype.2 hb, rfl⟩) @@ -1075,6 +1114,7 @@ f.sum $ λp c, single p.1 (single p.2 c) @[simp] lemma curry_apply (f : (α × β) →₀ M) (x : α) (y : β) : f.curry x y = f (x, y) := begin + classical, have : ∀ (b : α × β), single b.fst (single b.snd (f b)) x y = if b = (x, y) then f b else 0, { rintros ⟨b₁, b₂⟩, simp [single_apply, ite_apply, prod.ext_iff, ite_and], @@ -1115,6 +1155,7 @@ by refine ⟨finsupp.curry, finsupp.uncurry, λ f, _, λ f, _⟩; simp only [ lemma filter_curry (f : α × β →₀ M) (p : α → Prop) : (f.filter (λa:α×β, p a.1)).curry = f.curry.filter p := begin + classical, rw [finsupp.curry, finsupp.curry, finsupp.sum, finsupp.sum, filter_sum, support_filter, sum_filter], refine finset.sum_congr rfl _, @@ -1143,7 +1184,8 @@ section sum def sum_elim {α β γ : Type*} [has_zero γ] (f : α →₀ γ) (g : β →₀ γ) : α ⊕ β →₀ γ := on_finset - ((f.support.map ⟨_, sum.inl_injective⟩) ∪ g.support.map ⟨_, sum.inr_injective⟩) + (by haveI := classical.dec_eq α; haveI := classical.dec_eq β; + exact (f.support.map ⟨_, sum.inl_injective⟩) ∪ g.support.map ⟨_, sum.inr_injective⟩) (sum.elim f g) (λ ab h, by { cases ab with a b; simp only [sum.elim_inl, sum.elim_inr] at h; simpa }) @@ -1475,12 +1517,15 @@ between the subtype of finitely supported functions with support contained in `s the type of finitely supported functions from `s`. -/ def restrict_support_equiv (s : set α) (M : Type*) [add_comm_monoid M] : {f : α →₀ M // ↑f.support ⊆ s } ≃ (s →₀ M) := -begin - refine ⟨λf, subtype_domain (λx, x ∈ s) f.1, λ f, ⟨f.map_domain subtype.val, _⟩, _, _⟩, - { refine set.subset.trans (finset.coe_subset.2 map_domain_support) _, +{ to_fun := λ f, subtype_domain (λ x, x ∈ s) f.1, + inv_fun := λ f, ⟨f.map_domain subtype.val, begin + classical, + refine set.subset.trans (finset.coe_subset.2 map_domain_support) _, rw [finset.coe_image, set.image_subset_iff], - exact assume x hx, x.2 }, - { rintros ⟨f, hf⟩, + exact assume x hx, x.2, + end⟩, + left_inv := begin + rintros ⟨f, hf⟩, apply subtype.eq, ext a, dsimp only, @@ -1490,12 +1535,13 @@ begin { convert map_domain_notin_range _ _ h, rw [← not_mem_support_iff], refine mt _ h, - exact assume ha, ⟨⟨a, hf ha⟩, rfl⟩ } }, - { assume f, + exact assume ha, ⟨⟨a, hf ha⟩, rfl⟩ } + end, + right_inv := λ f, begin ext ⟨a, ha⟩, dsimp only, - rw [subtype_domain_apply, map_domain_apply subtype.val_injective] } -end + rw [subtype_domain_apply, map_domain_apply subtype.val_injective] + end } /-- Given `add_comm_monoid M` and `e : α ≃ β`, `dom_congr e` is the corresponding `equiv` between `α →₀ M` and `β →₀ M`. @@ -1556,7 +1602,8 @@ end /-- Given `l`, a finitely supported function from the sigma type `Σ (i : ι), αs i` to `β`, `split_support l` is the finset of indices in `ι` that appear in the support of `l`. -/ -def split_support : finset ι := l.support.image sigma.fst +def split_support (l : (Σ i, αs i) →₀ M) : finset ι := +by haveI := classical.dec_eq ι; exact l.support.image sigma.fst lemma mem_split_support_iff_nonzero (i : ι) : i ∈ split_support l ↔ split l i ≠ 0 := diff --git a/src/data/finsupp/defs.lean b/src/data/finsupp/defs.lean index ac9ae1845a73b..da86462d707b7 100644 --- a/src/data/finsupp/defs.lean +++ b/src/data/finsupp/defs.lean @@ -80,7 +80,7 @@ This file is a `noncomputable theory` and uses classical logic throughout. noncomputable theory open finset function -open_locale classical big_operators +open_locale big_operators variables {α β γ ι M M' N P G H R S : Type*} @@ -146,7 +146,7 @@ by rw [← coe_zero, coe_fn_inj] lemma ext_iff' {f g : α →₀ M} : f = g ↔ f.support = g.support ∧ ∀ x ∈ f.support, f x = g x := ⟨λ h, h ▸ ⟨rfl, λ _ _, rfl⟩, λ ⟨h₁, h₂⟩, ext $ λ a, - if h : a ∈ f.support then h₂ a h else + by classical; exact if h : a ∈ f.support then h₂ a h else have hf : f a = 0, from not_mem_support_iff.1 h, have hg : g a = 0, by rwa [h₁, not_mem_support_iff] at h, by rw [hf, hg]⟩ @@ -202,14 +202,17 @@ variables [has_zero M] {a a' : α} {b : M} /-- `single a b` is the finitely supported function with value `b` at `a` and zero otherwise. -/ def single (a : α) (b : M) : α →₀ M := -⟨if b = 0 then ∅ else {a}, pi.single a b, λ a', begin - obtain rfl | hb := eq_or_ne b 0, - { simp }, - rw [if_neg hb, mem_singleton], - obtain rfl | ha := eq_or_ne a' a, - { simp [hb] }, - simp [pi.single_eq_of_ne', ha], -end⟩ +{ support := by haveI := classical.dec_eq M; exact if b = 0 then ∅ else {a}, + to_fun := by haveI := classical.dec_eq α; exact pi.single a b, + mem_support_to_fun := λ a', begin + classical, + obtain rfl | hb := eq_or_ne b 0, + { simp }, + rw [if_neg hb, mem_singleton], + obtain rfl | ha := eq_or_ne a' a, + { simp [hb] }, + simp [pi.single_eq_of_ne', ha], + end } lemma single_apply [decidable (a = a')] : single a b a' = if a = a' then b else 0 := by { simp_rw [@eq_comm _ a a'], convert pi.single_apply _ _ _, } @@ -401,26 +404,29 @@ If `b = 0`, this amounts to removing `a` from the `finsupp.support`. Otherwise, if `a` was not in the `finsupp.support`, it is added to it. This is the finitely-supported version of `function.update`. -/ -def update : α →₀ M := -⟨if b = 0 then f.support.erase a else insert a f.support, - function.update f a b, - λ i, begin +def update (f : α →₀ M) (a : α) (b : M) : α →₀ M := +{ support := by haveI := classical.dec_eq α; haveI := classical.dec_eq M; exact + if b = 0 then f.support.erase a else insert a f.support, + to_fun := by haveI := classical.dec_eq α; exact + function.update f a b, + mem_support_to_fun := λ i, begin simp only [function.update_apply, ne.def], split_ifs with hb ha ha hb; simp [ha, hb] - end⟩ + end } @[simp] lemma coe_update [decidable_eq α] : (f.update a b : α → M) = function.update f a b := by convert rfl @[simp] lemma update_self : f.update a (f a) = f := -by { ext, simp } +by { classical, ext, simp } @[simp] lemma zero_update : update 0 a b = single a b := -by { ext, rw single_eq_update, refl } +by { classical, ext, rw single_eq_update, refl } -lemma support_update [decidable_eq α] : support (f.update a b) = - if b = 0 then f.support.erase a else insert a f.support := by convert rfl +lemma support_update [decidable_eq α] [decidable_eq M] : + support (f.update a b) = if b = 0 then f.support.erase a else insert a f.support := +by convert rfl @[simp] lemma support_update_zero [decidable_eq α] : support (f.update a 0) = f.support.erase a := by convert if_pos rfl @@ -428,7 +434,7 @@ lemma support_update [decidable_eq α] : support (f.update a b) = variables {b} lemma support_update_ne_zero [decidable_eq α] (h : b ≠ 0) : - support (f.update a b) = insert a f.support := by convert if_neg h + support (f.update a b) = insert a f.support := by { classical, convert if_neg h } end update @@ -443,20 +449,21 @@ variables [has_zero M] If `a` is not in the support of `f` then `erase a f = f`. -/ def erase (a : α) (f : α →₀ M) : α →₀ M := -⟨f.support.erase a, (λa', if a' = a then 0 else f a'), - assume a', by rw [mem_erase, mem_support_iff]; split_ifs; +{ support := by haveI := classical.dec_eq α; exact f.support.erase a, + to_fun := λ a', by haveI := classical.dec_eq α; exact if a' = a then 0 else f a', + mem_support_to_fun := assume a', by rw [mem_erase, mem_support_iff]; split_ifs; [exact ⟨λ H _, H.1 h, λ H, (H rfl).elim⟩, - exact and_iff_right h]⟩ + exact and_iff_right h] } @[simp] lemma support_erase [decidable_eq α] {a : α} {f : α →₀ M} : (f.erase a).support = f.support.erase a := by convert rfl @[simp] lemma erase_same {a : α} {f : α →₀ M} : (f.erase a) a = 0 := -if_pos rfl +by convert if_pos rfl @[simp] lemma erase_ne {a a' : α} {f : α →₀ M} (h : a' ≠ a) : (f.erase a) a' = f a' := -if_neg h +by { classical, convert if_neg h } @[simp] lemma erase_single {a : α} {b : M} : (erase a (single a b)) = 0 := begin @@ -480,7 +487,7 @@ begin end @[simp] lemma erase_zero (a : α) : erase a (0 : α →₀ M) = 0 := -by rw [← support_eq_empty, support_erase, support_zero, erase_empty] +by { classical, rw [← support_eq_empty, support_erase, support_zero, erase_empty] } end erase @@ -493,7 +500,9 @@ variables [has_zero M] The function must be `0` outside of `s`. Use this when the set needs to be filtered anyways, otherwise a better set representation is often available. -/ def on_finset (s : finset α) (f : α → M) (hf : ∀a, f a ≠ 0 → a ∈ s) : α →₀ M := -⟨s.filter (λa, f a ≠ 0), f, by simpa⟩ +{ support := by haveI := classical.dec_eq M; exact s.filter (λa, f a ≠ 0), + to_fun := f, + mem_support_to_fun := by simpa } @[simp] lemma on_finset_apply {s : finset α} {f : α → M} {hf a} : (on_finset s f hf : α →₀ M) a = f a := @@ -501,17 +510,17 @@ rfl @[simp] lemma support_on_finset_subset {s : finset α} {f : α → M} {hf} : (on_finset s f hf).support ⊆ s := -filter_subset _ _ +by convert filter_subset _ _ @[simp] lemma mem_support_on_finset {s : finset α} {f : α → M} (hf : ∀ (a : α), f a ≠ 0 → a ∈ s) {a : α} : a ∈ (finsupp.on_finset s f hf).support ↔ f a ≠ 0 := by rw [finsupp.mem_support_iff, finsupp.on_finset_apply] -lemma support_on_finset +lemma support_on_finset [decidable_eq M] {s : finset α} {f : α → M} (hf : ∀ (a : α), f a ≠ 0 → a ∈ s) : (finsupp.on_finset s f hf).support = s.filter (λ a, f a ≠ 0) := -rfl +by convert rfl end on_finset @@ -577,7 +586,10 @@ support_on_finset_subset @[simp] lemma map_range_single {f : M → N} {hf : f 0 = 0} {a : α} {b : M} : map_range f hf (single a b) = single a (f b) := -ext $ λ a', by simpa only [single_eq_pi_single] using pi.apply_single _ (λ _, hf) a _ a' +ext $ λ a', begin + classical, + simpa only [single_eq_pi_single] using pi.apply_single _ (λ _, hf) a _ a' +end lemma support_map_range_of_injective {e : M → N} (he0 : e 0 = 0) (f : ι →₀ M) (he : function.injective e) : @@ -599,18 +611,22 @@ variables [has_zero M] [has_zero N] is the finitely supported function whose value at `f a : β` is `v a`. For a `b : β` outside the range of `f`, it is zero. -/ def emb_domain (f : α ↪ β) (v : α →₀ M) : β →₀ M := -begin - refine ⟨v.support.map f, λa₂, - if h : a₂ ∈ v.support.map f then v (v.support.choose (λa₁, f a₁ = a₂) _) else 0, _⟩, - { rcases finset.mem_map.1 h with ⟨a, ha, rfl⟩, - exact exists_unique.intro a ⟨ha, rfl⟩ (assume b ⟨_, hb⟩, f.injective hb) }, - { assume a₂, +{ support := v.support.map f, + to_fun := λ a₂, + by haveI := classical.dec_eq β; exact + if h : a₂ ∈ v.support.map f + then v (v.support.choose (λa₁, f a₁ = a₂) begin + rcases finset.mem_map.1 h with ⟨a, ha, rfl⟩, + exact exists_unique.intro a ⟨ha, rfl⟩ (assume b ⟨_, hb⟩, f.injective hb) + end) + else 0, + mem_support_to_fun := λ a₂, begin split_ifs, { simp only [h, true_iff, ne.def], rw [← not_mem_support_iff, not_not], apply finset.choose_mem }, - { simp only [h, ne.def, ne_self_iff_false] } } -end + { simp only [h, ne.def, ne_self_iff_false] } + end } @[simp] lemma support_emb_domain (f : α ↪ β) (v : α →₀ M) : (emb_domain f v).support = v.support.map f := @@ -622,6 +638,7 @@ rfl @[simp] lemma emb_domain_apply (f : α ↪ β) (v : α →₀ M) (a : α) : emb_domain f v (f a) = v a := begin + classical, change dite _ _ _ = _, split_ifs; rw [finset.mem_map' f] at h, { refine congr_arg (v : α → M) (f.inj' _), @@ -632,6 +649,7 @@ end lemma emb_domain_notin_range (f : α ↪ β) (v : α →₀ M) (a : β) (h : a ∉ set.range f) : emb_domain f v a = 0 := begin + classical, refine dif_neg (mt (assume h, _) h), rcases finset.mem_map.1 h with ⟨a, h, rfl⟩, exact set.mem_range_self a @@ -665,6 +683,7 @@ lemma single_of_emb_domain_single (h : l.emb_domain f = single a b) : ∃ x, l = single x b ∧ f x = a := begin + classical, have h_map_support : finset.map f (l.support) = {a}, by rw [←support_emb_domain, h, support_single_ne_zero _ hb]; refl, have ha : a ∈ finset.map f (l.support), @@ -685,6 +704,7 @@ end @[simp] lemma emb_domain_single (f : α ↪ β) (a : α) (m : M) : emb_domain f (single a m) = single (f a) m := begin + classical, ext b, by_cases h : b ∈ set.range f, { rcases h with ⟨a', rfl⟩, @@ -706,11 +726,13 @@ variables [has_zero M] [has_zero N] [has_zero P] `zip_with f hf g₁ g₂` is the finitely supported function `α →₀ P` satisfying `zip_with f hf g₁ g₂ a = f (g₁ a) (g₂ a)`, which is well-defined when `f 0 0 = 0`. -/ def zip_with (f : M → N → P) (hf : f 0 0 = 0) (g₁ : α →₀ M) (g₂ : α →₀ N) : α →₀ P := -on_finset (g₁.support ∪ g₂.support) (λa, f (g₁ a) (g₂ a)) $ λ a H, -begin - simp only [mem_union, mem_support_iff, ne], rw [← not_and_distrib], - rintro ⟨h₁, h₂⟩, rw [h₁, h₂] at H, exact H hf -end +on_finset + (by haveI := classical.dec_eq α; exact g₁.support ∪ g₂.support) + (λa, f (g₁ a) (g₂ a)) $ λ a H, + begin + simp only [mem_union, mem_support_iff, ne], rw [← not_and_distrib], + rintro ⟨h₁, h₂⟩, rw [h₁, h₂] at H, exact H hf + end @[simp] lemma zip_with_apply {f : M → N → P} {hf : f 0 0 = 0} {g₁ : α →₀ M} {g₂ : α →₀ N} {a : α} : @@ -782,6 +804,7 @@ noncomputable def coe_fn_add_hom : (α →₀ M) →+ (α → M) := lemma update_eq_single_add_erase (f : α →₀ M) (a : α) (b : M) : f.update a b = single a b + f.erase a := begin + classical, ext j, rcases eq_or_ne a j with rfl|h, { simp }, @@ -791,6 +814,7 @@ end lemma update_eq_erase_add_single (f : α →₀ M) (a : α) (b : M) : f.update a b = f.erase a + single a b := begin + classical, ext j, rcases eq_or_ne a j with rfl|h, { simp }, @@ -820,30 +844,33 @@ protected theorem induction {p : (α →₀ M) → Prop} (f : α →₀ M) (h0 : p 0) (ha : ∀a b (f : α →₀ M), a ∉ f.support → b ≠ 0 → p f → p (single a b + f)) : p f := suffices ∀s (f : α →₀ M), f.support = s → p f, from this _ _ rfl, -assume s, finset.induction_on s (λ f hf, by rwa [support_eq_empty.1 hf]) $ +assume s, finset.cons_induction_on s (λ f hf, by rwa [support_eq_empty.1 hf]) $ assume a s has ih f hf, suffices p (single a (f a) + f.erase a), by rwa [single_add_erase] at this, begin + classical, apply ha, { rw [support_erase, mem_erase], exact λ H, H.1 rfl }, - { rw [← mem_support_iff, hf], exact mem_insert_self _ _ }, + { rw [← mem_support_iff, hf], exact mem_cons_self _ _ }, { apply ih _ _, - rw [support_erase, hf, finset.erase_insert has] } + rw [support_erase, hf, finset.erase_cons] } end lemma induction₂ {p : (α →₀ M) → Prop} (f : α →₀ M) (h0 : p 0) (ha : ∀a b (f : α →₀ M), a ∉ f.support → b ≠ 0 → p f → p (f + single a b)) : p f := suffices ∀s (f : α →₀ M), f.support = s → p f, from this _ _ rfl, -assume s, finset.induction_on s (λ f hf, by rwa [support_eq_empty.1 hf]) $ +assume s, finset.cons_induction_on s (λ f hf, by rwa [support_eq_empty.1 hf]) $ assume a s has ih f hf, suffices p (f.erase a + single a (f a)), by rwa [erase_add_single] at this, begin + classical, apply ha, { rw [support_erase, mem_erase], exact λ H, H.1 rfl }, - { rw [← mem_support_iff, hf], exact mem_insert_self _ _ }, + { rw [← mem_support_iff, hf], + exact mem_cons_self _ _ }, { apply ih _ _, - rw [support_erase, hf, finset.erase_insert has] } + rw [support_erase, hf, finset.erase_cons] } end lemma induction_linear {p : (α →₀ M) → Prop} (f : α →₀ M) @@ -988,6 +1015,7 @@ lemma single_add_single_eq_single_add_single [add_comm_monoid M] single k u + single l v = single m u + single n v ↔ (k = m ∧ l = n) ∨ (u = v ∧ k = n ∧ l = m) ∨ (u + v = 0 ∧ k = l ∧ m = n) := begin + classical, simp_rw [fun_like.ext_iff, coe_add, single_eq_pi_single, ←funext_iff], exact pi.single_add_single_eq_single_add_single hu hv, end diff --git a/src/data/fintype/array.lean b/src/data/fintype/array.lean new file mode 100644 index 0000000000000..5211c8cebc957 --- /dev/null +++ b/src/data/fintype/array.lean @@ -0,0 +1,20 @@ +/- +Copyright (c) 2017 Mario Carneiro. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Mario Carneiro +-/ +import data.fintype.pi +import logic.equiv.array + +/-! +# `array n α` is a fintype when `α` is. +-/ + +variables {α : Type*} + +instance d_array.fintype {n : ℕ} {α : fin n → Type*} + [∀ n, fintype (α n)] : fintype (d_array n α) := +fintype.of_equiv _ (equiv.d_array_equiv_fin _).symm + +instance array.fintype {n : ℕ} {α : Type*} [fintype α] : fintype (array n α) := +d_array.fintype diff --git a/src/data/fintype/list.lean b/src/data/fintype/list.lean index 1accc08def6a3..2c9201bc56e31 100644 --- a/src/data/fintype/list.lean +++ b/src/data/fintype/list.lean @@ -10,6 +10,9 @@ import data.finset.powerset # Fintype instance for nodup lists +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + The subtype of `{l : list α // l.nodup}` over a `[fintype α]` admits a `fintype` instance. diff --git a/src/data/fintype/sigma.lean b/src/data/fintype/sigma.lean index 17ca0c138b297..7cea94a731f50 100644 --- a/src/data/fintype/sigma.lean +++ b/src/data/fintype/sigma.lean @@ -9,6 +9,9 @@ import data.finset.sigma /-! # fintype instances for sigma types + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ open function diff --git a/src/data/fintype/vector.lean b/src/data/fintype/vector.lean index 07ce3c1ef85c8..9b4b98cf30adb 100644 --- a/src/data/fintype/vector.lean +++ b/src/data/fintype/vector.lean @@ -4,22 +4,14 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ import data.fintype.pi -import data.array.lemmas import data.sym.basic /-! -# `vector α n` is a fintype when `α` is. +# `vector α n` and `sym α n` are fintypes when `α` is. -/ variables {α : Type*} -instance d_array.fintype {n : ℕ} {α : fin n → Type*} - [∀ n, fintype (α n)] : fintype (d_array n α) := -fintype.of_equiv _ (equiv.d_array_equiv_fin _).symm - -instance array.fintype {n : ℕ} {α : Type*} [fintype α] : fintype (array n α) := -d_array.fintype - instance vector.fintype [fintype α] {n : ℕ} : fintype (vector α n) := fintype.of_equiv _ (equiv.vector_equiv_fin _ _).symm diff --git a/src/data/list/alist.lean b/src/data/list/alist.lean index 8648c1c2b183e..79b6778ecb5a8 100644 --- a/src/data/list/alist.lean +++ b/src/data/list/alist.lean @@ -205,6 +205,11 @@ theorem insert_entries_of_neg {a} {b : β a} {s : alist β} (h : a ∉ s) : (insert a b s).entries = ⟨a, b⟩ :: s.entries := by rw [insert_entries, kerase_of_not_mem_keys h] +-- Todo: rename to `insert_of_not_mem`. +theorem insert_of_neg {a} {b : β a} {s : alist β} (h : a ∉ s) : + insert a b s = ⟨⟨a, b⟩ :: s.entries, nodupkeys_cons.2 ⟨h, s.2⟩⟩ := +ext $ insert_entries_of_neg h + @[simp] theorem insert_empty (a) (b : β a) : insert a b ∅ = singleton a b := rfl @[simp] theorem mem_insert {a a'} {b' : β a'} (s : alist β) : @@ -250,6 +255,50 @@ ext $ by simp only [alist.insert_entries, list.kerase_cons_eq, and_self, alist.s theorem to_alist_cons (a : α) (b : β a) (xs : list (sigma β)) : list.to_alist (⟨a,b⟩ :: xs) = insert a b xs.to_alist := rfl +theorem mk_cons_eq_insert (c : sigma β) (l : list (sigma β)) (h : (c :: l).nodupkeys) : + (⟨c :: l, h⟩ : alist β) = insert c.1 c.2 ⟨l, nodupkeys_of_nodupkeys_cons h⟩ := +by simpa [insert] using (kerase_of_not_mem_keys $ not_mem_keys_of_nodupkeys_cons h).symm + +/-- Recursion on an `alist`, using `insert`. Use as `induction l using alist.insert_rec`. -/ +@[elab_as_eliminator] def insert_rec {C : alist β → Sort*} (H0 : C ∅) + (IH : Π (a : α) (b : β a) (l : alist β) (h : a ∉ l), C l → C (l.insert a b)) : Π l : alist β, C l +| ⟨[], _⟩ := H0 +| ⟨c :: l, h⟩ := begin + rw mk_cons_eq_insert, + refine IH _ _ _ _ (insert_rec _), + exact not_mem_keys_of_nodupkeys_cons h +end + +-- Test that the `induction` tactic works on `insert_rec`. +example (l : alist β) : true := by induction l using alist.insert_rec; trivial + +@[simp] theorem insert_rec_empty {C : alist β → Sort*} (H0 : C ∅) + (IH : Π (a : α) (b : β a) (l : alist β) (h : a ∉ l), C l → C (l.insert a b)) : + @insert_rec α β _ C H0 IH ∅ = H0 := +by { change @insert_rec α β _ C H0 IH ⟨[], _⟩ = H0, rw insert_rec } + +theorem insert_rec_insert {C : alist β → Sort*} (H0 : C ∅) + (IH : Π (a : α) (b : β a) (l : alist β) (h : a ∉ l), C l → C (l.insert a b)) + {c : sigma β} {l : alist β} (h : c.1 ∉ l) : + @insert_rec α β _ C H0 IH (l.insert c.1 c.2) = IH c.1 c.2 l h (@insert_rec α β _ C H0 IH l) := +begin + cases l with l hl, + suffices : @insert_rec α β _ C H0 IH ⟨c :: l, nodupkeys_cons.2 ⟨h, hl⟩⟩ == + IH c.1 c.2 ⟨l, hl⟩ h (@insert_rec α β _ C H0 IH ⟨l, hl⟩), + { cases c, + apply eq_of_heq, + convert this; + rw insert_of_neg h }, + rw insert_rec, + apply cast_heq +end + +theorem recursion_insert_mk {C : alist β → Sort*} (H0 : C ∅) + (IH : Π (a : α) (b : β a) (l : alist β) (h : a ∉ l), C l → C (l.insert a b)) + {a : α} (b : β a) {l : alist β} (h : a ∉ l) : + @insert_rec α β _ C H0 IH (l.insert a b) = IH a b l h (@insert_rec α β _ C H0 IH l) := +@insert_rec_insert α β _ C H0 IH ⟨a, b⟩ l h + /-! ### extract -/ /-- Erase a key from the map, and return the corresponding value, if found. -/ diff --git a/src/data/list/of_fn.lean b/src/data/list/of_fn.lean index 9a2410b3f0614..4a578e41e8680 100644 --- a/src/data/list/of_fn.lean +++ b/src/data/list/of_fn.lean @@ -10,6 +10,9 @@ import data.list.join /-! # Lists from functions +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + Theorems and lemmas for dealing with `list.of_fn`, which converts a function on `fin n` to a list of length `n`. diff --git a/src/data/list/sigma.lean b/src/data/list/sigma.lean index 10f58ae9d0eca..b24f8564ce41a 100644 --- a/src/data/list/sigma.lean +++ b/src/data/list/sigma.lean @@ -80,6 +80,12 @@ nodupkeys_iff_pairwise.1 h nodupkeys (s::l) ↔ s.1 ∉ l.keys ∧ nodupkeys l := by simp [keys, nodupkeys] +theorem not_mem_keys_of_nodupkeys_cons {s : sigma β} {l : list (sigma β)} (h : nodupkeys (s :: l)) : + s.1 ∉ l.keys := (nodupkeys_cons.1 h).1 + +theorem nodupkeys_of_nodupkeys_cons {s : sigma β} {l : list (sigma β)} (h : nodupkeys (s :: l)) : + nodupkeys l := (nodupkeys_cons.1 h).2 + theorem nodupkeys.eq_of_fst_eq {l : list (sigma β)} (nd : nodupkeys l) {s s' : sigma β} (h : s ∈ l) (h' : s' ∈ l) : s.1 = s'.1 → s = s' := diff --git a/src/data/matrix/dmatrix.lean b/src/data/matrix/dmatrix.lean index 2fe70c3422ed0..4a80dde055876 100644 --- a/src/data/matrix/dmatrix.lean +++ b/src/data/matrix/dmatrix.lean @@ -8,6 +8,9 @@ import data.fintype.basic /-! # Matrices + +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. -/ universes u u' v w z diff --git a/src/data/nat/basic.lean b/src/data/nat/basic.lean index 4b5a0b8a04040..d8f47d7509cdd 100644 --- a/src/data/nat/basic.lean +++ b/src/data/nat/basic.lean @@ -210,18 +210,14 @@ by { rw ←not_iff_not, push_neg, exact forall_lt_succ } @[simp] theorem add_def {a b : ℕ} : nat.add a b = a + b := rfl @[simp] theorem mul_def {a b : ℕ} : nat.mul a b = a * b := rfl -lemma exists_eq_add_of_le : ∀ {m n : ℕ}, m ≤ n → ∃ k : ℕ, n = m + k -| 0 0 h := ⟨0, by simp⟩ -| 0 (n+1) h := ⟨n+1, by simp⟩ -| (m+1) (n+1) h := - let ⟨k, hk⟩ := exists_eq_add_of_le (nat.le_of_succ_le_succ h) in - ⟨k, by simp [hk, add_comm, add_left_comm]⟩ - -lemma exists_eq_add_of_lt : ∀ {m n : ℕ}, m < n → ∃ k : ℕ, n = m + k + 1 -| 0 0 h := false.elim $ lt_irrefl _ h -| 0 (n+1) h := ⟨n, by simp⟩ -| (m+1) (n+1) h := let ⟨k, hk⟩ := exists_eq_add_of_le (nat.le_of_succ_le_succ h) in - ⟨k, by simp [hk]⟩ +lemma exists_eq_add_of_le (h : m ≤ n) : ∃ k : ℕ, n = m + k := +⟨n - m, (nat.add_sub_of_le h).symm⟩ + +lemma exists_eq_add_of_le' (h : m ≤ n) : ∃ k : ℕ, n = k + m := +⟨n - m, (nat.sub_add_cancel h).symm⟩ + +lemma exists_eq_add_of_lt (h : m < n) : ∃ k : ℕ, n = m + k + 1 := +⟨n - (m + 1), by rw [add_right_comm, nat.add_sub_of_le h]⟩ /-! ### `pred` -/ diff --git a/src/data/nat/digits.lean b/src/data/nat/digits.lean index ba1681b9897fb..248bf2a028399 100644 --- a/src/data/nat/digits.lean +++ b/src/data/nat/digits.lean @@ -79,8 +79,8 @@ by rcases b with _|⟨_|⟨_⟩⟩; simp [digits, digits_aux_0, digits_aux_1] @[simp] lemma digits_zero_succ (n : ℕ) : digits 0 (n.succ) = [n+1] := rfl -theorem digits_zero_succ' : ∀ {n : ℕ} (w : 0 < n), digits 0 n = [n] -| 0 h := absurd h dec_trivial +theorem digits_zero_succ' : ∀ {n : ℕ}, n ≠ 0 → digits 0 n = [n] +| 0 h := (h rfl).elim | (n+1) _ := rfl @[simp] lemma digits_one (n : ℕ) : digits 1 n = list.replicate n 1 := rfl @@ -91,40 +91,31 @@ theorem digits_zero_succ' : ∀ {n : ℕ} (w : 0 < n), digits 0 n = [n] digits (b+2) (n+1) = (((n+1) % (b+2)) :: digits (b+2) ((n+1) / (b+2))) := by { rw [digits, digits_aux_def], exact succ_pos n } -theorem digits_def' : ∀ {b : ℕ} (h : 2 ≤ b) {n : ℕ} (w : 0 < n), +theorem digits_def' : ∀ {b : ℕ} (h : 1 < b) {n : ℕ} (w : 0 < n), digits b n = n % b :: digits b (n/b) | 0 h := absurd h dec_trivial | 1 h := absurd h dec_trivial | (b+2) h := digits_aux_def _ _ -@[simp] -lemma digits_of_lt (b x : ℕ) (w₁ : 0 < x) (w₂ : x < b) : digits b x = [x] := +@[simp] lemma digits_of_lt (b x : ℕ) (hx : x ≠ 0) (hxb : x < b) : digits b x = [x] := begin - cases b, - { cases w₂ }, - { cases b, - { interval_cases x, }, - { cases x, - { cases w₁, }, - { rw [digits_add_two_add_one, nat.div_eq_of_lt w₂, digits_zero, nat.mod_eq_of_lt w₂] } } } + rcases exists_eq_succ_of_ne_zero hx with ⟨x, rfl⟩, + rcases exists_eq_add_of_le' ((nat.le_add_left 1 x).trans_lt hxb) with ⟨b, rfl⟩, + rw [digits_add_two_add_one, div_eq_of_lt hxb, digits_zero, mod_eq_of_lt hxb] end -lemma digits_add (b : ℕ) (h : 2 ≤ b) (x y : ℕ) (w : x < b) (w' : 0 < x ∨ 0 < y) : +lemma digits_add (b : ℕ) (h : 1 < b) (x y : ℕ) (hxb : x < b) (hxy : x ≠ 0 ∨ y ≠ 0) : digits b (x + b * y) = x :: digits b y := begin - cases b, - { cases h, }, - { cases b, - { norm_num at h, }, - { cases y, - { norm_num at w', - simp [w, w'], }, - dsimp [digits], - rw digits_aux_def, - { congr, - { simp [nat.add_mod, nat.mod_eq_of_lt w], }, - { simp [mul_comm (b+2), nat.add_mul_div_right, nat.div_eq_of_lt w], } }, - { apply nat.succ_pos, }, }, }, + rcases exists_eq_add_of_le' h with ⟨b, rfl : _ = _ + 2⟩, + cases y, + { simp [hxb, hxy.resolve_right (absurd rfl)] }, + dsimp [digits], + rw digits_aux_def, + { congr, + { simp [nat.add_mod, mod_eq_of_lt hxb], }, + { simp [add_mul_div_left, div_eq_of_lt hxb] } }, + { apply nat.succ_pos } end /-- @@ -201,22 +192,14 @@ begin { dsimp [of_digits], push_cast } end -lemma digits_zero_of_eq_zero {b : ℕ} (h : 1 ≤ b) {L : list ℕ} (w : of_digits b L = 0) : - ∀ l ∈ L, l = 0 := -begin - induction L with d L ih, - { intros l m, - cases m, }, - { intros l m, - dsimp [of_digits] at w, - rcases m with ⟨rfl⟩, - { apply nat.eq_zero_of_add_eq_zero_right w }, - { exact ih (mul_right_injective₀ (pos_iff_ne_zero.1 h) - (nat.eq_zero_of_add_eq_zero_left w)) _ m, }, } -end +lemma digits_zero_of_eq_zero {b : ℕ} (h : b ≠ 0) : + ∀ {L : list ℕ} (h0 : of_digits b L = 0) (l ∈ L), l = 0 +| (a :: L) h0 l (or.inl rfl) := nat.eq_zero_of_add_eq_zero_right h0 +| (a :: L) h0 l (or.inr hL) := + digits_zero_of_eq_zero (mul_right_injective₀ h (nat.eq_zero_of_add_eq_zero_left h0)) _ hL lemma digits_of_digits - (b : ℕ) (h : 2 ≤ b) (L : list ℕ) + (b : ℕ) (h : 1 < b) (L : list ℕ) (w₁ : ∀ l ∈ L, l < b) (w₂ : ∀ (h : L ≠ []), L.last h ≠ 0) : digits b (of_digits b L) = L := begin @@ -233,17 +216,13 @@ begin { exact w₁ d (list.mem_cons_self _ _) }, { by_cases h' : L = [], { rcases h' with rfl, - simp at w₂, left, - apply nat.pos_of_ne_zero, - exact w₂ }, + simpa using w₂ }, { right, - apply nat.pos_of_ne_zero, contrapose! w₂, - apply digits_zero_of_eq_zero _ w₂, - { rw list.last_cons h', - exact list.last_mem h', }, - { exact le_of_lt h, }, }, }, }, + refine digits_zero_of_eq_zero h.ne_bot w₂ _ _, + rw list.last_cons h', + exact list.last_mem h' } } } end lemma of_digits_digits (b n : ℕ) : of_digits b (digits b n) = n := @@ -295,7 +274,7 @@ end lemma digits_ne_nil_iff_ne_zero {b n : ℕ} : digits b n ≠ [] ↔ n ≠ 0 := not_congr digits_eq_nil_iff_eq_zero -lemma digits_eq_cons_digits_div {b n : ℕ} (h : 2 ≤ b) (w : 0 < n) : +lemma digits_eq_cons_digits_div {b n : ℕ} (h : 1 < b) (w : n ≠ 0) : digits b n = ((n % b) :: digits b (n / b)) := begin rcases b with _|_|b, @@ -306,12 +285,12 @@ begin simp, end -lemma digits_last {b : ℕ} (m : ℕ) (h : 2 ≤ b) (p q) : +lemma digits_last {b : ℕ} (m : ℕ) (h : 1 < b) (p q) : (digits b m).last p = (digits b (m/b)).last q := begin by_cases hm : m = 0, { simp [hm], }, - simp only [digits_eq_cons_digits_div h (nat.pos_of_ne_zero hm)], + simp only [digits_eq_cons_digits_div h hm], rw list.last_cons, end @@ -322,21 +301,20 @@ function.left_inverse.injective (of_digits_digits b) b.digits n = b.digits m ↔ n = m := (digits.injective b).eq_iff -lemma digits_len (b n : ℕ) (hb : 2 ≤ b) (hn : 0 < n) : +lemma digits_len (b n : ℕ) (hb : 1 < b) (hn : n ≠ 0) : (b.digits n).length = b.log n + 1 := begin induction n using nat.strong_induction_on with n IH, rw [digits_eq_cons_digits_div hb hn, list.length], - cases (n / b).eq_zero_or_pos with h h, - { have posb : 0 < b := zero_lt_two.trans_le hb, - simp [h, log_eq_zero_iff, ←nat.div_eq_zero_iff posb] }, + by_cases h : n / b = 0, + { have hb0 : b ≠ 0 := (nat.succ_le_iff.1 hb).ne_bot, + simp [h, log_eq_zero_iff, ← nat.div_eq_zero_iff hb0.bot_lt] }, { have hb' : 1 < b := one_lt_two.trans_le hb, - have : n / b < n := div_lt_self hn hb', + have : n / b < n := div_lt_self (nat.pos_of_ne_zero hn) hb', rw [IH _ this h, log_div_base, tsub_add_cancel_of_le], - rw [succ_le_iff], - refine log_pos hb' _, + refine nat.succ_le_of_lt (log_pos hb' _), contrapose! h, - rw div_eq_of_lt h } + exact div_eq_of_lt h } end lemma last_digit_ne_zero (b : ℕ) {m : ℕ} (hm : m ≠ 0) : @@ -351,12 +329,10 @@ begin revert hm, apply nat.strong_induction_on m, intros n IH hn, - have hnpos : 0 < n := nat.pos_of_ne_zero hn, by_cases hnb : n < b + 2, - { simp_rw [digits_of_lt b.succ.succ n hnpos hnb], - exact pos_iff_ne_zero.mp hnpos }, + { simpa only [digits_of_lt (b + 2) n hn hnb] }, { rw digits_last n (show 2 ≤ b + 2, from dec_trivial), - refine IH _ (nat.div_lt_self hnpos dec_trivial) _, + refine IH _ (nat.div_lt_self hn.bot_lt dec_trivial) _, { rw ←pos_iff_ne_zero, exact nat.div_pos (le_of_not_lt hnb) dec_trivial } }, end @@ -375,7 +351,7 @@ begin end /-- The digits in the base b expansion of n are all less than b, if b ≥ 2 -/ -lemma digits_lt_base {b m d : ℕ} (hb : 2 ≤ b) (hd : d ∈ digits b m) : d < b := +lemma digits_lt_base {b m d : ℕ} (hb : 1 < b) (hd : d ∈ digits b m) : d < b := begin rcases b with _ | _ | b; try {linarith}, exact digits_lt_base' hd, @@ -397,8 +373,8 @@ begin exact hl hd (list.mem_cons_self _ _) } end -/-- an n-digit number in base b is less than b^n if b ≥ 2 -/ -lemma of_digits_lt_base_pow_length {b : ℕ} {l : list ℕ} (hb : 2 ≤ b) (hl : ∀ x ∈ l, x < b) : +/-- an n-digit number in base b is less than b^n if b > 1 -/ +lemma of_digits_lt_base_pow_length {b : ℕ} {l : list ℕ} (hb : 1 < b) (hl : ∀ x ∈ l, x < b) : of_digits b l < b^l.length := begin rcases b with _ | _ | b; try { linarith }, @@ -413,7 +389,7 @@ begin end /-- Any number m is less than b^(number of digits in the base b representation of m) -/ -lemma lt_base_pow_length_digits {b m : ℕ} (hb : 2 ≤ b) : m < b^(digits b m).length := +lemma lt_base_pow_length_digits {b m : ℕ} (hb : 1 < b) : m < b^(digits b m).length := begin rcases b with _ | _ | b; try { linarith }, exact lt_base_pow_length_digits', @@ -425,13 +401,10 @@ by rw [of_digits_append, of_digits_digits, of_digits_digits] lemma digits_len_le_digits_len_succ (b n : ℕ) : (digits b n).length ≤ (digits b (n + 1)).length := begin - rcases n.eq_zero_or_pos with rfl|hn, + rcases decidable.eq_or_ne n 0 with rfl|hn, { simp }, - cases lt_or_le b 2 with hb hb, - { rcases b with _|_|b, - { simp [digits_zero_succ', hn] }, - { simp, }, - { simpa [succ_lt_succ_iff] using hb } }, + cases le_or_lt b 1 with hb hb, + { interval_cases b; simp [digits_zero_succ', hn] }, simpa [digits_len, hb, hn] using log_mono_right (le_succ _) end @@ -466,7 +439,7 @@ end Any non-zero natural number `m` is greater than b^((number of digits in the base b representation of m) - 1) -/ -lemma base_pow_length_digits_le (b m : ℕ) (hb : 2 ≤ b): m ≠ 0 → b ^ ((digits b m).length) ≤ b * m := +lemma base_pow_length_digits_le (b m : ℕ) (hb : 1 < b): m ≠ 0 → b ^ ((digits b m).length) ≤ b * m := begin rcases b with _ | _ | b; try { linarith }, exact base_pow_length_digits_le' b m, @@ -480,11 +453,10 @@ begin { simp, }, rw bits_append_bit _ _ (λ hn, absurd hn h), cases b, - { rw digits_def' (le_refl 2), + { rw digits_def' one_lt_two, { simpa [nat.bit, nat.bit0_val n], }, { simpa [pos_iff_ne_zero, bit_eq_zero_iff], }, }, - { simpa [nat.bit, nat.bit1_val n, add_comm, digits_add 2 le_rfl 1 n (by norm_num) - (by norm_num)] }, + { simpa [nat.bit, nat.bit1_val n, add_comm, digits_add 2 one_lt_two 1 n] }, end @@ -629,8 +601,8 @@ theorem digits_succ (b n m r l) (e : r + b * m = n) (hr : r < b) - (h : nat.digits b m = l ∧ 2 ≤ b ∧ 0 < m) : - nat.digits b n = r :: l ∧ 2 ≤ b ∧ 0 < n := + (h : nat.digits b m = l ∧ 1 < b ∧ 0 < m) : + nat.digits b n = r :: l ∧ 1 < b ∧ 0 < n := begin rcases h with ⟨h, b2, m0⟩, have b0 : 0 < b := by linarith, @@ -642,12 +614,12 @@ end theorem digits_one (b n) (n0 : 0 < n) (nb : n < b) : - nat.digits b n = [n] ∧ 2 ≤ b ∧ 0 < n := + nat.digits b n = [n] ∧ 1 < b ∧ 0 < n := begin - have b2 : 2 ≤ b := by linarith, + have b2 : 1 < b := by linarith, refine ⟨_, b2, n0⟩, rw [nat.digits_def' b2 n0, nat.mod_eq_of_lt nb, - (nat.div_eq_zero_iff (by linarith : 0 < b)).2 nb, nat.digits_zero], + (nat.div_eq_zero_iff ((zero_le n).trans_lt nb)).2 nb, nat.digits_zero], end open tactic @@ -685,11 +657,10 @@ example : nat.digits 10 123 = [3,2,1] := by norm_num if n = 0 then return (`([] : list ℕ), `(nat.digits_zero %%eb)) else if b = 0 then do ic ← mk_instance_cache `(ℕ), - (_, pn0) ← norm_num.prove_pos ic en, + (_, pn0) ← norm_num.prove_ne_zero' ic en, return (`([%%en] : list ℕ), `(@nat.digits_zero_succ' %%en %%pn0)) else if b = 1 then do ic ← mk_instance_cache `(ℕ), - (_, pn0) ← norm_num.prove_pos ic en, s ← simp_lemmas.add_simp simp_lemmas.mk `list.replicate, (rhs, p2, _) ← simplify s [] `(list.replicate %%en 1), p ← mk_eq_trans `(nat.digits_one %%en) p2, diff --git a/src/data/polynomial/basic.lean b/src/data/polynomial/basic.lean index 8b825846cbcb2..1bd981b8111b8 100644 --- a/src/data/polynomial/basic.lean +++ b/src/data/polynomial/basic.lean @@ -781,7 +781,7 @@ by { ext, rw [coeff_update_apply, coeff_erase] } lemma support_update (p : R[X]) (n : ℕ) (a : R) [decidable (a = 0)] : support (p.update n a) = if a = 0 then p.support.erase n else insert n p.support := -by { cases p, simp only [support, update, support_update], congr } +by { classical, cases p, simp only [support, update, support_update], congr } lemma support_update_zero (p : R[X]) (n : ℕ) : support (p.update n 0) = p.support.erase n := diff --git a/src/data/real/basic.lean b/src/data/real/basic.lean index 76aa36a70ac09..4009cb9258435 100644 --- a/src/data/real/basic.lean +++ b/src/data/real/basic.lean @@ -11,6 +11,9 @@ import data.real.cau_seq_completion /-! # Real numbers from Cauchy sequences +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines `ℝ` as the type of equivalence classes of Cauchy sequences of rational numbers. This choice is motivated by how easy it is to prove that `ℝ` is a commutative ring, by simply lifting everything to `ℚ`. diff --git a/src/data/real/conjugate_exponents.lean b/src/data/real/conjugate_exponents.lean index c11da7181c8e2..6661e83456482 100644 --- a/src/data/real/conjugate_exponents.lean +++ b/src/data/real/conjugate_exponents.lean @@ -112,4 +112,8 @@ lemma is_conjugate_exponent_conjugate_exponent {p : ℝ} (h : 1 < p) : p.is_conjugate_exponent (conjugate_exponent p) := (is_conjugate_exponent_iff h).2 rfl +lemma is_conjugate_exponent_one_div {a b : ℝ} (ha : 0 < a) (hb : 0 < b) (hab : a + b = 1) : + (1 / a).is_conjugate_exponent (1 / b) := +⟨by { rw [lt_div_iff ha, one_mul], linarith }, by { simp_rw one_div_one_div, exact hab }⟩ + end real diff --git a/src/data/real/sign.lean b/src/data/real/sign.lean index eca1c6d1892b1..34f156eb1119a 100644 --- a/src/data/real/sign.lean +++ b/src/data/real/sign.lean @@ -8,6 +8,9 @@ import data.real.basic /-! # Real sign function +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file introduces and contains some results about `real.sign` which maps negative real numbers to -1, positive real numbers to 1, and 0 to 0. diff --git a/src/data/real/sqrt.lean b/src/data/real/sqrt.lean index 2ba270b9a1395..d2cfa45030bf5 100644 --- a/src/data/real/sqrt.lean +++ b/src/data/real/sqrt.lean @@ -290,6 +290,9 @@ by rw [sqrt, real.to_nnreal_inv, nnreal.sqrt_inv, nnreal.coe_inv, sqrt] @[simp] theorem sqrt_div (hx : 0 ≤ x) (y : ℝ) : sqrt (x / y) = sqrt x / sqrt y := by rw [division_def, sqrt_mul hx, sqrt_inv, division_def] +@[simp] theorem sqrt_div' (x) {y : ℝ} (hy : 0 ≤ y) : sqrt (x / y) = sqrt x / sqrt y := +by rw [division_def, sqrt_mul' x (inv_nonneg.2 hy), sqrt_inv, division_def] + @[simp] theorem div_sqrt : x / sqrt x = sqrt x := begin cases le_or_lt x 0, diff --git a/src/data/sym/sym2.lean b/src/data/sym/sym2.lean index 153f56f273a38..f7287e12bd1cf 100644 --- a/src/data/sym/sym2.lean +++ b/src/data/sym/sym2.lean @@ -5,6 +5,7 @@ Authors: Kyle Miller -/ import data.finset.prod import data.sym.basic +import data.set_like.basic import tactic.linarith /-! @@ -199,24 +200,42 @@ end section membership -/-! ### Declarations about membership -/ +/-! ### Membership and set coercion -/ /-- This is a predicate that determines whether a given term is a member of a term of the symmetric square. From this point of view, the symmetric square is the subtype of cardinality-two multisets on `α`. -/ -def mem (x : α) (z : sym2 α) : Prop := +protected def mem (x : α) (z : sym2 α) : Prop := ∃ (y : α), z = ⟦(x, y)⟧ -instance : has_mem α (sym2 α) := ⟨mem⟩ +lemma mem_iff' {a b c : α} : sym2.mem a ⟦(b, c)⟧ ↔ a = b ∨ a = c := +{ mp := by { rintro ⟨_, h⟩, rw eq_iff at h, tidy }, + mpr := by { rintro (rfl|rfl), { exact ⟨_, rfl⟩ }, rw eq_swap, exact ⟨_, rfl⟩ } } + +instance : set_like (sym2 α) α := +{ coe := λ z, {x | z.mem x}, + coe_injective' := λ z z' h, begin + simp only [set.ext_iff, set.mem_set_of_eq] at h, + induction z using sym2.ind with x y, + induction z' using sym2.ind with x' y', + have hx := h x, have hy := h y, have hx' := h x', have hy' := h y', + simp only [mem_iff', eq_self_iff_true, or_true, iff_true, true_or, true_iff] at hx hy hx' hy', + cases hx; cases hy; cases hx'; cases hy'; subst_vars, + rw [sym2.eq_swap], + end } + +@[simp] lemma mem_iff_mem {x : α} {z : sym2 α} : sym2.mem x z ↔ x ∈ z := iff.rfl + +lemma mem_iff_exists {x : α} {z : sym2 α} : x ∈ z ↔ ∃ (y : α), z = ⟦(x, y)⟧ := iff.rfl + +@[ext] theorem ext {p q : sym2 α} (h : ∀ x, x ∈ p ↔ x ∈ q) : p = q := set_like.ext h lemma mem_mk_left (x y : α) : x ∈ ⟦(x, y)⟧ := ⟨y, rfl⟩ lemma mem_mk_right (x y : α) : y ∈ ⟦(x, y)⟧ := eq_swap.subst $ mem_mk_left y x -@[simp] lemma mem_iff {a b c : α} : a ∈ ⟦(b, c)⟧ ↔ a = b ∨ a = c := -{ mp := by { rintro ⟨_, h⟩, rw eq_iff at h, tidy }, - mpr := by { rintro ⟨_⟩; subst a, { apply mem_mk_left }, apply mem_mk_right } } +@[simp] lemma mem_iff {a b c : α} : a ∈ ⟦(b, c)⟧ ↔ a = b ∨ a = c := mem_iff' lemma out_fst_mem (e : sym2 α) : e.out.1 ∈ e := ⟨e.out.2, by rw [prod.mk.eta, e.out_eq]⟩ lemma out_snd_mem (e : sym2 α) : e.out.2 ∈ e := ⟨e.out.1, by rw [eq_swap, prod.mk.eta, e.out_eq]⟩ @@ -258,17 +277,6 @@ lemma eq_of_ne_mem {x y : α} {z z' : sym2 α} (h : x ≠ y) (h1 : x ∈ z) (h2 : y ∈ z) (h3 : x ∈ z') (h4 : y ∈ z') : z = z' := ((mem_and_mem_iff h).mp ⟨h1, h2⟩).trans ((mem_and_mem_iff h).mp ⟨h3, h4⟩).symm -@[ext] -protected lemma ext (z z' : sym2 α) (h : ∀ x, x ∈ z ↔ x ∈ z') : z = z' := -begin - induction z using sym2.ind with x y, - induction z' using sym2.ind with x' y', - have hx := h x, have hy := h y, have hx' := h x', have hy' := h y', - simp only [mem_iff, eq_self_iff_true, or_true, iff_true, true_or, true_iff] at hx hy hx' hy', - cases hx; cases hy; cases hx'; cases hy'; subst_vars, - simp only [sym2.eq_swap], -end - instance mem.decidable [decidable_eq α] (x : α) (z : sym2 α) : decidable (x ∈ z) := quotient.rec_on_subsingleton z (λ ⟨y₁, y₂⟩, decidable_of_iff' _ mem_iff) diff --git a/src/geometry/euclidean/angle/oriented/affine.lean b/src/geometry/euclidean/angle/oriented/affine.lean index c627fc75e2991..ccb652414fe37 100644 --- a/src/geometry/euclidean/angle/oriented/affine.lean +++ b/src/geometry/euclidean/angle/oriented/affine.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Myers -/ import analysis.convex.side -import geometry.euclidean.angle.oriented.basic +import geometry.euclidean.angle.oriented.rotation import geometry.euclidean.angle.unoriented.affine /-! diff --git a/src/geometry/euclidean/angle/oriented/basic.lean b/src/geometry/euclidean/angle/oriented/basic.lean index 8f1c052d56586..c4d739c33fcc8 100644 --- a/src/geometry/euclidean/angle/oriented/basic.lean +++ b/src/geometry/euclidean/angle/oriented/basic.lean @@ -16,8 +16,6 @@ This file defines oriented angles in real inner product spaces. * `orientation.oangle` is the oriented angle between two vectors with respect to an orientation. -* `orientation.rotation` is the rotation by an oriented angle with respect to an orientation. - ## Implementation notes The definitions here use the `real.angle` type, angles modulo `2 * π`. For some purposes, @@ -45,7 +43,6 @@ variables {V V' : Type*} [inner_product_space ℝ V] [inner_product_space ℝ V' variables [fact (finrank ℝ V = 2)] [fact (finrank ℝ V' = 2)] (o : orientation ℝ V (fin 2)) local notation `ω` := o.area_form -local notation `J` := o.right_angle_rotation /-- The oriented angle from `x` to `y`, modulo `2 * π`. If either vector is 0, this is 0. See `inner_product_geometry.angle` for the corresponding unoriented angle definition. -/ @@ -553,357 +550,6 @@ begin simp end -/-- Auxiliary construction to build a rotation by the oriented angle `θ`. -/ -def rotation_aux (θ : real.angle) : V →ₗᵢ[ℝ] V := -linear_map.isometry_of_inner - (real.angle.cos θ • linear_map.id - + real.angle.sin θ • ↑(linear_isometry_equiv.to_linear_equiv J)) - begin - intros x y, - simp only [is_R_or_C.conj_to_real, id.def, linear_map.smul_apply, linear_map.add_apply, - linear_map.id_coe, linear_equiv.coe_coe, linear_isometry_equiv.coe_to_linear_equiv, - orientation.area_form_right_angle_rotation_left, - orientation.inner_right_angle_rotation_left, - orientation.inner_right_angle_rotation_right, - inner_add_left, inner_smul_left, inner_add_right, inner_smul_right], - linear_combination inner x y * θ.cos_sq_add_sin_sq, - end - -@[simp] lemma rotation_aux_apply (θ : real.angle) (x : V) : - o.rotation_aux θ x = real.angle.cos θ • x + real.angle.sin θ • J x := -rfl - -/-- A rotation by the oriented angle `θ`. -/ -def rotation (θ : real.angle) : V ≃ₗᵢ[ℝ] V := -linear_isometry_equiv.of_linear_isometry - (o.rotation_aux θ) - (real.angle.cos θ • linear_map.id - real.angle.sin θ • ↑(linear_isometry_equiv.to_linear_equiv J)) - begin - ext x, - convert congr_arg (λ t : ℝ, t • x) θ.cos_sq_add_sin_sq using 1, - { simp only [o.right_angle_rotation_right_angle_rotation, o.rotation_aux_apply, - function.comp_app, id.def, linear_equiv.coe_coe, linear_isometry.coe_to_linear_map, - linear_isometry_equiv.coe_to_linear_equiv, map_smul, map_sub, linear_map.coe_comp, - linear_map.id_coe, linear_map.smul_apply, linear_map.sub_apply, ← mul_smul, add_smul, - smul_add, smul_neg, smul_sub, mul_comm, sq], - abel }, - { simp }, - end - begin - ext x, - convert congr_arg (λ t : ℝ, t • x) θ.cos_sq_add_sin_sq using 1, - { simp only [o.right_angle_rotation_right_angle_rotation, o.rotation_aux_apply, - function.comp_app, id.def, linear_equiv.coe_coe, linear_isometry.coe_to_linear_map, - linear_isometry_equiv.coe_to_linear_equiv, map_add, map_smul, linear_map.coe_comp, - linear_map.id_coe, linear_map.smul_apply, linear_map.sub_apply, add_smul, ← mul_smul, - mul_comm, smul_add, smul_neg, sq], - abel }, - { simp }, - end - -lemma rotation_apply (θ : real.angle) (x : V) : - o.rotation θ x = real.angle.cos θ • x + real.angle.sin θ • J x := -rfl - -lemma rotation_symm_apply (θ : real.angle) (x : V) : - (o.rotation θ).symm x = real.angle.cos θ • x - real.angle.sin θ • J x := -rfl - -attribute [irreducible] rotation - -lemma rotation_eq_matrix_to_lin (θ : real.angle) {x : V} (hx : x ≠ 0) : - (o.rotation θ).to_linear_map - = matrix.to_lin - (o.basis_right_angle_rotation x hx) (o.basis_right_angle_rotation x hx) - !![θ.cos, -θ.sin; θ.sin, θ.cos] := -begin - apply (o.basis_right_angle_rotation x hx).ext, - intros i, - fin_cases i, - { rw matrix.to_lin_self, - simp [rotation_apply, fin.sum_univ_succ] }, - { rw matrix.to_lin_self, - simp [rotation_apply, fin.sum_univ_succ, add_comm] }, -end - -/-- The determinant of `rotation` (as a linear map) is equal to `1`. -/ -@[simp] lemma det_rotation (θ : real.angle) : - (o.rotation θ).to_linear_map.det = 1 := -begin - haveI : nontrivial V := - finite_dimensional.nontrivial_of_finrank_eq_succ (fact.out (finrank ℝ V = 2)), - obtain ⟨x, hx⟩ : ∃ x, x ≠ (0:V) := exists_ne (0:V), - rw o.rotation_eq_matrix_to_lin θ hx, - simpa [sq] using θ.cos_sq_add_sin_sq, -end - -/-- The determinant of `rotation` (as a linear equiv) is equal to `1`. -/ -@[simp] lemma linear_equiv_det_rotation (θ : real.angle) : - (o.rotation θ).to_linear_equiv.det = 1 := -units.ext $ o.det_rotation θ - -/-- The inverse of `rotation` is rotation by the negation of the angle. -/ -@[simp] lemma rotation_symm (θ : real.angle) : (o.rotation θ).symm = o.rotation (-θ) := -by ext; simp [o.rotation_apply, o.rotation_symm_apply, sub_eq_add_neg] - -/-- Rotation by 0 is the identity. -/ -@[simp] lemma rotation_zero : o.rotation 0 = linear_isometry_equiv.refl ℝ V := -by ext; simp [rotation] - -/-- Rotation by π is negation. -/ -@[simp] lemma rotation_pi : o.rotation π = linear_isometry_equiv.neg ℝ := -begin - ext x, - simp [rotation] -end - -/-- Rotation by π is negation. -/ -lemma rotation_pi_apply (x : V) : o.rotation π x = -x := -by simp - -/-- Rotation by π / 2 is the "right-angle-rotation" map `J`. -/ -lemma rotation_pi_div_two : o.rotation (π / 2 : ℝ) = J := -begin - ext x, - simp [rotation], -end - -/-- Rotating twice is equivalent to rotating by the sum of the angles. -/ -@[simp] lemma rotation_rotation (θ₁ θ₂ : real.angle) (x : V) : - o.rotation θ₁ (o.rotation θ₂ x) = o.rotation (θ₁ + θ₂) x := -begin - simp only [o.rotation_apply, ←mul_smul, real.angle.cos_add, real.angle.sin_add, add_smul, - sub_smul, linear_isometry_equiv.trans_apply, smul_add, linear_isometry_equiv.map_add, - linear_isometry_equiv.map_smul, right_angle_rotation_right_angle_rotation, smul_neg], - ring_nf, - abel, -end - -/-- Rotating twice is equivalent to rotating by the sum of the angles. -/ -@[simp] lemma rotation_trans (θ₁ θ₂ : real.angle) : - (o.rotation θ₁).trans (o.rotation θ₂) = o.rotation (θ₂ + θ₁) := -linear_isometry_equiv.ext $ λ _, by rw [←rotation_rotation, linear_isometry_equiv.trans_apply] - -/-- Rotating the first of two vectors by `θ` scales their Kahler form by `cos θ - sin θ * I`. -/ -@[simp] lemma kahler_rotation_left (x y : V) (θ : real.angle) : - o.kahler (o.rotation θ x) y = conj (θ.exp_map_circle : ℂ) * o.kahler x y := -begin - simp only [o.rotation_apply, map_add, map_mul, linear_map.map_smulₛₗ, ring_hom.id_apply, - linear_map.add_apply, linear_map.smul_apply, real_smul, kahler_right_angle_rotation_left, - real.angle.coe_exp_map_circle, is_R_or_C.conj_of_real, conj_I], - ring, -end - -/-- Negating a rotation is equivalent to rotation by π plus the angle. -/ -lemma neg_rotation (θ : real.angle) (x : V) : -o.rotation θ x = o.rotation (π + θ) x := -by rw [←o.rotation_pi_apply, rotation_rotation] - -/-- Negating a rotation by -π / 2 is equivalent to rotation by π / 2. -/ -@[simp] lemma neg_rotation_neg_pi_div_two (x : V) : - -o.rotation (-π / 2 : ℝ) x = o.rotation (π / 2 : ℝ) x := -by rw [neg_rotation, ←real.angle.coe_add, neg_div, ←sub_eq_add_neg, sub_half] - -/-- Negating a rotation by π / 2 is equivalent to rotation by -π / 2. -/ -lemma neg_rotation_pi_div_two (x : V) : -o.rotation (π / 2 : ℝ) x = o.rotation (-π / 2 : ℝ) x := -neg_eq_iff_neg_eq.1 $ o.neg_rotation_neg_pi_div_two _ - -/-- Rotating the first of two vectors by `θ` scales their Kahler form by `cos (-θ) + sin (-θ) * I`. --/ -lemma kahler_rotation_left' (x y : V) (θ : real.angle) : - o.kahler (o.rotation θ x) y = (-θ).exp_map_circle * o.kahler x y := -by simpa [coe_inv_circle_eq_conj, -kahler_rotation_left] using o.kahler_rotation_left x y θ - -/-- Rotating the second of two vectors by `θ` scales their Kahler form by `cos θ + sin θ * I`. -/ -@[simp] lemma kahler_rotation_right (x y : V) (θ : real.angle) : - o.kahler x (o.rotation θ y) = θ.exp_map_circle * o.kahler x y := -begin - simp only [o.rotation_apply, map_add, linear_map.map_smulₛₗ, ring_hom.id_apply, real_smul, - kahler_right_angle_rotation_right, real.angle.coe_exp_map_circle], - ring, -end - -/-- Rotating the first vector by `θ` subtracts `θ` from the angle between two vectors. -/ -@[simp] lemma oangle_rotation_left {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) (θ : real.angle) : - o.oangle (o.rotation θ x) y = o.oangle x y - θ := -begin - simp only [oangle, o.kahler_rotation_left'], - rw [complex.arg_mul_coe_angle, real.angle.arg_exp_map_circle], - { abel }, - { exact ne_zero_of_mem_circle _ }, - { exact o.kahler_ne_zero hx hy }, -end - -/-- Rotating the second vector by `θ` adds `θ` to the angle between two vectors. -/ -@[simp] lemma oangle_rotation_right {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) (θ : real.angle) : - o.oangle x (o.rotation θ y) = o.oangle x y + θ := -begin - simp only [oangle, o.kahler_rotation_right], - rw [complex.arg_mul_coe_angle, real.angle.arg_exp_map_circle], - { abel }, - { exact ne_zero_of_mem_circle _ }, - { exact o.kahler_ne_zero hx hy }, -end - -/-- The rotation of a vector by `θ` has an angle of `-θ` from that vector. -/ -@[simp] lemma oangle_rotation_self_left {x : V} (hx : x ≠ 0) (θ : real.angle) : - o.oangle (o.rotation θ x) x = -θ := -by simp [hx] - -/-- A vector has an angle of `θ` from the rotation of that vector by `θ`. -/ -@[simp] lemma oangle_rotation_self_right {x : V} (hx : x ≠ 0) (θ : real.angle) : - o.oangle x (o.rotation θ x) = θ := -by simp [hx] - -/-- Rotating the first vector by the angle between the two vectors results an an angle of 0. -/ -@[simp] lemma oangle_rotation_oangle_left (x y : V) : - o.oangle (o.rotation (o.oangle x y) x) y = 0 := -begin - by_cases hx : x = 0, - { simp [hx] }, - { by_cases hy : y = 0, - { simp [hy] }, - { simp [hx, hy] } } -end - -/-- Rotating the first vector by the angle between the two vectors and swapping the vectors -results an an angle of 0. -/ -@[simp] lemma oangle_rotation_oangle_right (x y : V) : - o.oangle y (o.rotation (o.oangle x y) x) = 0 := -begin - rw [oangle_rev], - simp -end - -/-- Rotating both vectors by the same angle does not change the angle between those vectors. -/ -@[simp] lemma oangle_rotation (x y : V) (θ : real.angle) : - o.oangle (o.rotation θ x) (o.rotation θ y) = o.oangle x y := -begin - by_cases hx : x = 0; by_cases hy : y = 0; - simp [hx, hy] -end - -/-- A rotation of a nonzero vector equals that vector if and only if the angle is zero. -/ -@[simp] lemma rotation_eq_self_iff_angle_eq_zero {x : V} (hx : x ≠ 0) (θ : real.angle) : - o.rotation θ x = x ↔ θ = 0 := -begin - split, - { intro h, - rw eq_comm, - simpa [hx, h] using o.oangle_rotation_right hx hx θ }, - { intro h, - simp [h] } -end - -/-- A nonzero vector equals a rotation of that vector if and only if the angle is zero. -/ -@[simp] lemma eq_rotation_self_iff_angle_eq_zero {x : V} (hx : x ≠ 0) (θ : real.angle) : - x = o.rotation θ x ↔ θ = 0 := -by rw [←o.rotation_eq_self_iff_angle_eq_zero hx, eq_comm] - -/-- A rotation of a vector equals that vector if and only if the vector or the angle is zero. -/ -lemma rotation_eq_self_iff (x : V) (θ : real.angle) : - o.rotation θ x = x ↔ x = 0 ∨ θ = 0 := -begin - by_cases h : x = 0; - simp [h] -end - -/-- A vector equals a rotation of that vector if and only if the vector or the angle is zero. -/ -lemma eq_rotation_self_iff (x : V) (θ : real.angle) : - x = o.rotation θ x ↔ x = 0 ∨ θ = 0 := -by rw [←rotation_eq_self_iff, eq_comm] - -/-- Rotating a vector by the angle to another vector gives the second vector if and only if the -norms are equal. -/ -@[simp] lemma rotation_oangle_eq_iff_norm_eq (x y : V) : - o.rotation (o.oangle x y) x = y ↔ ‖x‖ = ‖y‖ := -begin - split, - { intro h, - rw [←h, linear_isometry_equiv.norm_map] }, - { intro h, - rw o.eq_iff_oangle_eq_zero_of_norm_eq; - simp [h] } -end - -/-- The angle between two nonzero vectors is `θ` if and only if the second vector is the first -rotated by `θ` and scaled by the ratio of the norms. -/ -lemma oangle_eq_iff_eq_norm_div_norm_smul_rotation_of_ne_zero {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) - (θ : real.angle) : o.oangle x y = θ ↔ y = (‖y‖ / ‖x‖) • o.rotation θ x := -begin - have hp := div_pos (norm_pos_iff.2 hy) (norm_pos_iff.2 hx), - split, - { rintro rfl, - rw [←linear_isometry_equiv.map_smul, ←o.oangle_smul_left_of_pos x y hp, - eq_comm, rotation_oangle_eq_iff_norm_eq, norm_smul, real.norm_of_nonneg hp.le, - div_mul_cancel _ (norm_ne_zero_iff.2 hx)] }, - { intro hye, - rw [hye, o.oangle_smul_right_of_pos _ _ hp, o.oangle_rotation_self_right hx] } -end - -/-- The angle between two nonzero vectors is `θ` if and only if the second vector is the first -rotated by `θ` and scaled by a positive real. -/ -lemma oangle_eq_iff_eq_pos_smul_rotation_of_ne_zero {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) - (θ : real.angle) : o.oangle x y = θ ↔ ∃ r : ℝ, 0 < r ∧ y = r • o.rotation θ x := -begin - split, - { intro h, - rw o.oangle_eq_iff_eq_norm_div_norm_smul_rotation_of_ne_zero hx hy at h, - exact ⟨‖y‖ / ‖x‖, div_pos (norm_pos_iff.2 hy) (norm_pos_iff.2 hx), h⟩ }, - { rintro ⟨r, hr, rfl⟩, - rw [o.oangle_smul_right_of_pos _ _ hr, o.oangle_rotation_self_right hx] } -end - -/-- The angle between two vectors is `θ` if and only if they are nonzero and the second vector -is the first rotated by `θ` and scaled by the ratio of the norms, or `θ` and at least one of the -vectors are zero. -/ -lemma oangle_eq_iff_eq_norm_div_norm_smul_rotation_or_eq_zero {x y : V} (θ : real.angle) : - o.oangle x y = θ ↔ - (x ≠ 0 ∧ y ≠ 0 ∧ y = (‖y‖ / ‖x‖) • o.rotation θ x) ∨ (θ = 0 ∧ (x = 0 ∨ y = 0)) := -begin - by_cases hx : x = 0, - { simp [hx, eq_comm] }, - { by_cases hy : y = 0, - { simp [hy, eq_comm] }, - { rw o.oangle_eq_iff_eq_norm_div_norm_smul_rotation_of_ne_zero hx hy, - simp [hx, hy] } } -end - -/-- The angle between two vectors is `θ` if and only if they are nonzero and the second vector -is the first rotated by `θ` and scaled by a positive real, or `θ` and at least one of the -vectors are zero. -/ -lemma oangle_eq_iff_eq_pos_smul_rotation_or_eq_zero {x y : V} (θ : real.angle) : - o.oangle x y = θ ↔ - (x ≠ 0 ∧ y ≠ 0 ∧ ∃ r : ℝ, 0 < r ∧ y = r • o.rotation θ x) ∨ (θ = 0 ∧ (x = 0 ∨ y = 0)) := -begin - by_cases hx : x = 0, - { simp [hx, eq_comm] }, - { by_cases hy : y = 0, - { simp [hy, eq_comm] }, - { rw o.oangle_eq_iff_eq_pos_smul_rotation_of_ne_zero hx hy, - simp [hx, hy] } } -end - -/-- Any linear isometric equivalence in `V` with positive determinant is `rotation`. -/ -lemma exists_linear_isometry_equiv_eq_of_det_pos {f : V ≃ₗᵢ[ℝ] V} - (hd : 0 < (f.to_linear_equiv : V →ₗ[ℝ] V).det) : ∃ θ : real.angle, f = o.rotation θ := -begin - haveI : nontrivial V := - finite_dimensional.nontrivial_of_finrank_eq_succ (fact.out (finrank ℝ V = 2)), - obtain ⟨x, hx⟩ : ∃ x, x ≠ (0:V) := exists_ne (0:V), - use o.oangle x (f x), - apply linear_isometry_equiv.to_linear_equiv_injective, - apply linear_equiv.to_linear_map_injective, - apply (o.basis_right_angle_rotation x hx).ext, - intros i, - symmetry, - fin_cases i, - { simp }, - have : o.oangle (J x) (f (J x)) = o.oangle x (f x), - { simp only [oangle, o.linear_isometry_equiv_comp_right_angle_rotation f hd, - o.kahler_comp_right_angle_rotation] }, - simp [← this], -end - /-- The angle between two vectors, with respect to an orientation given by `orientation.map` with a linear isometric equivalence, equals the angle between those two vectors, transformed by the inverse of that equivalence, with respect to the original orientation. -/ @@ -929,34 +575,6 @@ end lemma oangle_neg_orientation_eq_neg (x y : V) : (-o).oangle x y = -(o.oangle x y) := by simp [oangle] -lemma rotation_map (θ : real.angle) (f : V ≃ₗᵢ[ℝ] V') (x : V') : - (orientation.map (fin 2) f.to_linear_equiv o).rotation θ x - = f (o.rotation θ (f.symm x)) := -by simp [rotation_apply, o.right_angle_rotation_map] - -@[simp] protected lemma _root_.complex.rotation (θ : real.angle) (z : ℂ) : - complex.orientation.rotation θ z = θ.exp_map_circle * z := -begin - simp only [rotation_apply, complex.right_angle_rotation, real.angle.coe_exp_map_circle, - real_smul], - ring -end - -/-- Rotation in an oriented real inner product space of dimension 2 can be evaluated in terms of a -complex-number representation of the space. -/ -lemma rotation_map_complex (θ : real.angle) (f : V ≃ₗᵢ[ℝ] ℂ) - (hf : (orientation.map (fin 2) f.to_linear_equiv o) = complex.orientation) (x : V) : - f (o.rotation θ x) = θ.exp_map_circle * f x := -begin - rw [← complex.rotation, ← hf, o.rotation_map], - simp, -end - -/-- Negating the orientation negates the angle in `rotation`. -/ -lemma rotation_neg_orientation_eq_neg (θ : real.angle) : - (-o).rotation θ = o.rotation (-θ) := -linear_isometry_equiv.ext $ by simp [rotation_apply] - /-- The inner product of two vectors is the product of the norms and the cosine of the oriented angle between the vectors. -/ lemma inner_eq_norm_mul_norm_mul_cos_oangle (x y : V) : @@ -1168,81 +786,6 @@ lemma inner_rev_eq_zero_of_oangle_eq_neg_pi_div_two {x y : V} (h : o.oangle x y ⟪y, x⟫ = 0 := by rw [real_inner_comm, o.inner_eq_zero_of_oangle_eq_neg_pi_div_two h] -/-- The inner product between a `π / 2` rotation of a vector and that vector is zero. -/ -@[simp] lemma inner_rotation_pi_div_two_left (x : V) : ⟪o.rotation (π / 2 : ℝ) x, x⟫ = 0 := -by rw [rotation_pi_div_two, inner_right_angle_rotation_self] - -/-- The inner product between a vector and a `π / 2` rotation of that vector is zero. -/ -@[simp] lemma inner_rotation_pi_div_two_right (x : V) : ⟪x, o.rotation (π / 2 : ℝ) x⟫ = 0 := -by rw [real_inner_comm, inner_rotation_pi_div_two_left] - -/-- The inner product between a multiple of a `π / 2` rotation of a vector and that vector is -zero. -/ -@[simp] lemma inner_smul_rotation_pi_div_two_left (x : V) (r : ℝ) : - ⟪r • o.rotation (π / 2 : ℝ) x, x⟫ = 0 := -by rw [inner_smul_left, inner_rotation_pi_div_two_left, mul_zero] - -/-- The inner product between a vector and a multiple of a `π / 2` rotation of that vector is -zero. -/ -@[simp] lemma inner_smul_rotation_pi_div_two_right (x : V) (r : ℝ) : - ⟪x, r • o.rotation (π / 2 : ℝ) x⟫ = 0 := -by rw [real_inner_comm, inner_smul_rotation_pi_div_two_left] - -/-- The inner product between a `π / 2` rotation of a vector and a multiple of that vector is -zero. -/ -@[simp] lemma inner_rotation_pi_div_two_left_smul (x : V) (r : ℝ) : - ⟪o.rotation (π / 2 : ℝ) x, r • x⟫ = 0 := -by rw [inner_smul_right, inner_rotation_pi_div_two_left, mul_zero] - -/-- The inner product between a multiple of a vector and a `π / 2` rotation of that vector is -zero. -/ -@[simp] lemma inner_rotation_pi_div_two_right_smul (x : V) (r : ℝ) : - ⟪r • x, o.rotation (π / 2 : ℝ) x⟫ = 0 := -by rw [real_inner_comm, inner_rotation_pi_div_two_left_smul] - -/-- The inner product between a multiple of a `π / 2` rotation of a vector and a multiple of -that vector is zero. -/ -@[simp] lemma inner_smul_rotation_pi_div_two_smul_left (x : V) (r₁ r₂ : ℝ) : - ⟪r₁ • o.rotation (π / 2 : ℝ) x, r₂ • x⟫ = 0 := -by rw [inner_smul_right, inner_smul_rotation_pi_div_two_left, mul_zero] - -/-- The inner product between a multiple of a vector and a multiple of a `π / 2` rotation of -that vector is zero. -/ -@[simp] lemma inner_smul_rotation_pi_div_two_smul_right (x : V) (r₁ r₂ : ℝ) : - ⟪r₂ • x, r₁ • o.rotation (π / 2 : ℝ) x⟫ = 0 := -by rw [real_inner_comm, inner_smul_rotation_pi_div_two_smul_left] - -/-- The inner product between two vectors is zero if and only if the first vector is zero or -the second is a multiple of a `π / 2` rotation of that vector. -/ -lemma inner_eq_zero_iff_eq_zero_or_eq_smul_rotation_pi_div_two {x y : V} : - ⟪x, y⟫ = 0 ↔ (x = 0 ∨ ∃ r : ℝ, r • o.rotation (π / 2 : ℝ) x = y) := -begin - rw ←o.eq_zero_or_oangle_eq_iff_inner_eq_zero, - refine ⟨λ h, _, λ h, _⟩, - { rcases h with rfl | rfl | h | h, - { exact or.inl rfl }, - { exact or.inr ⟨0, zero_smul _ _⟩ }, - { obtain ⟨r, hr, rfl⟩ := (o.oangle_eq_iff_eq_pos_smul_rotation_of_ne_zero - (o.left_ne_zero_of_oangle_eq_pi_div_two h) - (o.right_ne_zero_of_oangle_eq_pi_div_two h) _).1 h, - exact or.inr ⟨r, rfl⟩ }, - { obtain ⟨r, hr, rfl⟩ := (o.oangle_eq_iff_eq_pos_smul_rotation_of_ne_zero - (o.left_ne_zero_of_oangle_eq_neg_pi_div_two h) - (o.right_ne_zero_of_oangle_eq_neg_pi_div_two h) _).1 h, - refine or.inr ⟨-r, _⟩, - rw [neg_smul, ←smul_neg, o.neg_rotation_pi_div_two] } }, - { rcases h with rfl | ⟨r, rfl⟩, - { exact or.inl rfl }, - { by_cases hx : x = 0, { exact or.inl hx }, - rcases lt_trichotomy r 0 with hr | rfl | hr, - { refine or.inr (or.inr (or.inr _)), - rw [o.oangle_smul_right_of_neg _ _ hr, o.neg_rotation_pi_div_two, - o.oangle_rotation_self_right hx] }, - { exact or.inr (or.inl (zero_smul _ _)) }, - { refine or.inr (or.inr (or.inl _)), - rw [o.oangle_smul_right_of_pos _ _ hr, o.oangle_rotation_self_right hx] } } } -end - /-- Negating the first vector passed to `oangle` negates the sign of the angle. -/ @[simp] lemma oangle_sign_neg_left (x y : V) : (o.oangle (-x) y).sign = -((o.oangle x y).sign) := diff --git a/src/geometry/euclidean/angle/oriented/rotation.lean b/src/geometry/euclidean/angle/oriented/rotation.lean new file mode 100644 index 0000000000000..588502241a5e2 --- /dev/null +++ b/src/geometry/euclidean/angle/oriented/rotation.lean @@ -0,0 +1,488 @@ +/- +Copyright (c) 2022 Joseph Myers. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Joseph Myers, Heather Macbeth +-/ +import geometry.euclidean.angle.oriented.basic + +/-! +# Rotations by oriented angles. + +This file defines rotations by oriented angles in real inner product spaces. + +## Main definitions + +* `orientation.rotation` is the rotation by an oriented angle with respect to an orientation. + +-/ + +noncomputable theory + +open finite_dimensional complex +open_locale real real_inner_product_space complex_conjugate + +namespace orientation + +local attribute [instance] fact_finite_dimensional_of_finrank_eq_succ +local attribute [instance] complex.finrank_real_complex_fact + +variables {V V' : Type*} [inner_product_space ℝ V] [inner_product_space ℝ V'] +variables [fact (finrank ℝ V = 2)] [fact (finrank ℝ V' = 2)] (o : orientation ℝ V (fin 2)) + +local notation `J` := o.right_angle_rotation + +/-- Auxiliary construction to build a rotation by the oriented angle `θ`. -/ +def rotation_aux (θ : real.angle) : V →ₗᵢ[ℝ] V := +linear_map.isometry_of_inner + (real.angle.cos θ • linear_map.id + + real.angle.sin θ • ↑(linear_isometry_equiv.to_linear_equiv J)) + begin + intros x y, + simp only [is_R_or_C.conj_to_real, id.def, linear_map.smul_apply, linear_map.add_apply, + linear_map.id_coe, linear_equiv.coe_coe, linear_isometry_equiv.coe_to_linear_equiv, + orientation.area_form_right_angle_rotation_left, + orientation.inner_right_angle_rotation_left, + orientation.inner_right_angle_rotation_right, + inner_add_left, inner_smul_left, inner_add_right, inner_smul_right], + linear_combination inner x y * θ.cos_sq_add_sin_sq, + end + +@[simp] lemma rotation_aux_apply (θ : real.angle) (x : V) : + o.rotation_aux θ x = real.angle.cos θ • x + real.angle.sin θ • J x := +rfl + +/-- A rotation by the oriented angle `θ`. -/ +def rotation (θ : real.angle) : V ≃ₗᵢ[ℝ] V := +linear_isometry_equiv.of_linear_isometry + (o.rotation_aux θ) + (real.angle.cos θ • linear_map.id - real.angle.sin θ • ↑(linear_isometry_equiv.to_linear_equiv J)) + begin + ext x, + convert congr_arg (λ t : ℝ, t • x) θ.cos_sq_add_sin_sq using 1, + { simp only [o.right_angle_rotation_right_angle_rotation, o.rotation_aux_apply, + function.comp_app, id.def, linear_equiv.coe_coe, linear_isometry.coe_to_linear_map, + linear_isometry_equiv.coe_to_linear_equiv, map_smul, map_sub, linear_map.coe_comp, + linear_map.id_coe, linear_map.smul_apply, linear_map.sub_apply, ← mul_smul, add_smul, + smul_add, smul_neg, smul_sub, mul_comm, sq], + abel }, + { simp }, + end + begin + ext x, + convert congr_arg (λ t : ℝ, t • x) θ.cos_sq_add_sin_sq using 1, + { simp only [o.right_angle_rotation_right_angle_rotation, o.rotation_aux_apply, + function.comp_app, id.def, linear_equiv.coe_coe, linear_isometry.coe_to_linear_map, + linear_isometry_equiv.coe_to_linear_equiv, map_add, map_smul, linear_map.coe_comp, + linear_map.id_coe, linear_map.smul_apply, linear_map.sub_apply, add_smul, ← mul_smul, + mul_comm, smul_add, smul_neg, sq], + abel }, + { simp }, + end + +lemma rotation_apply (θ : real.angle) (x : V) : + o.rotation θ x = real.angle.cos θ • x + real.angle.sin θ • J x := +rfl + +lemma rotation_symm_apply (θ : real.angle) (x : V) : + (o.rotation θ).symm x = real.angle.cos θ • x - real.angle.sin θ • J x := +rfl + +attribute [irreducible] rotation + +lemma rotation_eq_matrix_to_lin (θ : real.angle) {x : V} (hx : x ≠ 0) : + (o.rotation θ).to_linear_map + = matrix.to_lin + (o.basis_right_angle_rotation x hx) (o.basis_right_angle_rotation x hx) + !![θ.cos, -θ.sin; θ.sin, θ.cos] := +begin + apply (o.basis_right_angle_rotation x hx).ext, + intros i, + fin_cases i, + { rw matrix.to_lin_self, + simp [rotation_apply, fin.sum_univ_succ] }, + { rw matrix.to_lin_self, + simp [rotation_apply, fin.sum_univ_succ, add_comm] }, +end + +/-- The determinant of `rotation` (as a linear map) is equal to `1`. -/ +@[simp] lemma det_rotation (θ : real.angle) : + (o.rotation θ).to_linear_map.det = 1 := +begin + haveI : nontrivial V := + finite_dimensional.nontrivial_of_finrank_eq_succ (fact.out (finrank ℝ V = 2)), + obtain ⟨x, hx⟩ : ∃ x, x ≠ (0:V) := exists_ne (0:V), + rw o.rotation_eq_matrix_to_lin θ hx, + simpa [sq] using θ.cos_sq_add_sin_sq, +end + +/-- The determinant of `rotation` (as a linear equiv) is equal to `1`. -/ +@[simp] lemma linear_equiv_det_rotation (θ : real.angle) : + (o.rotation θ).to_linear_equiv.det = 1 := +units.ext $ o.det_rotation θ + +/-- The inverse of `rotation` is rotation by the negation of the angle. -/ +@[simp] lemma rotation_symm (θ : real.angle) : (o.rotation θ).symm = o.rotation (-θ) := +by ext; simp [o.rotation_apply, o.rotation_symm_apply, sub_eq_add_neg] + +/-- Rotation by 0 is the identity. -/ +@[simp] lemma rotation_zero : o.rotation 0 = linear_isometry_equiv.refl ℝ V := +by ext; simp [rotation] + +/-- Rotation by π is negation. -/ +@[simp] lemma rotation_pi : o.rotation π = linear_isometry_equiv.neg ℝ := +begin + ext x, + simp [rotation] +end + +/-- Rotation by π is negation. -/ +lemma rotation_pi_apply (x : V) : o.rotation π x = -x := +by simp + +/-- Rotation by π / 2 is the "right-angle-rotation" map `J`. -/ +lemma rotation_pi_div_two : o.rotation (π / 2 : ℝ) = J := +begin + ext x, + simp [rotation], +end + +/-- Rotating twice is equivalent to rotating by the sum of the angles. -/ +@[simp] lemma rotation_rotation (θ₁ θ₂ : real.angle) (x : V) : + o.rotation θ₁ (o.rotation θ₂ x) = o.rotation (θ₁ + θ₂) x := +begin + simp only [o.rotation_apply, ←mul_smul, real.angle.cos_add, real.angle.sin_add, add_smul, + sub_smul, linear_isometry_equiv.trans_apply, smul_add, linear_isometry_equiv.map_add, + linear_isometry_equiv.map_smul, right_angle_rotation_right_angle_rotation, smul_neg], + ring_nf, + abel, +end + +/-- Rotating twice is equivalent to rotating by the sum of the angles. -/ +@[simp] lemma rotation_trans (θ₁ θ₂ : real.angle) : + (o.rotation θ₁).trans (o.rotation θ₂) = o.rotation (θ₂ + θ₁) := +linear_isometry_equiv.ext $ λ _, by rw [←rotation_rotation, linear_isometry_equiv.trans_apply] + +/-- Rotating the first of two vectors by `θ` scales their Kahler form by `cos θ - sin θ * I`. -/ +@[simp] lemma kahler_rotation_left (x y : V) (θ : real.angle) : + o.kahler (o.rotation θ x) y = conj (θ.exp_map_circle : ℂ) * o.kahler x y := +begin + simp only [o.rotation_apply, map_add, map_mul, linear_map.map_smulₛₗ, ring_hom.id_apply, + linear_map.add_apply, linear_map.smul_apply, real_smul, kahler_right_angle_rotation_left, + real.angle.coe_exp_map_circle, is_R_or_C.conj_of_real, conj_I], + ring, +end + +/-- Negating a rotation is equivalent to rotation by π plus the angle. -/ +lemma neg_rotation (θ : real.angle) (x : V) : -o.rotation θ x = o.rotation (π + θ) x := +by rw [←o.rotation_pi_apply, rotation_rotation] + +/-- Negating a rotation by -π / 2 is equivalent to rotation by π / 2. -/ +@[simp] lemma neg_rotation_neg_pi_div_two (x : V) : + -o.rotation (-π / 2 : ℝ) x = o.rotation (π / 2 : ℝ) x := +by rw [neg_rotation, ←real.angle.coe_add, neg_div, ←sub_eq_add_neg, sub_half] + +/-- Negating a rotation by π / 2 is equivalent to rotation by -π / 2. -/ +lemma neg_rotation_pi_div_two (x : V) : -o.rotation (π / 2 : ℝ) x = o.rotation (-π / 2 : ℝ) x := +neg_eq_iff_neg_eq.1 $ o.neg_rotation_neg_pi_div_two _ + +/-- Rotating the first of two vectors by `θ` scales their Kahler form by `cos (-θ) + sin (-θ) * I`. +-/ +lemma kahler_rotation_left' (x y : V) (θ : real.angle) : + o.kahler (o.rotation θ x) y = (-θ).exp_map_circle * o.kahler x y := +by simpa [coe_inv_circle_eq_conj, -kahler_rotation_left] using o.kahler_rotation_left x y θ + +/-- Rotating the second of two vectors by `θ` scales their Kahler form by `cos θ + sin θ * I`. -/ +@[simp] lemma kahler_rotation_right (x y : V) (θ : real.angle) : + o.kahler x (o.rotation θ y) = θ.exp_map_circle * o.kahler x y := +begin + simp only [o.rotation_apply, map_add, linear_map.map_smulₛₗ, ring_hom.id_apply, real_smul, + kahler_right_angle_rotation_right, real.angle.coe_exp_map_circle], + ring, +end + +/-- Rotating the first vector by `θ` subtracts `θ` from the angle between two vectors. -/ +@[simp] lemma oangle_rotation_left {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) (θ : real.angle) : + o.oangle (o.rotation θ x) y = o.oangle x y - θ := +begin + simp only [oangle, o.kahler_rotation_left'], + rw [complex.arg_mul_coe_angle, real.angle.arg_exp_map_circle], + { abel }, + { exact ne_zero_of_mem_circle _ }, + { exact o.kahler_ne_zero hx hy }, +end + +/-- Rotating the second vector by `θ` adds `θ` to the angle between two vectors. -/ +@[simp] lemma oangle_rotation_right {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) (θ : real.angle) : + o.oangle x (o.rotation θ y) = o.oangle x y + θ := +begin + simp only [oangle, o.kahler_rotation_right], + rw [complex.arg_mul_coe_angle, real.angle.arg_exp_map_circle], + { abel }, + { exact ne_zero_of_mem_circle _ }, + { exact o.kahler_ne_zero hx hy }, +end + +/-- The rotation of a vector by `θ` has an angle of `-θ` from that vector. -/ +@[simp] lemma oangle_rotation_self_left {x : V} (hx : x ≠ 0) (θ : real.angle) : + o.oangle (o.rotation θ x) x = -θ := +by simp [hx] + +/-- A vector has an angle of `θ` from the rotation of that vector by `θ`. -/ +@[simp] lemma oangle_rotation_self_right {x : V} (hx : x ≠ 0) (θ : real.angle) : + o.oangle x (o.rotation θ x) = θ := +by simp [hx] + +/-- Rotating the first vector by the angle between the two vectors results an an angle of 0. -/ +@[simp] lemma oangle_rotation_oangle_left (x y : V) : + o.oangle (o.rotation (o.oangle x y) x) y = 0 := +begin + by_cases hx : x = 0, + { simp [hx] }, + { by_cases hy : y = 0, + { simp [hy] }, + { simp [hx, hy] } } +end + +/-- Rotating the first vector by the angle between the two vectors and swapping the vectors +results an an angle of 0. -/ +@[simp] lemma oangle_rotation_oangle_right (x y : V) : + o.oangle y (o.rotation (o.oangle x y) x) = 0 := +begin + rw [oangle_rev], + simp +end + +/-- Rotating both vectors by the same angle does not change the angle between those vectors. -/ +@[simp] lemma oangle_rotation (x y : V) (θ : real.angle) : + o.oangle (o.rotation θ x) (o.rotation θ y) = o.oangle x y := +begin + by_cases hx : x = 0; by_cases hy : y = 0; + simp [hx, hy] +end + +/-- A rotation of a nonzero vector equals that vector if and only if the angle is zero. -/ +@[simp] lemma rotation_eq_self_iff_angle_eq_zero {x : V} (hx : x ≠ 0) (θ : real.angle) : + o.rotation θ x = x ↔ θ = 0 := +begin + split, + { intro h, + rw eq_comm, + simpa [hx, h] using o.oangle_rotation_right hx hx θ }, + { intro h, + simp [h] } +end + +/-- A nonzero vector equals a rotation of that vector if and only if the angle is zero. -/ +@[simp] lemma eq_rotation_self_iff_angle_eq_zero {x : V} (hx : x ≠ 0) (θ : real.angle) : + x = o.rotation θ x ↔ θ = 0 := +by rw [←o.rotation_eq_self_iff_angle_eq_zero hx, eq_comm] + +/-- A rotation of a vector equals that vector if and only if the vector or the angle is zero. -/ +lemma rotation_eq_self_iff (x : V) (θ : real.angle) : + o.rotation θ x = x ↔ x = 0 ∨ θ = 0 := +begin + by_cases h : x = 0; + simp [h] +end + +/-- A vector equals a rotation of that vector if and only if the vector or the angle is zero. -/ +lemma eq_rotation_self_iff (x : V) (θ : real.angle) : + x = o.rotation θ x ↔ x = 0 ∨ θ = 0 := +by rw [←rotation_eq_self_iff, eq_comm] + +/-- Rotating a vector by the angle to another vector gives the second vector if and only if the +norms are equal. -/ +@[simp] lemma rotation_oangle_eq_iff_norm_eq (x y : V) : + o.rotation (o.oangle x y) x = y ↔ ‖x‖ = ‖y‖ := +begin + split, + { intro h, + rw [←h, linear_isometry_equiv.norm_map] }, + { intro h, + rw o.eq_iff_oangle_eq_zero_of_norm_eq; + simp [h] } +end + +/-- The angle between two nonzero vectors is `θ` if and only if the second vector is the first +rotated by `θ` and scaled by the ratio of the norms. -/ +lemma oangle_eq_iff_eq_norm_div_norm_smul_rotation_of_ne_zero {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) + (θ : real.angle) : o.oangle x y = θ ↔ y = (‖y‖ / ‖x‖) • o.rotation θ x := +begin + have hp := div_pos (norm_pos_iff.2 hy) (norm_pos_iff.2 hx), + split, + { rintro rfl, + rw [←linear_isometry_equiv.map_smul, ←o.oangle_smul_left_of_pos x y hp, + eq_comm, rotation_oangle_eq_iff_norm_eq, norm_smul, real.norm_of_nonneg hp.le, + div_mul_cancel _ (norm_ne_zero_iff.2 hx)] }, + { intro hye, + rw [hye, o.oangle_smul_right_of_pos _ _ hp, o.oangle_rotation_self_right hx] } +end + +/-- The angle between two nonzero vectors is `θ` if and only if the second vector is the first +rotated by `θ` and scaled by a positive real. -/ +lemma oangle_eq_iff_eq_pos_smul_rotation_of_ne_zero {x y : V} (hx : x ≠ 0) (hy : y ≠ 0) + (θ : real.angle) : o.oangle x y = θ ↔ ∃ r : ℝ, 0 < r ∧ y = r • o.rotation θ x := +begin + split, + { intro h, + rw o.oangle_eq_iff_eq_norm_div_norm_smul_rotation_of_ne_zero hx hy at h, + exact ⟨‖y‖ / ‖x‖, div_pos (norm_pos_iff.2 hy) (norm_pos_iff.2 hx), h⟩ }, + { rintro ⟨r, hr, rfl⟩, + rw [o.oangle_smul_right_of_pos _ _ hr, o.oangle_rotation_self_right hx] } +end + +/-- The angle between two vectors is `θ` if and only if they are nonzero and the second vector +is the first rotated by `θ` and scaled by the ratio of the norms, or `θ` and at least one of the +vectors are zero. -/ +lemma oangle_eq_iff_eq_norm_div_norm_smul_rotation_or_eq_zero {x y : V} (θ : real.angle) : + o.oangle x y = θ ↔ + (x ≠ 0 ∧ y ≠ 0 ∧ y = (‖y‖ / ‖x‖) • o.rotation θ x) ∨ (θ = 0 ∧ (x = 0 ∨ y = 0)) := +begin + by_cases hx : x = 0, + { simp [hx, eq_comm] }, + { by_cases hy : y = 0, + { simp [hy, eq_comm] }, + { rw o.oangle_eq_iff_eq_norm_div_norm_smul_rotation_of_ne_zero hx hy, + simp [hx, hy] } } +end + +/-- The angle between two vectors is `θ` if and only if they are nonzero and the second vector +is the first rotated by `θ` and scaled by a positive real, or `θ` and at least one of the +vectors are zero. -/ +lemma oangle_eq_iff_eq_pos_smul_rotation_or_eq_zero {x y : V} (θ : real.angle) : + o.oangle x y = θ ↔ + (x ≠ 0 ∧ y ≠ 0 ∧ ∃ r : ℝ, 0 < r ∧ y = r • o.rotation θ x) ∨ (θ = 0 ∧ (x = 0 ∨ y = 0)) := +begin + by_cases hx : x = 0, + { simp [hx, eq_comm] }, + { by_cases hy : y = 0, + { simp [hy, eq_comm] }, + { rw o.oangle_eq_iff_eq_pos_smul_rotation_of_ne_zero hx hy, + simp [hx, hy] } } +end + +/-- Any linear isometric equivalence in `V` with positive determinant is `rotation`. -/ +lemma exists_linear_isometry_equiv_eq_of_det_pos {f : V ≃ₗᵢ[ℝ] V} + (hd : 0 < (f.to_linear_equiv : V →ₗ[ℝ] V).det) : ∃ θ : real.angle, f = o.rotation θ := +begin + haveI : nontrivial V := + finite_dimensional.nontrivial_of_finrank_eq_succ (fact.out (finrank ℝ V = 2)), + obtain ⟨x, hx⟩ : ∃ x, x ≠ (0:V) := exists_ne (0:V), + use o.oangle x (f x), + apply linear_isometry_equiv.to_linear_equiv_injective, + apply linear_equiv.to_linear_map_injective, + apply (o.basis_right_angle_rotation x hx).ext, + intros i, + symmetry, + fin_cases i, + { simp }, + have : o.oangle (J x) (f (J x)) = o.oangle x (f x), + { simp only [oangle, o.linear_isometry_equiv_comp_right_angle_rotation f hd, + o.kahler_comp_right_angle_rotation] }, + simp [← this], +end + +lemma rotation_map (θ : real.angle) (f : V ≃ₗᵢ[ℝ] V') (x : V') : + (orientation.map (fin 2) f.to_linear_equiv o).rotation θ x + = f (o.rotation θ (f.symm x)) := +by simp [rotation_apply, o.right_angle_rotation_map] + +@[simp] protected lemma _root_.complex.rotation (θ : real.angle) (z : ℂ) : + complex.orientation.rotation θ z = θ.exp_map_circle * z := +begin + simp only [rotation_apply, complex.right_angle_rotation, real.angle.coe_exp_map_circle, + real_smul], + ring +end + +/-- Rotation in an oriented real inner product space of dimension 2 can be evaluated in terms of a +complex-number representation of the space. -/ +lemma rotation_map_complex (θ : real.angle) (f : V ≃ₗᵢ[ℝ] ℂ) + (hf : (orientation.map (fin 2) f.to_linear_equiv o) = complex.orientation) (x : V) : + f (o.rotation θ x) = θ.exp_map_circle * f x := +begin + rw [← complex.rotation, ← hf, o.rotation_map], + simp, +end + +/-- Negating the orientation negates the angle in `rotation`. -/ +lemma rotation_neg_orientation_eq_neg (θ : real.angle) : + (-o).rotation θ = o.rotation (-θ) := +linear_isometry_equiv.ext $ by simp [rotation_apply] + +/-- The inner product between a `π / 2` rotation of a vector and that vector is zero. -/ +@[simp] lemma inner_rotation_pi_div_two_left (x : V) : ⟪o.rotation (π / 2 : ℝ) x, x⟫ = 0 := +by rw [rotation_pi_div_two, inner_right_angle_rotation_self] + +/-- The inner product between a vector and a `π / 2` rotation of that vector is zero. -/ +@[simp] lemma inner_rotation_pi_div_two_right (x : V) : ⟪x, o.rotation (π / 2 : ℝ) x⟫ = 0 := +by rw [real_inner_comm, inner_rotation_pi_div_two_left] + +/-- The inner product between a multiple of a `π / 2` rotation of a vector and that vector is +zero. -/ +@[simp] lemma inner_smul_rotation_pi_div_two_left (x : V) (r : ℝ) : + ⟪r • o.rotation (π / 2 : ℝ) x, x⟫ = 0 := +by rw [inner_smul_left, inner_rotation_pi_div_two_left, mul_zero] + +/-- The inner product between a vector and a multiple of a `π / 2` rotation of that vector is +zero. -/ +@[simp] lemma inner_smul_rotation_pi_div_two_right (x : V) (r : ℝ) : + ⟪x, r • o.rotation (π / 2 : ℝ) x⟫ = 0 := +by rw [real_inner_comm, inner_smul_rotation_pi_div_two_left] + +/-- The inner product between a `π / 2` rotation of a vector and a multiple of that vector is +zero. -/ +@[simp] lemma inner_rotation_pi_div_two_left_smul (x : V) (r : ℝ) : + ⟪o.rotation (π / 2 : ℝ) x, r • x⟫ = 0 := +by rw [inner_smul_right, inner_rotation_pi_div_two_left, mul_zero] + +/-- The inner product between a multiple of a vector and a `π / 2` rotation of that vector is +zero. -/ +@[simp] lemma inner_rotation_pi_div_two_right_smul (x : V) (r : ℝ) : + ⟪r • x, o.rotation (π / 2 : ℝ) x⟫ = 0 := +by rw [real_inner_comm, inner_rotation_pi_div_two_left_smul] + +/-- The inner product between a multiple of a `π / 2` rotation of a vector and a multiple of +that vector is zero. -/ +@[simp] lemma inner_smul_rotation_pi_div_two_smul_left (x : V) (r₁ r₂ : ℝ) : + ⟪r₁ • o.rotation (π / 2 : ℝ) x, r₂ • x⟫ = 0 := +by rw [inner_smul_right, inner_smul_rotation_pi_div_two_left, mul_zero] + +/-- The inner product between a multiple of a vector and a multiple of a `π / 2` rotation of +that vector is zero. -/ +@[simp] lemma inner_smul_rotation_pi_div_two_smul_right (x : V) (r₁ r₂ : ℝ) : + ⟪r₂ • x, r₁ • o.rotation (π / 2 : ℝ) x⟫ = 0 := +by rw [real_inner_comm, inner_smul_rotation_pi_div_two_smul_left] + +/-- The inner product between two vectors is zero if and only if the first vector is zero or +the second is a multiple of a `π / 2` rotation of that vector. -/ +lemma inner_eq_zero_iff_eq_zero_or_eq_smul_rotation_pi_div_two {x y : V} : + ⟪x, y⟫ = 0 ↔ (x = 0 ∨ ∃ r : ℝ, r • o.rotation (π / 2 : ℝ) x = y) := +begin + rw ←o.eq_zero_or_oangle_eq_iff_inner_eq_zero, + refine ⟨λ h, _, λ h, _⟩, + { rcases h with rfl | rfl | h | h, + { exact or.inl rfl }, + { exact or.inr ⟨0, zero_smul _ _⟩ }, + { obtain ⟨r, hr, rfl⟩ := (o.oangle_eq_iff_eq_pos_smul_rotation_of_ne_zero + (o.left_ne_zero_of_oangle_eq_pi_div_two h) + (o.right_ne_zero_of_oangle_eq_pi_div_two h) _).1 h, + exact or.inr ⟨r, rfl⟩ }, + { obtain ⟨r, hr, rfl⟩ := (o.oangle_eq_iff_eq_pos_smul_rotation_of_ne_zero + (o.left_ne_zero_of_oangle_eq_neg_pi_div_two h) + (o.right_ne_zero_of_oangle_eq_neg_pi_div_two h) _).1 h, + refine or.inr ⟨-r, _⟩, + rw [neg_smul, ←smul_neg, o.neg_rotation_pi_div_two] } }, + { rcases h with rfl | ⟨r, rfl⟩, + { exact or.inl rfl }, + { by_cases hx : x = 0, { exact or.inl hx }, + rcases lt_trichotomy r 0 with hr | rfl | hr, + { refine or.inr (or.inr (or.inr _)), + rw [o.oangle_smul_right_of_neg _ _ hr, o.neg_rotation_pi_div_two, + o.oangle_rotation_self_right hx] }, + { exact or.inr (or.inl (zero_smul _ _)) }, + { refine or.inr (or.inr (or.inl _)), + rw [o.oangle_smul_right_of_pos _ _ hr, o.oangle_rotation_self_right hx] } } } +end + +end orientation diff --git a/src/geometry/euclidean/angle/unoriented/basic.lean b/src/geometry/euclidean/angle/unoriented/basic.lean index b361b70d0be13..6afbd85c4ec5c 100644 --- a/src/geometry/euclidean/angle/unoriented/basic.lean +++ b/src/geometry/euclidean/angle/unoriented/basic.lean @@ -3,7 +3,6 @@ Copyright (c) 2020 Joseph Myers. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Myers, Manuel Candales -/ -import analysis.calculus.conformal.normed_space import analysis.inner_product_space.basic import analysis.special_functions.trigonometric.inverse @@ -18,6 +17,9 @@ This file defines unoriented angles in real inner product spaces. -/ +assert_not_exists has_fderiv_at +assert_not_exists conformal_at + noncomputable theory open_locale big_operators open_locale real @@ -53,24 +55,6 @@ by rw [angle, angle, f.inner_map_map, f.norm_map, f.norm_map] angle (x : V) (y : V) = angle x y := s.subtypeₗᵢ.angle_map x y -lemma is_conformal_map.preserves_angle {E F : Type*} - [inner_product_space ℝ E] [inner_product_space ℝ F] - {f' : E →L[ℝ] F} (h : is_conformal_map f') (u v : E) : - angle (f' u) (f' v) = angle u v := -begin - obtain ⟨c, hc, li, rfl⟩ := h, - exact (angle_smul_smul hc _ _).trans (li.angle_map _ _) -end - -/-- If a real differentiable map `f` is conformal at a point `x`, - then it preserves the angles at that point. -/ -lemma conformal_at.preserves_angle {E F : Type*} - [inner_product_space ℝ E] [inner_product_space ℝ F] - {f : E → F} {x : E} {f' : E →L[ℝ] F} - (h : has_fderiv_at f f' x) (H : conformal_at f x) (u v : E) : - angle (f' u) (f' v) = angle u v := -let ⟨f₁, h₁, c⟩ := H in h₁.unique h ▸ is_conformal_map.preserves_angle c u v - /-- The cosine of the angle between two vectors. -/ lemma cos_angle (x y : V) : real.cos (angle x y) = ⟪x, y⟫ / (‖x‖ * ‖y‖) := real.cos_arccos (abs_le.mp (abs_real_inner_div_norm_mul_norm_le_one x y)).1 diff --git a/src/geometry/euclidean/angle/unoriented/conformal.lean b/src/geometry/euclidean/angle/unoriented/conformal.lean new file mode 100644 index 0000000000000..cd65fa3ca0463 --- /dev/null +++ b/src/geometry/euclidean/angle/unoriented/conformal.lean @@ -0,0 +1,38 @@ +/- +Copyright (c) 2021 Yourong Zang. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Yourong Zang +-/ +import analysis.calculus.conformal.normed_space +import geometry.euclidean.angle.unoriented.basic + +/-! +# Angles and conformal maps + +This file proves that conformal maps preserve angles. + +-/ + +namespace inner_product_geometry + +variables {V : Type*} [inner_product_space ℝ V] + +lemma is_conformal_map.preserves_angle {E F : Type*} + [inner_product_space ℝ E] [inner_product_space ℝ F] + {f' : E →L[ℝ] F} (h : is_conformal_map f') (u v : E) : + angle (f' u) (f' v) = angle u v := +begin + obtain ⟨c, hc, li, rfl⟩ := h, + exact (angle_smul_smul hc _ _).trans (li.angle_map _ _) +end + +/-- If a real differentiable map `f` is conformal at a point `x`, + then it preserves the angles at that point. -/ +lemma conformal_at.preserves_angle {E F : Type*} + [inner_product_space ℝ E] [inner_product_space ℝ F] + {f : E → F} {x : E} {f' : E →L[ℝ] F} + (h : has_fderiv_at f f' x) (H : conformal_at f x) (u v : E) : + angle (f' u) (f' v) = angle u v := +let ⟨f₁, h₁, c⟩ := H in h₁.unique h ▸ is_conformal_map.preserves_angle c u v + +end inner_product_geometry diff --git a/src/geometry/euclidean/circumcenter.lean b/src/geometry/euclidean/circumcenter.lean index f03398d2a09fc..da7ffb6e7c325 100644 --- a/src/geometry/euclidean/circumcenter.lean +++ b/src/geometry/euclidean/circumcenter.lean @@ -188,10 +188,11 @@ end /-- Given a finite nonempty affinely independent family of points, there is a unique (circumcenter, circumradius) pair for those points in the affine subspace they span. -/ -lemma _root_.affine_independent.exists_unique_dist_eq {ι : Type*} [hne : nonempty ι] [fintype ι] +lemma _root_.affine_independent.exists_unique_dist_eq {ι : Type*} [hne : nonempty ι] [finite ι] {p : ι → P} (ha : affine_independent ℝ p) : ∃! cs : sphere P, cs.center ∈ affine_span ℝ (set.range p) ∧ set.range p ⊆ (cs : set P) := begin + casesI nonempty_fintype ι, unfreezingI { induction hn : fintype.card ι with m hm generalizing ι }, { exfalso, have h := fintype.card_pos_iff.2 hne, @@ -225,7 +226,7 @@ begin { simp } }, haveI : nonempty ι2 := fintype.card_pos_iff.1 (hc.symm ▸ nat.zero_lt_succ _), have ha2 : affine_independent ℝ (λ i2 : ι2, p i2) := ha.subtype _, - replace hm := hm ha2 hc, + replace hm := hm ha2 _ hc, have hr : set.range p = insert (p i) (set.range (λ i2 : ι2, p i2)), { change _ = insert _ (set.range (λ i2 : {x | x ≠ i}, p i2)), rw [←set.image_eq_range, ←set.image_univ, ←set.image_insert_eq], diff --git a/src/group_theory/commutator.lean b/src/group_theory/commutator.lean index a70850a6c1604..3206dab1db37a 100644 --- a/src/group_theory/commutator.lean +++ b/src/group_theory/commutator.lean @@ -5,7 +5,7 @@ Authors: Jordan Brown, Thomas Browning, Patrick Lutz -/ import data.bracket -import group_theory.subgroup.basic +import group_theory.subgroup.finite import tactic.group /-! diff --git a/src/group_theory/congruence.lean b/src/group_theory/congruence.lean index 11748fd953fb4..857995ec74330 100644 --- a/src/group_theory/congruence.lean +++ b/src/group_theory/congruence.lean @@ -11,6 +11,9 @@ import group_theory.submonoid.operations /-! # Congruence relations +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines congruence relations: equivalence relations that preserve a binary operation, which in this case is multiplication or addition. The principal definition is a `structure` extending a `setoid` (an equivalence relation), and the inductive definition of the smallest diff --git a/src/group_theory/coset.lean b/src/group_theory/coset.lean index 48d24dd1d9afa..69843a35df0c7 100644 --- a/src/group_theory/coset.lean +++ b/src/group_theory/coset.lean @@ -5,7 +5,9 @@ Authors: Mitchell Rowett, Scott Morrison -/ import algebra.quotient +import data.fintype.prod import group_theory.group_action.basic +import group_theory.subgroup.mul_opposite import tactic.group /-! diff --git a/src/group_theory/free_group.lean b/src/group_theory/free_group.lean index 8c477a913ba27..ed5d485a08040 100644 --- a/src/group_theory/free_group.lean +++ b/src/group_theory/free_group.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Kenny Lau -/ import data.fintype.basic +import data.list.sublists import group_theory.subgroup.basic /-! diff --git a/src/group_theory/group_action/basic.lean b/src/group_theory/group_action/basic.lean index d3e5fd0d59cb6..faec650307fa1 100644 --- a/src/group_theory/group_action/basic.lean +++ b/src/group_theory/group_action/basic.lean @@ -3,6 +3,7 @@ Copyright (c) 2018 Chris Hughes. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ +import data.fintype.card import group_theory.group_action.defs import group_theory.group_action.group import data.setoid.basic @@ -28,7 +29,7 @@ of `•` belong elsewhere. universes u v w variables {α : Type u} {β : Type v} {γ : Type w} -open_locale big_operators pointwise +open_locale pointwise open function namespace mul_action diff --git a/src/group_theory/group_action/conj_act.lean b/src/group_theory/group_action/conj_act.lean index 6c216817766fe..5a9ed7a2ffb8c 100644 --- a/src/group_theory/group_action/conj_act.lean +++ b/src/group_theory/group_action/conj_act.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Chris Hughes -/ import group_theory.group_action.basic -import group_theory.subgroup.basic +import group_theory.subgroup.zpowers import algebra.group_ring_action.basic /-! # Conjugation action of a group on itself diff --git a/src/group_theory/group_action/fixing_subgroup.lean b/src/group_theory/group_action/fixing_subgroup.lean index 3718d4c8ac4da..ba7d078189e48 100644 --- a/src/group_theory/group_action/fixing_subgroup.lean +++ b/src/group_theory/group_action/fixing_subgroup.lean @@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Antoine Chambert-Loir -/ -import group_theory.subgroup.basic +import group_theory.subgroup.actions import group_theory.group_action.basic /-! diff --git a/src/group_theory/noncomm_pi_coprod.lean b/src/group_theory/noncomm_pi_coprod.lean index 500d7f5ad3f3e..3e2db788d60d9 100644 --- a/src/group_theory/noncomm_pi_coprod.lean +++ b/src/group_theory/noncomm_pi_coprod.lean @@ -7,6 +7,7 @@ import group_theory.order_of_element import data.finset.noncomm_prod import data.fintype.big_operators import data.nat.gcd.big_operators +import order.sup_indep /-! # Canonical homomorphism from a finite family of monoids @@ -42,6 +43,44 @@ images of different morphisms commute, we obtain a canonical morphism open_locale big_operators +namespace subgroup + +variables {G : Type*} [group G] + +/-- `finset.noncomm_prod` is “injective” in `f` if `f` maps into independent subgroups. This +generalizes (one direction of) `subgroup.disjoint_iff_mul_eq_one`. -/ +@[to_additive "`finset.noncomm_sum` is “injective” in `f` if `f` maps into independent subgroups. +This generalizes (one direction of) `add_subgroup.disjoint_iff_add_eq_zero`. "] +lemma eq_one_of_noncomm_prod_eq_one_of_independent {ι : Type*} (s : finset ι) (f : ι → G) (comm) + (K : ι → subgroup G) (hind : complete_lattice.independent K) (hmem : ∀ (x ∈ s), f x ∈ K x) + (heq1 : s.noncomm_prod f comm = 1) : ∀ (i ∈ s), f i = 1 := +begin + classical, + revert heq1, + induction s using finset.induction_on with i s hnmem ih, + { simp, }, + { have hcomm := comm.mono (finset.coe_subset.2 $ finset.subset_insert _ _), + simp only [finset.forall_mem_insert] at hmem, + have hmem_bsupr: s.noncomm_prod f hcomm ∈ ⨆ (i ∈ (s : set ι)), K i, + { refine subgroup.noncomm_prod_mem _ _ _, + intros x hx, + have : K x ≤ ⨆ (i ∈ (s : set ι)), K i := le_supr₂ x hx, + exact this (hmem.2 x hx), }, + intro heq1, + rw finset.noncomm_prod_insert_of_not_mem _ _ _ _ hnmem at heq1, + have hnmem' : i ∉ (s : set ι), by simpa, + obtain ⟨heq1i : f i = 1, heq1S : s.noncomm_prod f _ = 1⟩ := + subgroup.disjoint_iff_mul_eq_one.mp (hind.disjoint_bsupr hnmem') hmem.1 hmem_bsupr heq1, + intros i h, + simp only [finset.mem_insert] at h, + rcases h with ⟨rfl | _⟩, + { exact heq1i }, + { exact ih hcomm hmem.2 heq1S _ h } } +end + +end subgroup + + section family_of_monoids variables {M : Type*} [monoid M] diff --git a/src/group_theory/perm/subgroup.lean b/src/group_theory/perm/subgroup.lean index 6bfaa6fafde50..cf8f21c6a40c3 100644 --- a/src/group_theory/perm/subgroup.lean +++ b/src/group_theory/perm/subgroup.lean @@ -5,7 +5,7 @@ Authors: Eric Wieser -/ import group_theory.perm.basic import data.fintype.perm -import group_theory.subgroup.basic +import group_theory.subgroup.finite /-! # Lemmas about subgroups within the permutations (self-equivalences) of a type `α` diff --git a/src/group_theory/perm/support.lean b/src/group_theory/perm/support.lean index f8aac0b2bf1e7..78fa4c4dc079a 100644 --- a/src/group_theory/perm/support.lean +++ b/src/group_theory/perm/support.lean @@ -10,6 +10,9 @@ import group_theory.perm.basic /-! # Support of a permutation +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + ## Main definitions In the following, `f g : equiv.perm α`. diff --git a/src/group_theory/quotient_group.lean b/src/group_theory/quotient_group.lean index 14fc5da61645d..7d80d10bef0c1 100644 --- a/src/group_theory/quotient_group.lean +++ b/src/group_theory/quotient_group.lean @@ -7,6 +7,7 @@ This file is to a certain extent based on `quotient_module.lean` by Johannes Hö -/ import group_theory.congruence import group_theory.coset +import group_theory.subgroup.finite import group_theory.subgroup.pointwise /-! diff --git a/src/group_theory/solvable.lean b/src/group_theory/solvable.lean index e9ae039668dda..d02f872a54c4d 100644 --- a/src/group_theory/solvable.lean +++ b/src/group_theory/solvable.lean @@ -7,6 +7,7 @@ Authors: Jordan Brown, Thomas Browning, Patrick Lutz import data.fin.vec_notation import group_theory.abelianization import group_theory.perm.via_embedding +import group_theory.subgroup.simple import set_theory.cardinal.basic /-! diff --git a/src/group_theory/specific_groups/alternating.lean b/src/group_theory/specific_groups/alternating.lean index 5fddc8b655f95..5f4a13d2e7245 100644 --- a/src/group_theory/specific_groups/alternating.lean +++ b/src/group_theory/specific_groups/alternating.lean @@ -6,6 +6,7 @@ Authors: Aaron Anderson import algebra.group.conj_finite import group_theory.perm.fin +import group_theory.subgroup.simple import tactic.interval_cases /-! diff --git a/src/group_theory/specific_groups/cyclic.lean b/src/group_theory/specific_groups/cyclic.lean index 21e961156bfdb..9a48297f65a54 100644 --- a/src/group_theory/specific_groups/cyclic.lean +++ b/src/group_theory/specific_groups/cyclic.lean @@ -7,6 +7,7 @@ Authors: Johannes Hölzl import algebra.big_operators.order import data.nat.totient import group_theory.order_of_element +import group_theory.subgroup.simple import tactic.group import group_theory.exponent diff --git a/src/group_theory/subgroup/actions.lean b/src/group_theory/subgroup/actions.lean new file mode 100644 index 0000000000000..c3d8d0ad3fa79 --- /dev/null +++ b/src/group_theory/subgroup/actions.lean @@ -0,0 +1,71 @@ +/- +Copyright (c) 2021 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ + +import group_theory.subgroup.basic + +/-! +# Actions by `subgroup`s + +These are just copies of the definitions about `submonoid` starting from `submonoid.mul_action`. + +## Tags +subgroup, subgroups + +-/ + +namespace subgroup + +variables {G : Type*} [group G] +variables {α β : Type*} + +/-- The action by a subgroup is the action by the underlying group. -/ +@[to_additive /-"The additive action by an add_subgroup is the action by the underlying +add_group. "-/] +instance [mul_action G α] (S : subgroup G) : mul_action S α := +S.to_submonoid.mul_action + +@[to_additive] +lemma smul_def [mul_action G α] {S : subgroup G} (g : S) (m : α) : g • m = (g : G) • m := rfl + +@[to_additive] +instance smul_comm_class_left + [mul_action G β] [has_smul α β] [smul_comm_class G α β] (S : subgroup G) : + smul_comm_class S α β := +S.to_submonoid.smul_comm_class_left + +@[to_additive] +instance smul_comm_class_right + [has_smul α β] [mul_action G β] [smul_comm_class α G β] (S : subgroup G) : + smul_comm_class α S β := +S.to_submonoid.smul_comm_class_right + +/-- Note that this provides `is_scalar_tower S G G` which is needed by `smul_mul_assoc`. -/ +instance + [has_smul α β] [mul_action G α] [mul_action G β] [is_scalar_tower G α β] (S : subgroup G) : + is_scalar_tower S α β := +S.to_submonoid.is_scalar_tower + +instance [mul_action G α] [has_faithful_smul G α] (S : subgroup G) : + has_faithful_smul S α := +S.to_submonoid.has_faithful_smul + +/-- The action by a subgroup is the action by the underlying group. -/ +instance [add_monoid α] [distrib_mul_action G α] (S : subgroup G) : distrib_mul_action S α := +S.to_submonoid.distrib_mul_action + +/-- The action by a subgroup is the action by the underlying group. -/ +instance [monoid α] [mul_distrib_mul_action G α] (S : subgroup G) : mul_distrib_mul_action S α := +S.to_submonoid.mul_distrib_mul_action + +/-- The center of a group acts commutatively on that group. -/ +instance center.smul_comm_class_left : smul_comm_class (center G) G G := +submonoid.center.smul_comm_class_left + +/-- The center of a group acts commutatively on that group. -/ +instance center.smul_comm_class_right : smul_comm_class G (center G) G := +submonoid.center.smul_comm_class_right + +end subgroup diff --git a/src/group_theory/subgroup/basic.lean b/src/group_theory/subgroup/basic.lean index a5c2c2982e21d..8e4c2f27970bd 100644 --- a/src/group_theory/subgroup/basic.lean +++ b/src/group_theory/subgroup/basic.lean @@ -7,12 +7,10 @@ import algebra.group.conj import algebra.module.basic import algebra.order.group.inj_surj import data.countable.basic -import data.set.finite import group_theory.submonoid.centralizer -import group_theory.submonoid.membership import logic.encodable.basic import order.atoms -import order.sup_indep +import tactic.apply_fun /-! # Subgroups @@ -76,8 +74,6 @@ Definitions in the file: * `monoid_hom.eq_locus f g` : given group homomorphisms `f`, `g`, the elements of `G` such that `f x = g x` form a subgroup of `G` -* `is_simple_group G` : a class indicating that a group has exactly two normal subgroups - ## Implementation notes Subgroup inclusion is denoted `≤` rather than `⊆`, although `∈` is defined as @@ -88,7 +84,6 @@ subgroup, subgroups -/ open function -open_locale big_operators variables {G G' : Type*} [group G] [group G'] variables {A : Type*} [add_group A] @@ -326,14 +321,6 @@ lemma coe_to_submonoid (K : subgroup G) : (K.to_submonoid : set G) = K := rfl @[simp, to_additive] lemma mem_to_submonoid (K : subgroup G) (x : G) : x ∈ K.to_submonoid ↔ x ∈ K := iff.rfl -@[to_additive] -instance (K : subgroup G) [d : decidable_pred (∈ K)] [fintype G] : fintype K := -show fintype {g : G // g ∈ K}, from infer_instance - -@[to_additive] -instance (K : subgroup G) [finite G] : finite K := -subtype.finite - @[to_additive] theorem to_submonoid_injective : function.injective (to_submonoid : subgroup G → submonoid G) := @@ -459,37 +446,6 @@ mul_mem_cancel_right h @[to_additive] protected lemma mul_mem_cancel_left {x y : G} (h : x ∈ H) : x * y ∈ H ↔ y ∈ H := mul_mem_cancel_left h -/-- Product of a list of elements in a subgroup is in the subgroup. -/ -@[to_additive "Sum of a list of elements in an `add_subgroup` is in the `add_subgroup`."] -protected lemma list_prod_mem {l : list G} : (∀ x ∈ l, x ∈ K) → l.prod ∈ K := -list_prod_mem - -/-- Product of a multiset of elements in a subgroup of a `comm_group` is in the subgroup. -/ -@[to_additive "Sum of a multiset of elements in an `add_subgroup` of an `add_comm_group` -is in the `add_subgroup`."] -protected lemma multiset_prod_mem {G} [comm_group G] (K : subgroup G) (g : multiset G) : - (∀ a ∈ g, a ∈ K) → g.prod ∈ K := multiset_prod_mem g - -@[to_additive] -lemma multiset_noncomm_prod_mem (K : subgroup G) (g : multiset G) (comm) : - (∀ a ∈ g, a ∈ K) → g.noncomm_prod comm ∈ K := -K.to_submonoid.multiset_noncomm_prod_mem g comm - -/-- Product of elements of a subgroup of a `comm_group` indexed by a `finset` is in the - subgroup. -/ -@[to_additive "Sum of elements in an `add_subgroup` of an `add_comm_group` indexed by a `finset` -is in the `add_subgroup`."] -protected lemma prod_mem {G : Type*} [comm_group G] (K : subgroup G) - {ι : Type*} {t : finset ι} {f : ι → G} (h : ∀ c ∈ t, f c ∈ K) : - ∏ c in t, f c ∈ K := -prod_mem h - -@[to_additive] -lemma noncomm_prod_mem (K : subgroup G) {ι : Type*} {t : finset ι} {f : ι → G} (comm) : - (∀ c ∈ t, f c ∈ K) → t.noncomm_prod f comm ∈ K := -K.to_submonoid.noncomm_prod_mem t f comm - - @[to_additive add_subgroup.nsmul_mem] protected lemma pow_mem {x : G} (hx : x ∈ K) : ∀ n : ℕ, x ^ n ∈ K := pow_mem hx @@ -583,19 +539,6 @@ def subtype : H →* G := ⟨coe, rfl, λ _ _, rfl⟩ @[to_additive] lemma subtype_injective : function.injective (subtype H) := subtype.coe_injective -@[simp, norm_cast, to_additive] theorem coe_list_prod (l : list H) : - (l.prod : G) = (l.map coe).prod := -submonoid_class.coe_list_prod l - -@[simp, norm_cast, to_additive] theorem coe_multiset_prod {G} [comm_group G] (H : subgroup G) - (m : multiset H) : (m.prod : G) = (m.map coe).prod := -submonoid_class.coe_multiset_prod m - -@[simp, norm_cast, to_additive] theorem coe_finset_prod {ι G} [comm_group G] (H : subgroup G) - (f : ι → H) (s : finset ι) : - ↑(∏ i in s, f i) = (∏ i in s, f i : G) := -submonoid_class.coe_finset_prod f s - /-- The inclusion homomorphism from a subgroup `H` contained in `K` to `K`. -/ @[to_additive "The inclusion homomorphism from a additive subgroup `H` contained in `K` to `K`."] def inclusion {H K : subgroup G} (h : H ≤ K) : H →* K := @@ -666,35 +609,6 @@ end ⟨λ ⟨g, hg⟩, by { haveI : subsingleton (H : set G) := by { rw hg, apply_instance }, exact H.eq_bot_of_subsingleton }, λ h, ⟨1, set_like.ext'_iff.mp h⟩⟩ -@[to_additive] instance fintype_bot : fintype (⊥ : subgroup G) := ⟨{1}, -by {rintro ⟨x, ⟨hx⟩⟩, exact finset.mem_singleton_self _}⟩ - -/- curly brackets `{}` are used here instead of instance brackets `[]` because - the instance in a goal is often not the same as the one inferred by type class inference. -/ -@[simp, to_additive] lemma card_bot {_ : fintype ↥(⊥ : subgroup G)} : - fintype.card (⊥ : subgroup G) = 1 := -fintype.card_eq_one_iff.2 - ⟨⟨(1 : G), set.mem_singleton 1⟩, λ ⟨y, hy⟩, subtype.eq $ subgroup.mem_bot.1 hy⟩ - -@[to_additive] lemma eq_top_of_card_eq [fintype H] [fintype G] - (h : fintype.card H = fintype.card G) : H = ⊤ := -begin - haveI : fintype (H : set G) := ‹fintype H›, - rw [set_like.ext'_iff, coe_top, ← finset.coe_univ, ← (H : set G).coe_to_finset, finset.coe_inj, - ← finset.card_eq_iff_eq_univ, ← h, set.to_finset_card], - congr -end - -@[to_additive] lemma eq_top_of_le_card [fintype H] [fintype G] - (h : fintype.card G ≤ fintype.card H) : H = ⊤ := -eq_top_of_card_eq H (le_antisymm (fintype.card_le_of_injective coe subtype.coe_injective) h) - -@[to_additive] lemma eq_bot_of_card_le [fintype H] (h : fintype.card H ≤ 1) : H = ⊥ := -let _ := fintype.card_le_one_iff_subsingleton.mp h in by exactI eq_bot_of_subsingleton H - -@[to_additive] lemma eq_bot_of_card_eq [fintype H] (h : fintype.card H = 1) : H = ⊥ := -H.eq_bot_of_card_le (le_of_eq h) - @[to_additive] lemma nontrivial_iff_exists_ne_one (H : subgroup G) : nontrivial H ↔ ∃ x ∈ H, x ≠ (1:G) := subtype.nontrivial_iff_exists_ne (λ x, x ∈ H) (1 : H) @@ -720,14 +634,6 @@ begin rw nontrivial_iff_exists_ne_one end -@[to_additive] lemma card_le_one_iff_eq_bot [fintype H] : fintype.card H ≤ 1 ↔ H = ⊥ := -⟨λ h, (eq_bot_iff_forall _).2 - (λ x hx, by simpa [subtype.ext_iff] using fintype.card_le_one_iff.1 h ⟨x, hx⟩ 1), - λ h, by simp [h]⟩ - -@[to_additive] lemma one_lt_card_iff_ne_bot [fintype H] : 1 < fintype.card H ↔ H ≠ ⊥ := -lt_iff_not_le.trans H.card_le_one_iff_eq_bot.not - /-- The inf of two subgroups is their intersection. -/ @[to_additive "The inf of two `add_subgroups`s is their intersection."] instance : has_inf (subgroup G) := @@ -1384,51 +1290,6 @@ begin { simp [heq, one_mem], }, } end -@[to_additive] -lemma pi_mem_of_mul_single_mem_aux [decidable_eq η] (I : finset η) {H : subgroup (Π i, f i) } - (x : Π i, f i) (h1 : ∀ i, i ∉ I → x i = 1) (h2 : ∀ i, i ∈ I → pi.mul_single i (x i) ∈ H ) : - x ∈ H := -begin - induction I using finset.induction_on with i I hnmem ih generalizing x, - { convert one_mem H, - ext i, - exact (h1 i (not_mem_empty i)) }, - { have : x = function.update x i 1 * pi.mul_single i (x i), - { ext j, - by_cases heq : j = i, - { subst heq, simp, }, - { simp [heq], }, }, - rw this, clear this, - apply mul_mem, - { apply ih; clear ih, - { intros j hj, - by_cases heq : j = i, - { subst heq, simp, }, - { simp [heq], apply h1 j, simpa [heq] using hj, } }, - { intros j hj, - have : j ≠ i, by { rintro rfl, contradiction }, - simp [this], - exact h2 _ (finset.mem_insert_of_mem hj), }, }, - { apply h2, simp, } } -end - -@[to_additive] -lemma pi_mem_of_mul_single_mem [finite η] [decidable_eq η] {H : subgroup (Π i, f i)} - (x : Π i, f i) (h : ∀ i, pi.mul_single i (x i) ∈ H) : x ∈ H := -by { casesI nonempty_fintype η, - exact pi_mem_of_mul_single_mem_aux finset.univ x (by simp) (λ i _, h i) } - -/-- For finite index types, the `subgroup.pi` is generated by the embeddings of the groups. -/ -@[to_additive "For finite index types, the `subgroup.pi` is generated by the embeddings of the -additive groups."] -lemma pi_le_iff [decidable_eq η] [finite η] {H : Π i, subgroup (f i)} {J : subgroup (Π i, f i)} : - pi univ H ≤ J ↔ ∀ i : η, map (monoid_hom.single f i) (H i) ≤ J := -begin - split, - { rintros h i _ ⟨x, hx, rfl⟩, apply h, simpa using hx }, - { exact λ h x hx, pi_mem_of_mul_single_mem x (λ i, h i (mem_map_of_mem _ (hx i trivial))), } -end - @[to_additive] lemma pi_eq_bot_iff (H : Π i, subgroup (f i)) : pi set.univ H = ⊥ ↔ ∀ i, H i = ⊥ := @@ -1578,8 +1439,8 @@ variable {G} @[to_additive] lemma mem_center_iff {z : G} : z ∈ center G ↔ ∀ g, g * z = z * g := iff.rfl -instance decidable_mem_center [decidable_eq G] [fintype G] : decidable_pred (∈ center G) := -λ _, decidable_of_iff' _ mem_center_iff +instance decidable_mem_center (z : G) [decidable (∀ g, g * z = z * g)] : decidable (z ∈ center G) := +decidable_of_iff' _ mem_center_iff @[to_additive] instance center_characteristic : (center G).characteristic := begin @@ -1623,16 +1484,6 @@ def set_normalizer (S : set G) : subgroup G := inv_mem' := λ a (ha : ∀ n, n ∈ S ↔ a * n * a⁻¹ ∈ S) n, by { rw [ha (a⁻¹ * n * a⁻¹⁻¹)], simp [mul_assoc] } } -lemma mem_normalizer_fintype {S : set G} [finite S] {x : G} - (h : ∀ n, n ∈ S → x * n * x⁻¹ ∈ S) : x ∈ subgroup.set_normalizer S := -by haveI := classical.prop_decidable; casesI nonempty_fintype S; -haveI := set.fintype_image S (λ n, x * n * x⁻¹); exact -λ n, ⟨h n, λ h₁, -have heq : (λ n, x * n * x⁻¹) '' S = S := set.eq_of_subset_of_card_le - (λ n ⟨y, hy⟩, hy.2 ▸ h y hy.1) (by rw set.card_image_of_injective S conj_injective), -have x * n * x⁻¹ ∈ (λ n, x * n * x⁻¹) '' S := heq.symm ▸ h₁, -let ⟨y, hy⟩ := this in conj_injective hy.2 ▸ hy.1⟩ - variable {H} @[to_additive] lemma mem_normalizer_iff {g : G} : g ∈ H.normalizer ↔ ∀ h, h ∈ H ↔ g * h * g⁻¹ ∈ H := @@ -1962,11 +1813,6 @@ open subgroup def range (f : G →* N) : subgroup N := subgroup.copy ((⊤ : subgroup G).map f) (set.range f) (by simp [set.ext_iff]) -@[to_additive] -instance decidable_mem_range (f : G →* N) [fintype G] [decidable_eq N] : - decidable_pred (∈ f.range) := -λ x, fintype.decidable_exists_fintype - @[simp, to_additive] lemma coe_range (f : G →* N) : (f.range : set N) = set.range f := rfl @@ -2200,32 +2046,6 @@ lemma map_closure (f : G →* N) (s : set G) : set.image_preimage.l_comm_of_u_comm (subgroup.gc_map_comap f) (subgroup.gi N).gc (subgroup.gi G).gc (λ t, rfl) --- this instance can't go just after the definition of `mrange` because `fintype` is --- not imported at that stage - -/-- The range of a finite monoid under a monoid homomorphism is finite. -Note: this instance can form a diamond with `subtype.fintype` in the -presence of `fintype N`. -/ -@[to_additive "The range of a finite additive monoid under an additive monoid homomorphism is -finite. - -Note: this instance can form a diamond with `subtype.fintype` or `subgroup.fintype` in the -presence of `fintype N`."] -instance fintype_mrange {M N : Type*} [monoid M] [monoid N] [fintype M] [decidable_eq N] - (f : M →* N) : fintype (mrange f) := -set.fintype_range f - -/-- The range of a finite group under a group homomorphism is finite. - -Note: this instance can form a diamond with `subtype.fintype` or `subgroup.fintype` in the -presence of `fintype N`. -/ -@[to_additive "The range of a finite additive group under an additive group homomorphism is finite. - -Note: this instance can form a diamond with `subtype.fintype` or `subgroup.fintype` in the -presence of `fintype N`."] -instance fintype_range [fintype G] [decidable_eq N] (f : G →* N) : fintype (range f) := -set.fintype_range f - end monoid_hom namespace subgroup @@ -2613,157 +2433,6 @@ hH.comap _ instance subgroup.normal_subgroup_of {H N : subgroup G} [N.normal] : (N.subgroup_of H).normal := subgroup.normal_comap _ -namespace subgroup - -/-- The subgroup generated by an element. -/ -def zpowers (g : G) : subgroup G := -subgroup.copy (zpowers_hom G g).range (set.range ((^) g : ℤ → G)) rfl - -@[simp] lemma mem_zpowers (g : G) : g ∈ zpowers g := ⟨1, zpow_one _⟩ - -lemma zpowers_eq_closure (g : G) : zpowers g = closure {g} := -by { ext, exact mem_closure_singleton.symm } - -@[simp] lemma range_zpowers_hom (g : G) : (zpowers_hom G g).range = zpowers g := rfl - -lemma zpowers_subset {a : G} {K : subgroup G} (h : a ∈ K) : zpowers a ≤ K := -λ x hx, match x, hx with _, ⟨i, rfl⟩ := K.zpow_mem h i end - -lemma mem_zpowers_iff {g h : G} : - h ∈ zpowers g ↔ ∃ (k : ℤ), g ^ k = h := -iff.rfl - -@[simp] lemma zpow_mem_zpowers (g : G) (k : ℤ) : g^k ∈ zpowers g := -mem_zpowers_iff.mpr ⟨k, rfl⟩ - -@[simp] lemma npow_mem_zpowers (g : G) (k : ℕ) : g^k ∈ zpowers g := -(zpow_coe_nat g k) ▸ zpow_mem_zpowers g k - -@[simp] lemma forall_zpowers {x : G} {p : zpowers x → Prop} : - (∀ g, p g) ↔ ∀ m : ℤ, p ⟨x ^ m, m, rfl⟩ := -set.forall_subtype_range_iff - -@[simp] lemma exists_zpowers {x : G} {p : zpowers x → Prop} : - (∃ g, p g) ↔ ∃ m : ℤ, p ⟨x ^ m, m, rfl⟩ := -set.exists_subtype_range_iff - -lemma forall_mem_zpowers {x : G} {p : G → Prop} : - (∀ g ∈ zpowers x, p g) ↔ ∀ m : ℤ, p (x ^ m) := -set.forall_range_iff - -lemma exists_mem_zpowers {x : G} {p : G → Prop} : - (∃ g ∈ zpowers x, p g) ↔ ∃ m : ℤ, p (x ^ m) := -set.exists_range_iff - -instance (a : G) : countable (zpowers a) := -((zpowers_hom G a).range_restrict_surjective.comp multiplicative.of_add.surjective).countable - -end subgroup - -namespace add_subgroup - -/-- The subgroup generated by an element. -/ -def zmultiples (a : A) : add_subgroup A := -add_subgroup.copy (zmultiples_hom A a).range (set.range ((• a) : ℤ → A)) rfl - -@[simp] lemma range_zmultiples_hom (a : A) : (zmultiples_hom A a).range = zmultiples a := rfl - -attribute [to_additive add_subgroup.zmultiples] subgroup.zpowers -attribute [to_additive add_subgroup.mem_zmultiples] subgroup.mem_zpowers -attribute [to_additive add_subgroup.zmultiples_eq_closure] subgroup.zpowers_eq_closure -attribute [to_additive add_subgroup.range_zmultiples_hom] subgroup.range_zpowers_hom -attribute [to_additive add_subgroup.zmultiples_subset] subgroup.zpowers_subset -attribute [to_additive add_subgroup.mem_zmultiples_iff] subgroup.mem_zpowers_iff -attribute [to_additive add_subgroup.zsmul_mem_zmultiples] subgroup.zpow_mem_zpowers -attribute [to_additive add_subgroup.nsmul_mem_zmultiples] subgroup.npow_mem_zpowers -attribute [to_additive add_subgroup.forall_zmultiples] subgroup.forall_zpowers -attribute [to_additive add_subgroup.forall_mem_zmultiples] subgroup.forall_mem_zpowers -attribute [to_additive add_subgroup.exists_zmultiples] subgroup.exists_zpowers -attribute [to_additive add_subgroup.exists_mem_zmultiples] subgroup.exists_mem_zpowers - -instance (a : A) : countable (zmultiples a) := -(zmultiples_hom A a).range_restrict_surjective.countable - -section ring - -variables {R : Type*} [ring R] (r : R) (k : ℤ) - -@[simp] lemma int_cast_mul_mem_zmultiples : - ↑(k : ℤ) * r ∈ zmultiples r := -by simpa only [← zsmul_eq_mul] using zsmul_mem_zmultiples r k - -@[simp] lemma int_cast_mem_zmultiples_one : - ↑(k : ℤ) ∈ zmultiples (1 : R) := -mem_zmultiples_iff.mp ⟨k, by simp⟩ - -end ring - -end add_subgroup - -@[simp, to_additive map_zmultiples] lemma monoid_hom.map_zpowers (f : G →* N) (x : G) : - (subgroup.zpowers x).map f = subgroup.zpowers (f x) := -by rw [subgroup.zpowers_eq_closure, subgroup.zpowers_eq_closure, f.map_closure, set.image_singleton] - -lemma int.mem_zmultiples_iff {a b : ℤ} : - b ∈ add_subgroup.zmultiples a ↔ a ∣ b := -exists_congr (λ k, by rw [mul_comm, eq_comm, ← smul_eq_mul]) - -lemma of_mul_image_zpowers_eq_zmultiples_of_mul { x : G } : - additive.of_mul '' ((subgroup.zpowers x) : set G) = add_subgroup.zmultiples (additive.of_mul x) := -begin - ext y, - split, - { rintro ⟨z, ⟨m, hm⟩, hz2⟩, - use m, - simp only, - rwa [← of_mul_zpow, hm] }, - { rintros ⟨n, hn⟩, - refine ⟨x ^ n, ⟨n, rfl⟩, _⟩, - rwa of_mul_zpow } -end - -lemma of_add_image_zmultiples_eq_zpowers_of_add {x : A} : - multiplicative.of_add '' ((add_subgroup.zmultiples x) : set A) = - subgroup.zpowers (multiplicative.of_add x) := -begin - symmetry, - rw equiv.eq_image_iff_symm_image_eq, - exact of_mul_image_zpowers_eq_zmultiples_of_mul, -end - -namespace subgroup - -@[to_additive zmultiples_is_commutative] -instance zpowers_is_commutative (g : G) : (zpowers g).is_commutative := -⟨⟨λ ⟨_, _, h₁⟩ ⟨_, _, h₂⟩, by rw [subtype.ext_iff, coe_mul, coe_mul, - subtype.coe_mk, subtype.coe_mk, ←h₁, ←h₂, zpow_mul_comm]⟩⟩ - -@[simp, to_additive zmultiples_le] -lemma zpowers_le {g : G} {H : subgroup G} : zpowers g ≤ H ↔ g ∈ H := -by rw [zpowers_eq_closure, closure_le, set.singleton_subset_iff, set_like.mem_coe] - -@[simp, to_additive zmultiples_eq_bot] lemma zpowers_eq_bot {g : G} : zpowers g = ⊥ ↔ g = 1 := -by rw [eq_bot_iff, zpowers_le, mem_bot] - -@[simp, to_additive zmultiples_zero_eq_bot] lemma zpowers_one_eq_bot : - subgroup.zpowers (1 : G) = ⊥ := -subgroup.zpowers_eq_bot.mpr rfl - -@[to_additive] lemma centralizer_closure (S : set G) : - (closure S).centralizer = ⨅ g ∈ S, (zpowers g).centralizer := -le_antisymm (le_infi $ λ g, le_infi $ λ hg, centralizer_le $ zpowers_le.2 $ subset_closure hg) - $ le_centralizer_iff.1 $ (closure_le _).2 - $ λ g, set_like.mem_coe.2 ∘ zpowers_le.1 ∘ le_centralizer_iff.1 ∘ infi_le_of_le g ∘ infi_le _ - -@[to_additive] lemma center_eq_infi (S : set G) (hS : closure S = ⊤) : - center G = ⨅ g ∈ S, centralizer (zpowers g) := -by rw [←centralizer_top, ←hS, centralizer_closure] - -@[to_additive] lemma center_eq_infi' (S : set G) (hS : closure S = ⊤) : - center G = ⨅ g : S, centralizer (zpowers g) := -by rw [center_eq_infi S hS, ←infi_subtype''] - -end subgroup namespace monoid_hom @@ -2857,49 +2526,6 @@ end⟩ end subgroup -section -variables (G) (A) - -/-- A `group` is simple when it has exactly two normal `subgroup`s. -/ -class is_simple_group extends nontrivial G : Prop := -(eq_bot_or_eq_top_of_normal : ∀ H : subgroup G, H.normal → H = ⊥ ∨ H = ⊤) - -/-- An `add_group` is simple when it has exactly two normal `add_subgroup`s. -/ -class is_simple_add_group extends nontrivial A : Prop := -(eq_bot_or_eq_top_of_normal : ∀ H : add_subgroup A, H.normal → H = ⊥ ∨ H = ⊤) - -attribute [to_additive] is_simple_group - -variables {G} {A} - -@[to_additive] -lemma subgroup.normal.eq_bot_or_eq_top [is_simple_group G] {H : subgroup G} (Hn : H.normal) : - H = ⊥ ∨ H = ⊤ := -is_simple_group.eq_bot_or_eq_top_of_normal H Hn - -namespace is_simple_group - -@[to_additive] -instance {C : Type*} [comm_group C] [is_simple_group C] : - is_simple_order (subgroup C) := -⟨λ H, H.normal_of_comm.eq_bot_or_eq_top⟩ - -open _root_.subgroup - -@[to_additive] -lemma is_simple_group_of_surjective {H : Type*} [group H] [is_simple_group G] - [nontrivial H] (f : G →* H) (hf : function.surjective f) : - is_simple_group H := -⟨nontrivial.exists_pair_ne, λ H iH, begin - refine ((iH.comap f).eq_bot_or_eq_top).imp (λ h, _) (λ h, _), - { rw [←map_bot f, ←h, map_comap_eq_self_of_surjective hf] }, - { rw [←comap_top f] at h, exact comap_injective hf h } -end⟩ - -end is_simple_group - -end - namespace subgroup section subgroup_normal @@ -3008,37 +2634,6 @@ begin ←subtype.ext_iff, ←subtype.ext_iff, eq_comm, ←prod.ext_iff] at hxy, end -/-- `finset.noncomm_prod` is “injective” in `f` if `f` maps into independent subgroups. This -generalizes (one direction of) `subgroup.disjoint_iff_mul_eq_one`. -/ -@[to_additive "`finset.noncomm_sum` is “injective” in `f` if `f` maps into independent subgroups. -This generalizes (one direction of) `add_subgroup.disjoint_iff_add_eq_zero`. "] -lemma eq_one_of_noncomm_prod_eq_one_of_independent {ι : Type*} (s : finset ι) (f : ι → G) (comm) - (K : ι → subgroup G) (hind : complete_lattice.independent K) (hmem : ∀ (x ∈ s), f x ∈ K x) - (heq1 : s.noncomm_prod f comm = 1) : ∀ (i ∈ s), f i = 1 := -begin - classical, - revert heq1, - induction s using finset.induction_on with i s hnmem ih, - { simp, }, - { have hcomm := comm.mono (finset.coe_subset.2 $ finset.subset_insert _ _), - simp only [finset.forall_mem_insert] at hmem, - have hmem_bsupr: s.noncomm_prod f hcomm ∈ ⨆ (i ∈ (s : set ι)), K i, - { refine subgroup.noncomm_prod_mem _ _ _, - intros x hx, - have : K x ≤ ⨆ (i ∈ (s : set ι)), K i := le_supr₂ x hx, - exact this (hmem.2 x hx), }, - intro heq1, - rw finset.noncomm_prod_insert_of_not_mem _ _ _ _ hnmem at heq1, - have hnmem' : i ∉ (s : set ι), by simpa, - obtain ⟨heq1i : f i = 1, heq1S : s.noncomm_prod f _ = 1⟩ := - subgroup.disjoint_iff_mul_eq_one.mp (hind.disjoint_bsupr hnmem') hmem.1 hmem_bsupr heq1, - intros i h, - simp only [finset.mem_insert] at h, - rcases h with ⟨rfl | _⟩, - { exact heq1i }, - { exact ih hcomm hmem.2 heq1S _ h } } -end - end subgroup namespace is_conj @@ -3072,153 +2667,4 @@ end end is_conj -/-! ### Actions by `subgroup`s - -These are just copies of the definitions about `submonoid` starting from `submonoid.mul_action`. --/ -section actions - -namespace subgroup - -variables {α β : Type*} - -/-- The action by a subgroup is the action by the underlying group. -/ -@[to_additive /-"The additive action by an add_subgroup is the action by the underlying -add_group. "-/] -instance [mul_action G α] (S : subgroup G) : mul_action S α := -S.to_submonoid.mul_action - -@[to_additive] -lemma smul_def [mul_action G α] {S : subgroup G} (g : S) (m : α) : g • m = (g : G) • m := rfl - -@[to_additive] -instance smul_comm_class_left - [mul_action G β] [has_smul α β] [smul_comm_class G α β] (S : subgroup G) : - smul_comm_class S α β := -S.to_submonoid.smul_comm_class_left - -@[to_additive] -instance smul_comm_class_right - [has_smul α β] [mul_action G β] [smul_comm_class α G β] (S : subgroup G) : - smul_comm_class α S β := -S.to_submonoid.smul_comm_class_right - -/-- Note that this provides `is_scalar_tower S G G` which is needed by `smul_mul_assoc`. -/ -instance - [has_smul α β] [mul_action G α] [mul_action G β] [is_scalar_tower G α β] (S : subgroup G) : - is_scalar_tower S α β := -S.to_submonoid.is_scalar_tower - -instance [mul_action G α] [has_faithful_smul G α] (S : subgroup G) : - has_faithful_smul S α := -S.to_submonoid.has_faithful_smul - -/-- The action by a subgroup is the action by the underlying group. -/ -instance [add_monoid α] [distrib_mul_action G α] (S : subgroup G) : distrib_mul_action S α := -S.to_submonoid.distrib_mul_action - -/-- The action by a subgroup is the action by the underlying group. -/ -instance [monoid α] [mul_distrib_mul_action G α] (S : subgroup G) : mul_distrib_mul_action S α := -S.to_submonoid.mul_distrib_mul_action - -/-- The center of a group acts commutatively on that group. -/ -instance center.smul_comm_class_left : smul_comm_class (center G) G G := -submonoid.center.smul_comm_class_left - -/-- The center of a group acts commutatively on that group. -/ -instance center.smul_comm_class_right : smul_comm_class G (center G) G := -submonoid.center.smul_comm_class_right - -end subgroup - -end actions - -/-! ### Mul-opposite subgroups -/ - -section mul_opposite - -namespace subgroup - -/-- A subgroup `H` of `G` determines a subgroup `H.opposite` of the opposite group `Gᵐᵒᵖ`. -/ -@[to_additive "An additive subgroup `H` of `G` determines an additive subgroup `H.opposite` of the - opposite additive group `Gᵃᵒᵖ`."] -def opposite : subgroup G ≃ subgroup Gᵐᵒᵖ := -{ to_fun := λ H, { carrier := mul_opposite.unop ⁻¹' (H : set G), - one_mem' := H.one_mem, - mul_mem' := λ a b ha hb, H.mul_mem hb ha, - inv_mem' := λ a, H.inv_mem }, - inv_fun := λ H, { carrier := mul_opposite.op ⁻¹' (H : set Gᵐᵒᵖ), - one_mem' := H.one_mem, - mul_mem' := λ a b ha hb, H.mul_mem hb ha, - inv_mem' := λ a, H.inv_mem }, - left_inv := λ H, set_like.coe_injective rfl, - right_inv := λ H, set_like.coe_injective rfl } - -/-- Bijection between a subgroup `H` and its opposite. -/ -@[to_additive "Bijection between an additive subgroup `H` and its opposite.", simps] -def opposite_equiv (H : subgroup G) : H ≃ H.opposite := -mul_opposite.op_equiv.subtype_equiv $ λ _, iff.rfl - -@[to_additive] instance (H : subgroup G) [encodable H] : encodable H.opposite := -encodable.of_equiv H H.opposite_equiv.symm - -@[to_additive] instance (H : subgroup G) [countable H] : countable H.opposite := -countable.of_equiv H H.opposite_equiv - -@[to_additive] lemma smul_opposite_mul {H : subgroup G} (x g : G) (h : H.opposite) : - h • (g * x) = g * (h • x) := -begin - cases h, - simp [(•), mul_assoc], -end - -end subgroup - -end mul_opposite - -/-! ### Saturated subgroups -/ - -section saturated - -namespace subgroup - -/-- A subgroup `H` of `G` is *saturated* if for all `n : ℕ` and `g : G` with `g^n ∈ H` -we have `n = 0` or `g ∈ H`. -/ -@[to_additive "An additive subgroup `H` of `G` is *saturated* if -for all `n : ℕ` and `g : G` with `n•g ∈ H` we have `n = 0` or `g ∈ H`."] -def saturated (H : subgroup G) : Prop := ∀ ⦃n g⦄, g ^ n ∈ H → n = 0 ∨ g ∈ H - -@[to_additive] lemma saturated_iff_npow {H : subgroup G} : - saturated H ↔ (∀ (n : ℕ) (g : G), g ^ n ∈ H → n = 0 ∨ g ∈ H) := iff.rfl - -@[to_additive] lemma saturated_iff_zpow {H : subgroup G} : - saturated H ↔ (∀ (n : ℤ) (g : G), g ^ n ∈ H → n = 0 ∨ g ∈ H) := -begin - split, - { rintros hH ⟨n⟩ g hgn, - { simp only [int.coe_nat_eq_zero, int.of_nat_eq_coe, zpow_coe_nat] at hgn ⊢, - exact hH hgn }, - { suffices : g ^ (n+1) ∈ H, - { refine (hH this).imp _ id, simp only [is_empty.forall_iff, nat.succ_ne_zero], }, - simpa only [inv_mem_iff, zpow_neg_succ_of_nat] using hgn, } }, - { intros h n g hgn, - specialize h n g, - simp only [int.coe_nat_eq_zero, zpow_coe_nat] at h, - apply h hgn } -end - -end subgroup - -namespace add_subgroup - -lemma ker_saturated {A₁ A₂ : Type*} [add_comm_group A₁] [add_comm_group A₂] - [no_zero_smul_divisors ℕ A₂] (f : A₁ →+ A₂) : - (f.ker).saturated := -begin - intros n g hg, - simpa only [f.mem_ker, nsmul_eq_smul, f.map_nsmul, smul_eq_zero] using hg -end - -end add_subgroup - -end saturated +assert_not_exists multiset diff --git a/src/group_theory/subgroup/finite.lean b/src/group_theory/subgroup/finite.lean new file mode 100644 index 0000000000000..b8de0c882f0db --- /dev/null +++ b/src/group_theory/subgroup/finite.lean @@ -0,0 +1,238 @@ +/- +Copyright (c) 2020 Kexing Ying. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Kexing Ying +-/ + +import data.set.finite +import group_theory.subgroup.basic +import group_theory.submonoid.membership + +/-! +# Subgroups + +This file provides some result on multiplicative and additive subgroups in the finite context. + +## Tags +subgroup, subgroups +-/ + +open_locale big_operators + +variables {G : Type*} [group G] +variables {A : Type*} [add_group A] + +namespace subgroup + +@[to_additive] +instance (K : subgroup G) [d : decidable_pred (∈ K)] [fintype G] : fintype K := +show fintype {g : G // g ∈ K}, from infer_instance + +@[to_additive] +instance (K : subgroup G) [finite G] : finite K := +subtype.finite + +end subgroup + +/-! +### Conversion to/from `additive`/`multiplicative` +-/ +namespace subgroup + +variables (H K : subgroup G) + +/-- Product of a list of elements in a subgroup is in the subgroup. -/ +@[to_additive "Sum of a list of elements in an `add_subgroup` is in the `add_subgroup`."] +protected lemma list_prod_mem {l : list G} : (∀ x ∈ l, x ∈ K) → l.prod ∈ K := +list_prod_mem + +/-- Product of a multiset of elements in a subgroup of a `comm_group` is in the subgroup. -/ +@[to_additive "Sum of a multiset of elements in an `add_subgroup` of an `add_comm_group` +is in the `add_subgroup`."] +protected lemma multiset_prod_mem {G} [comm_group G] (K : subgroup G) (g : multiset G) : + (∀ a ∈ g, a ∈ K) → g.prod ∈ K := multiset_prod_mem g + +@[to_additive] +lemma multiset_noncomm_prod_mem (K : subgroup G) (g : multiset G) (comm) : + (∀ a ∈ g, a ∈ K) → g.noncomm_prod comm ∈ K := +K.to_submonoid.multiset_noncomm_prod_mem g comm + +/-- Product of elements of a subgroup of a `comm_group` indexed by a `finset` is in the + subgroup. -/ +@[to_additive "Sum of elements in an `add_subgroup` of an `add_comm_group` indexed by a `finset` +is in the `add_subgroup`."] +protected lemma prod_mem {G : Type*} [comm_group G] (K : subgroup G) + {ι : Type*} {t : finset ι} {f : ι → G} (h : ∀ c ∈ t, f c ∈ K) : + ∏ c in t, f c ∈ K := +prod_mem h + +@[to_additive] +lemma noncomm_prod_mem (K : subgroup G) {ι : Type*} {t : finset ι} {f : ι → G} (comm) : + (∀ c ∈ t, f c ∈ K) → t.noncomm_prod f comm ∈ K := +K.to_submonoid.noncomm_prod_mem t f comm + +@[simp, norm_cast, to_additive] theorem coe_list_prod (l : list H) : + (l.prod : G) = (l.map coe).prod := +submonoid_class.coe_list_prod l + +@[simp, norm_cast, to_additive] theorem coe_multiset_prod {G} [comm_group G] (H : subgroup G) + (m : multiset H) : (m.prod : G) = (m.map coe).prod := +submonoid_class.coe_multiset_prod m + +@[simp, norm_cast, to_additive] theorem coe_finset_prod {ι G} [comm_group G] (H : subgroup G) + (f : ι → H) (s : finset ι) : + ↑(∏ i in s, f i) = (∏ i in s, f i : G) := +submonoid_class.coe_finset_prod f s + +@[to_additive] instance fintype_bot : fintype (⊥ : subgroup G) := ⟨{1}, +by {rintro ⟨x, ⟨hx⟩⟩, exact finset.mem_singleton_self _}⟩ + +/- curly brackets `{}` are used here instead of instance brackets `[]` because + the instance in a goal is often not the same as the one inferred by type class inference. -/ +@[simp, to_additive] lemma card_bot {_ : fintype ↥(⊥ : subgroup G)} : + fintype.card (⊥ : subgroup G) = 1 := +fintype.card_eq_one_iff.2 + ⟨⟨(1 : G), set.mem_singleton 1⟩, λ ⟨y, hy⟩, subtype.eq $ subgroup.mem_bot.1 hy⟩ + +@[to_additive] lemma eq_top_of_card_eq [fintype H] [fintype G] + (h : fintype.card H = fintype.card G) : H = ⊤ := +begin + haveI : fintype (H : set G) := ‹fintype H›, + rw [set_like.ext'_iff, coe_top, ← finset.coe_univ, ← (H : set G).coe_to_finset, finset.coe_inj, + ← finset.card_eq_iff_eq_univ, ← h, set.to_finset_card], + congr +end + +@[to_additive] lemma eq_top_of_le_card [fintype H] [fintype G] + (h : fintype.card G ≤ fintype.card H) : H = ⊤ := +eq_top_of_card_eq H (le_antisymm (fintype.card_le_of_injective coe subtype.coe_injective) h) + +@[to_additive] lemma eq_bot_of_card_le [fintype H] (h : fintype.card H ≤ 1) : H = ⊥ := +let _ := fintype.card_le_one_iff_subsingleton.mp h in by exactI eq_bot_of_subsingleton H + +@[to_additive] lemma eq_bot_of_card_eq [fintype H] (h : fintype.card H = 1) : H = ⊥ := +H.eq_bot_of_card_le (le_of_eq h) + +@[to_additive] lemma card_le_one_iff_eq_bot [fintype H] : fintype.card H ≤ 1 ↔ H = ⊥ := +⟨λ h, (eq_bot_iff_forall _).2 + (λ x hx, by simpa [subtype.ext_iff] using fintype.card_le_one_iff.1 h ⟨x, hx⟩ 1), + λ h, by simp [h]⟩ + +@[to_additive] lemma one_lt_card_iff_ne_bot [fintype H] : 1 < fintype.card H ↔ H ≠ ⊥ := +lt_iff_not_le.trans H.card_le_one_iff_eq_bot.not + +end subgroup + +namespace subgroup + +section pi + +open set + +variables {η : Type*} {f : η → Type*} [∀ i, group (f i)] + +@[to_additive] +lemma pi_mem_of_mul_single_mem_aux [decidable_eq η] (I : finset η) {H : subgroup (Π i, f i) } + (x : Π i, f i) (h1 : ∀ i, i ∉ I → x i = 1) (h2 : ∀ i, i ∈ I → pi.mul_single i (x i) ∈ H ) : + x ∈ H := +begin + induction I using finset.induction_on with i I hnmem ih generalizing x, + { convert one_mem H, + ext i, + exact (h1 i (not_mem_empty i)) }, + { have : x = function.update x i 1 * pi.mul_single i (x i), + { ext j, + by_cases heq : j = i, + { subst heq, simp, }, + { simp [heq], }, }, + rw this, clear this, + apply mul_mem, + { apply ih; clear ih, + { intros j hj, + by_cases heq : j = i, + { subst heq, simp, }, + { simp [heq], apply h1 j, simpa [heq] using hj, } }, + { intros j hj, + have : j ≠ i, by { rintro rfl, contradiction }, + simp [this], + exact h2 _ (finset.mem_insert_of_mem hj), }, }, + { apply h2, simp, } } +end + +@[to_additive] +lemma pi_mem_of_mul_single_mem [finite η] [decidable_eq η] {H : subgroup (Π i, f i)} + (x : Π i, f i) (h : ∀ i, pi.mul_single i (x i) ∈ H) : x ∈ H := +by { casesI nonempty_fintype η, + exact pi_mem_of_mul_single_mem_aux finset.univ x (by simp) (λ i _, h i) } + +/-- For finite index types, the `subgroup.pi` is generated by the embeddings of the groups. -/ +@[to_additive "For finite index types, the `subgroup.pi` is generated by the embeddings of the +additive groups."] +lemma pi_le_iff [decidable_eq η] [finite η] {H : Π i, subgroup (f i)} {J : subgroup (Π i, f i)} : + pi univ H ≤ J ↔ ∀ i : η, map (monoid_hom.single f i) (H i) ≤ J := +begin + split, + { rintros h i _ ⟨x, hx, rfl⟩, apply h, simpa using hx }, + { exact λ h x hx, pi_mem_of_mul_single_mem x (λ i, h i (mem_map_of_mem _ (hx i trivial))), } +end + +end pi + +end subgroup + +namespace subgroup + +section normalizer + +lemma mem_normalizer_fintype {S : set G} [finite S] {x : G} + (h : ∀ n, n ∈ S → x * n * x⁻¹ ∈ S) : x ∈ subgroup.set_normalizer S := +by haveI := classical.prop_decidable; casesI nonempty_fintype S; +haveI := set.fintype_image S (λ n, x * n * x⁻¹); exact +λ n, ⟨h n, λ h₁, +have heq : (λ n, x * n * x⁻¹) '' S = S := set.eq_of_subset_of_card_le + (λ n ⟨y, hy⟩, hy.2 ▸ h y hy.1) (by rw set.card_image_of_injective S conj_injective), +have x * n * x⁻¹ ∈ (λ n, x * n * x⁻¹) '' S := heq.symm ▸ h₁, +let ⟨y, hy⟩ := this in conj_injective hy.2 ▸ hy.1⟩ + +end normalizer + +end subgroup + +namespace monoid_hom + +variables {N : Type*} [group N] + +open subgroup + +@[to_additive] +instance decidable_mem_range (f : G →* N) [fintype G] [decidable_eq N] : + decidable_pred (∈ f.range) := +λ x, fintype.decidable_exists_fintype + +-- this instance can't go just after the definition of `mrange` because `fintype` is +-- not imported at that stage + +/-- The range of a finite monoid under a monoid homomorphism is finite. +Note: this instance can form a diamond with `subtype.fintype` in the +presence of `fintype N`. -/ +@[to_additive "The range of a finite additive monoid under an additive monoid homomorphism is +finite. + +Note: this instance can form a diamond with `subtype.fintype` or `subgroup.fintype` in the +presence of `fintype N`."] +instance fintype_mrange {M N : Type*} [monoid M] [monoid N] [fintype M] [decidable_eq N] + (f : M →* N) : fintype (mrange f) := +set.fintype_range f + +/-- The range of a finite group under a group homomorphism is finite. + +Note: this instance can form a diamond with `subtype.fintype` or `subgroup.fintype` in the +presence of `fintype N`. -/ +@[to_additive "The range of a finite additive group under an additive group homomorphism is finite. + +Note: this instance can form a diamond with `subtype.fintype` or `subgroup.fintype` in the +presence of `fintype N`."] +instance fintype_range [fintype G] [decidable_eq N] (f : G →* N) : fintype (range f) := +set.fintype_range f + +end monoid_hom diff --git a/src/group_theory/subgroup/mul_opposite.lean b/src/group_theory/subgroup/mul_opposite.lean new file mode 100644 index 0000000000000..93f590238c684 --- /dev/null +++ b/src/group_theory/subgroup/mul_opposite.lean @@ -0,0 +1,54 @@ +/- +Copyright (c) 2022 Alex Kontorovich. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Alex Kontorovich +-/ + +import group_theory.subgroup.actions + +/-! +# Mul-opposite subgroups + +## Tags +subgroup, subgroups + +-/ + +variables {G : Type*} [group G] + +namespace subgroup + +/-- A subgroup `H` of `G` determines a subgroup `H.opposite` of the opposite group `Gᵐᵒᵖ`. -/ +@[to_additive "An additive subgroup `H` of `G` determines an additive subgroup `H.opposite` of the + opposite additive group `Gᵃᵒᵖ`."] +def opposite : subgroup G ≃ subgroup Gᵐᵒᵖ := +{ to_fun := λ H, { carrier := mul_opposite.unop ⁻¹' (H : set G), + one_mem' := H.one_mem, + mul_mem' := λ a b ha hb, H.mul_mem hb ha, + inv_mem' := λ a, H.inv_mem }, + inv_fun := λ H, { carrier := mul_opposite.op ⁻¹' (H : set Gᵐᵒᵖ), + one_mem' := H.one_mem, + mul_mem' := λ a b ha hb, H.mul_mem hb ha, + inv_mem' := λ a, H.inv_mem }, + left_inv := λ H, set_like.coe_injective rfl, + right_inv := λ H, set_like.coe_injective rfl } + +/-- Bijection between a subgroup `H` and its opposite. -/ +@[to_additive "Bijection between an additive subgroup `H` and its opposite.", simps] +def opposite_equiv (H : subgroup G) : H ≃ H.opposite := +mul_opposite.op_equiv.subtype_equiv $ λ _, iff.rfl + +@[to_additive] instance (H : subgroup G) [encodable H] : encodable H.opposite := +encodable.of_equiv H H.opposite_equiv.symm + +@[to_additive] instance (H : subgroup G) [countable H] : countable H.opposite := +countable.of_equiv H H.opposite_equiv + +@[to_additive] lemma smul_opposite_mul {H : subgroup G} (x g : G) (h : H.opposite) : + h • (g * x) = g * (h • x) := +begin + cases h, + simp [(•), mul_assoc], +end + +end subgroup diff --git a/src/group_theory/subgroup/pointwise.lean b/src/group_theory/subgroup/pointwise.lean index 67d21096238a2..3a565dfb0313f 100644 --- a/src/group_theory/subgroup/pointwise.lean +++ b/src/group_theory/subgroup/pointwise.lean @@ -3,7 +3,7 @@ Copyright (c) 2021 Eric Wieser. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Eric Wieser -/ -import group_theory.subgroup.basic +import group_theory.subgroup.mul_opposite import group_theory.submonoid.pointwise import group_theory.group_action.conj_act diff --git a/src/group_theory/subgroup/saturated.lean b/src/group_theory/subgroup/saturated.lean new file mode 100644 index 0000000000000..e22b4921d7165 --- /dev/null +++ b/src/group_theory/subgroup/saturated.lean @@ -0,0 +1,58 @@ +/- +Copyright (c) 2021 Johan Commelin. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Johan Commelin +-/ + +import group_theory.subgroup.basic + +/-! +# Saturated subgroups + +## Tags +subgroup, subgroups + +-/ + +namespace subgroup + +variables {G : Type*} [group G] + +/-- A subgroup `H` of `G` is *saturated* if for all `n : ℕ` and `g : G` with `g^n ∈ H` +we have `n = 0` or `g ∈ H`. -/ +@[to_additive "An additive subgroup `H` of `G` is *saturated* if +for all `n : ℕ` and `g : G` with `n•g ∈ H` we have `n = 0` or `g ∈ H`."] +def saturated (H : subgroup G) : Prop := ∀ ⦃n g⦄, g ^ n ∈ H → n = 0 ∨ g ∈ H + +@[to_additive] lemma saturated_iff_npow {H : subgroup G} : + saturated H ↔ (∀ (n : ℕ) (g : G), g ^ n ∈ H → n = 0 ∨ g ∈ H) := iff.rfl + +@[to_additive] lemma saturated_iff_zpow {H : subgroup G} : + saturated H ↔ (∀ (n : ℤ) (g : G), g ^ n ∈ H → n = 0 ∨ g ∈ H) := +begin + split, + { rintros hH ⟨n⟩ g hgn, + { simp only [int.coe_nat_eq_zero, int.of_nat_eq_coe, zpow_coe_nat] at hgn ⊢, + exact hH hgn }, + { suffices : g ^ (n+1) ∈ H, + { refine (hH this).imp _ id, simp only [is_empty.forall_iff, nat.succ_ne_zero], }, + simpa only [inv_mem_iff, zpow_neg_succ_of_nat] using hgn, } }, + { intros h n g hgn, + specialize h n g, + simp only [int.coe_nat_eq_zero, zpow_coe_nat] at h, + apply h hgn } +end + +end subgroup + +namespace add_subgroup + +lemma ker_saturated {A₁ A₂ : Type*} [add_comm_group A₁] [add_comm_group A₂] + [no_zero_smul_divisors ℕ A₂] (f : A₁ →+ A₂) : + (f.ker).saturated := +begin + intros n g hg, + simpa only [f.mem_ker, nsmul_eq_smul, f.map_nsmul, smul_eq_zero] using hg +end + +end add_subgroup diff --git a/src/group_theory/subgroup/simple.lean b/src/group_theory/subgroup/simple.lean new file mode 100644 index 0000000000000..f043081c9206e --- /dev/null +++ b/src/group_theory/subgroup/simple.lean @@ -0,0 +1,70 @@ +/- +Copyright (c) 2021 Aaron Anderson. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Aaron Anderson +-/ + +import group_theory.subgroup.actions + +/-! +# Simple groups + +This file defines `is_simple_group G`, a class indicating that a group has exactly two normal +subgroups. + +## Main definitions + +- `is_simple_group G`, a class indicating that a group has exactly two normal subgroups. + +## Tags +subgroup, subgroups + +-/ + +set_option old_structure_cmd true + +variables {G : Type*} [group G] +variables {A : Type*} [add_group A] + +section +variables (G) (A) + +/-- A `group` is simple when it has exactly two normal `subgroup`s. -/ +class is_simple_group extends nontrivial G : Prop := +(eq_bot_or_eq_top_of_normal : ∀ H : subgroup G, H.normal → H = ⊥ ∨ H = ⊤) + +/-- An `add_group` is simple when it has exactly two normal `add_subgroup`s. -/ +class is_simple_add_group extends nontrivial A : Prop := +(eq_bot_or_eq_top_of_normal : ∀ H : add_subgroup A, H.normal → H = ⊥ ∨ H = ⊤) + +attribute [to_additive] is_simple_group + +variables {G} {A} + +@[to_additive] +lemma subgroup.normal.eq_bot_or_eq_top [is_simple_group G] {H : subgroup G} (Hn : H.normal) : + H = ⊥ ∨ H = ⊤ := +is_simple_group.eq_bot_or_eq_top_of_normal H Hn + +namespace is_simple_group + +@[to_additive] +instance {C : Type*} [comm_group C] [is_simple_group C] : + is_simple_order (subgroup C) := +⟨λ H, H.normal_of_comm.eq_bot_or_eq_top⟩ + +open _root_.subgroup + +@[to_additive] +lemma is_simple_group_of_surjective {H : Type*} [group H] [is_simple_group G] + [nontrivial H] (f : G →* H) (hf : function.surjective f) : + is_simple_group H := +⟨nontrivial.exists_pair_ne, λ H iH, begin + refine ((iH.comap f).eq_bot_or_eq_top).imp (λ h, _) (λ h, _), + { rw [←map_bot f, ←h, map_comap_eq_self_of_surjective hf] }, + { rw [←comap_top f] at h, exact comap_injective hf h } +end⟩ + +end is_simple_group + +end diff --git a/src/group_theory/subgroup/zpowers.lean b/src/group_theory/subgroup/zpowers.lean new file mode 100644 index 0000000000000..b7a89fb456edd --- /dev/null +++ b/src/group_theory/subgroup/zpowers.lean @@ -0,0 +1,171 @@ +/- +Copyright (c) 2020 Chris Hughes. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Chris Hughes +-/ + +import group_theory.subgroup.basic + +/-! +# Subgroups generated by an element + +## Tags +subgroup, subgroups + +-/ + +variables {G : Type*} [group G] +variables {A : Type*} [add_group A] +variables {N : Type*} [group N] + +namespace subgroup + +/-- The subgroup generated by an element. -/ +def zpowers (g : G) : subgroup G := +subgroup.copy (zpowers_hom G g).range (set.range ((^) g : ℤ → G)) rfl + +@[simp] lemma mem_zpowers (g : G) : g ∈ zpowers g := ⟨1, zpow_one _⟩ + +lemma zpowers_eq_closure (g : G) : zpowers g = closure {g} := +by { ext, exact mem_closure_singleton.symm } + +@[simp] lemma range_zpowers_hom (g : G) : (zpowers_hom G g).range = zpowers g := rfl + +lemma zpowers_subset {a : G} {K : subgroup G} (h : a ∈ K) : zpowers a ≤ K := +λ x hx, match x, hx with _, ⟨i, rfl⟩ := K.zpow_mem h i end + +lemma mem_zpowers_iff {g h : G} : + h ∈ zpowers g ↔ ∃ (k : ℤ), g ^ k = h := +iff.rfl + +@[simp] lemma zpow_mem_zpowers (g : G) (k : ℤ) : g^k ∈ zpowers g := +mem_zpowers_iff.mpr ⟨k, rfl⟩ + +@[simp] lemma npow_mem_zpowers (g : G) (k : ℕ) : g^k ∈ zpowers g := +(zpow_coe_nat g k) ▸ zpow_mem_zpowers g k + +@[simp] lemma forall_zpowers {x : G} {p : zpowers x → Prop} : + (∀ g, p g) ↔ ∀ m : ℤ, p ⟨x ^ m, m, rfl⟩ := +set.forall_subtype_range_iff + +@[simp] lemma exists_zpowers {x : G} {p : zpowers x → Prop} : + (∃ g, p g) ↔ ∃ m : ℤ, p ⟨x ^ m, m, rfl⟩ := +set.exists_subtype_range_iff + +lemma forall_mem_zpowers {x : G} {p : G → Prop} : + (∀ g ∈ zpowers x, p g) ↔ ∀ m : ℤ, p (x ^ m) := +set.forall_range_iff + +lemma exists_mem_zpowers {x : G} {p : G → Prop} : + (∃ g ∈ zpowers x, p g) ↔ ∃ m : ℤ, p (x ^ m) := +set.exists_range_iff + +instance (a : G) : countable (zpowers a) := +((zpowers_hom G a).range_restrict_surjective.comp multiplicative.of_add.surjective).countable + +end subgroup + +namespace add_subgroup + +/-- The subgroup generated by an element. -/ +def zmultiples (a : A) : add_subgroup A := +add_subgroup.copy (zmultiples_hom A a).range (set.range ((• a) : ℤ → A)) rfl + +@[simp] lemma range_zmultiples_hom (a : A) : (zmultiples_hom A a).range = zmultiples a := rfl + +attribute [to_additive add_subgroup.zmultiples] subgroup.zpowers +attribute [to_additive add_subgroup.mem_zmultiples] subgroup.mem_zpowers +attribute [to_additive add_subgroup.zmultiples_eq_closure] subgroup.zpowers_eq_closure +attribute [to_additive add_subgroup.range_zmultiples_hom] subgroup.range_zpowers_hom +attribute [to_additive add_subgroup.zmultiples_subset] subgroup.zpowers_subset +attribute [to_additive add_subgroup.mem_zmultiples_iff] subgroup.mem_zpowers_iff +attribute [to_additive add_subgroup.zsmul_mem_zmultiples] subgroup.zpow_mem_zpowers +attribute [to_additive add_subgroup.nsmul_mem_zmultiples] subgroup.npow_mem_zpowers +attribute [to_additive add_subgroup.forall_zmultiples] subgroup.forall_zpowers +attribute [to_additive add_subgroup.forall_mem_zmultiples] subgroup.forall_mem_zpowers +attribute [to_additive add_subgroup.exists_zmultiples] subgroup.exists_zpowers +attribute [to_additive add_subgroup.exists_mem_zmultiples] subgroup.exists_mem_zpowers + +instance (a : A) : countable (zmultiples a) := +(zmultiples_hom A a).range_restrict_surjective.countable + +section ring + +variables {R : Type*} [ring R] (r : R) (k : ℤ) + +@[simp] lemma int_cast_mul_mem_zmultiples : + ↑(k : ℤ) * r ∈ zmultiples r := +by simpa only [← zsmul_eq_mul] using zsmul_mem_zmultiples r k + +@[simp] lemma int_cast_mem_zmultiples_one : + ↑(k : ℤ) ∈ zmultiples (1 : R) := +mem_zmultiples_iff.mp ⟨k, by simp⟩ + +end ring + +end add_subgroup + +@[simp, to_additive map_zmultiples] lemma monoid_hom.map_zpowers (f : G →* N) (x : G) : + (subgroup.zpowers x).map f = subgroup.zpowers (f x) := +by rw [subgroup.zpowers_eq_closure, subgroup.zpowers_eq_closure, f.map_closure, set.image_singleton] + +lemma int.mem_zmultiples_iff {a b : ℤ} : + b ∈ add_subgroup.zmultiples a ↔ a ∣ b := +exists_congr (λ k, by rw [mul_comm, eq_comm, ← smul_eq_mul]) + +lemma of_mul_image_zpowers_eq_zmultiples_of_mul { x : G } : + additive.of_mul '' ((subgroup.zpowers x) : set G) = add_subgroup.zmultiples (additive.of_mul x) := +begin + ext y, + split, + { rintro ⟨z, ⟨m, hm⟩, hz2⟩, + use m, + simp only, + rwa [← of_mul_zpow, hm] }, + { rintros ⟨n, hn⟩, + refine ⟨x ^ n, ⟨n, rfl⟩, _⟩, + rwa of_mul_zpow } +end + +lemma of_add_image_zmultiples_eq_zpowers_of_add {x : A} : + multiplicative.of_add '' ((add_subgroup.zmultiples x) : set A) = + subgroup.zpowers (multiplicative.of_add x) := +begin + symmetry, + rw equiv.eq_image_iff_symm_image_eq, + exact of_mul_image_zpowers_eq_zmultiples_of_mul, +end + +namespace subgroup + +@[to_additive zmultiples_is_commutative] +instance zpowers_is_commutative (g : G) : (zpowers g).is_commutative := +⟨⟨λ ⟨_, _, h₁⟩ ⟨_, _, h₂⟩, by rw [subtype.ext_iff, coe_mul, coe_mul, + subtype.coe_mk, subtype.coe_mk, ←h₁, ←h₂, zpow_mul_comm]⟩⟩ + +@[simp, to_additive zmultiples_le] +lemma zpowers_le {g : G} {H : subgroup G} : zpowers g ≤ H ↔ g ∈ H := +by rw [zpowers_eq_closure, closure_le, set.singleton_subset_iff, set_like.mem_coe] + +@[simp, to_additive zmultiples_eq_bot] lemma zpowers_eq_bot {g : G} : zpowers g = ⊥ ↔ g = 1 := +by rw [eq_bot_iff, zpowers_le, mem_bot] + +@[simp, to_additive zmultiples_zero_eq_bot] lemma zpowers_one_eq_bot : + subgroup.zpowers (1 : G) = ⊥ := +subgroup.zpowers_eq_bot.mpr rfl + +@[to_additive] lemma centralizer_closure (S : set G) : + (closure S).centralizer = ⨅ g ∈ S, (zpowers g).centralizer := +le_antisymm (le_infi $ λ g, le_infi $ λ hg, centralizer_le $ zpowers_le.2 $ subset_closure hg) + $ le_centralizer_iff.1 $ (closure_le _).2 + $ λ g, set_like.mem_coe.2 ∘ zpowers_le.1 ∘ le_centralizer_iff.1 ∘ infi_le_of_le g ∘ infi_le _ + +@[to_additive] lemma center_eq_infi (S : set G) (hS : closure S = ⊤) : + center G = ⨅ g ∈ S, centralizer (zpowers g) := +by rw [←centralizer_top, ←hS, centralizer_closure] + +@[to_additive] lemma center_eq_infi' (S : set G) (hS : closure S = ⊤) : + center G = ⨅ g : S, centralizer (zpowers g) := +by rw [center_eq_infi S hS, ←infi_subtype''] + +end subgroup diff --git a/src/linear_algebra/dual.lean b/src/linear_algebra/dual.lean index 425bd69a89e60..632e5effca548 100644 --- a/src/linear_algebra/dual.lean +++ b/src/linear_algebra/dual.lean @@ -854,6 +854,16 @@ begin rw [dual_map_apply, hx, map_zero] end +/-- if a linear map is surjective, then its dual is injective -/ +lemma dual_map_injective_of_surjective {f : M₁ →ₗ[R] M₂} (hf : function.surjective f) : + function.injective f.dual_map := +begin + obtain ⟨g, hg⟩ := hf.has_right_inverse, + intros x y hxy, + ext w, + simpa only [dual_map_apply, hg w] using fun_like.congr_fun hxy (g w), +end + section finite_dimensional variables {K : Type*} [field K] {V₁ : Type*} {V₂ : Type*} @@ -886,6 +896,41 @@ begin exact finrank_range_add_finrank_ker f, end +/-- `f.dual_map` is injective if and only if `f` is surjective -/ +@[simp] lemma dual_map_injective_iff {f : V₁ →ₗ[K] V₂} : + function.injective f.dual_map ↔ function.surjective f := +begin + refine ⟨_, λ h, dual_map_injective_of_surjective h⟩, + rw [← range_eq_top, ← ker_eq_bot], + intro h, + apply finite_dimensional.eq_top_of_finrank_eq, + rw ← finrank_eq_zero at h, + rw [← add_zero (finite_dimensional.finrank K f.range), ← h, + ← linear_map.finrank_range_dual_map_eq_finrank_range, + linear_map.finrank_range_add_finrank_ker, subspace.dual_finrank_eq], +end + +/-- `f.dual_map` is surjective if and only if `f` is injective -/ +@[simp] lemma dual_map_surjective_iff [finite_dimensional K V₁] {f : V₁ →ₗ[K] V₂} : + function.surjective f.dual_map ↔ function.injective f := +begin + rw [← range_eq_top, ← ker_eq_bot], + split, + { intro h, + rw [← finrank_eq_zero, ← add_right_inj (finite_dimensional.finrank K f.range), + add_zero, linear_map.finrank_range_add_finrank_ker, + ← linear_map.finrank_range_dual_map_eq_finrank_range, h, + finrank_top, subspace.dual_finrank_eq], }, + { intro h, + rw [range_dual_map_eq_dual_annihilator_ker, h], + exact submodule.dual_annihilator_bot, }, +end + +/-- `f.dual_map` is bijective if and only if `f` is -/ +@[simp] lemma dual_map_bijective_iff [finite_dimensional K V₁] {f : V₁ →ₗ[K] V₂} : + function.bijective f.dual_map ↔ function.bijective f := +by simp_rw [function.bijective, dual_map_surjective_iff, dual_map_injective_iff, and.comm] + end finite_dimensional section field diff --git a/src/linear_algebra/ray.lean b/src/linear_algebra/ray.lean index 9d9a2801a262f..c12aec99af056 100644 --- a/src/linear_algebra/ray.lean +++ b/src/linear_algebra/ray.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Joseph Myers. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Joseph Myers -/ +import group_theory.subgroup.actions import linear_algebra.linear_independent /-! diff --git a/src/logic/equiv/array.lean b/src/logic/equiv/array.lean new file mode 100644 index 0000000000000..5e1f98c7d71cf --- /dev/null +++ b/src/logic/equiv/array.lean @@ -0,0 +1,52 @@ +/- +Copyright (c) 2018 Mario Carneiro. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Mario Carneiro +-/ +import data.vector.basic +import logic.equiv.list +import control.traversable.equiv + +/-! +# Equivalences involving `array` + +We keep this separate from the file containing `list`-like equivalences as those have no future +in mathlib4. +-/ + +namespace equiv + +/-- The natural equivalence between length-`n` heterogeneous arrays +and dependent functions from `fin n`. -/ +def d_array_equiv_fin {n : ℕ} (α : fin n → Type*) : d_array n α ≃ (Π i, α i) := +⟨d_array.read, d_array.mk, λ ⟨f⟩, rfl, λ f, rfl⟩ + +/-- The natural equivalence between length-`n` arrays and functions from `fin n`. -/ +def array_equiv_fin (n : ℕ) (α : Type*) : array n α ≃ (fin n → α) := +d_array_equiv_fin _ + +/-- The natural equivalence between length-`n` vectors and length-`n` arrays. -/ +def vector_equiv_array (α : Type*) (n : ℕ) : vector α n ≃ array n α := +(vector_equiv_fin _ _).trans (array_equiv_fin _ _).symm + +end equiv + +namespace array +open function +variable {n : ℕ} + +instance : traversable (array n) := +@equiv.traversable (flip vector n) _ (λ α, equiv.vector_equiv_array α n) _ + +instance : is_lawful_traversable (array n) := +@equiv.is_lawful_traversable (flip vector n) _ (λ α, equiv.vector_equiv_array α n) _ _ + +end array + +/-- If `α` is encodable, then so is `array n α`. -/ +instance _root_.array.encodable {α} [encodable α] {n} : encodable (array n α) := +encodable.of_equiv _ (equiv.array_equiv_fin _ _) + +/-- If `α` is countable, then so is `array n α`. -/ +instance _root_.array.countable {α} [countable α] {n} : countable (array n α) := +countable.of_equiv _ (equiv.vector_equiv_array _ _) diff --git a/src/logic/equiv/list.lean b/src/logic/equiv/list.lean index 63cda4d0f5634..71a3d9e03d6fd 100644 --- a/src/logic/equiv/list.lean +++ b/src/logic/equiv/list.lean @@ -3,8 +3,8 @@ Copyright (c) 2018 Mario Carneiro. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Mario Carneiro -/ -import data.array.lemmas import data.finset.sort +import data.vector.basic import logic.denumerable /-! @@ -130,14 +130,6 @@ of_equiv _ (equiv.vector_equiv_fin _ _).symm instance fin_pi (n) (π : fin n → Type*) [∀ i, encodable (π i)] : encodable (Π i, π i) := of_equiv _ (equiv.pi_equiv_subtype_sigma (fin n) π) -/-- If `α` is encodable, then so is `array n α`. -/ -instance _root_.array.encodable [encodable α] {n} : encodable (array n α) := -of_equiv _ (equiv.array_equiv_fin _ _) - -/-- If `α` is countable, then so is `array n α`. -/ -instance _root_.array.countable [countable α] {n} : countable (array n α) := -countable.of_equiv _ (equiv.vector_equiv_array _ _) - /-- If `α` is encodable, then so is `finset α`. -/ instance _root_.finset.encodable [encodable α] : encodable (finset α) := by haveI := decidable_eq_of_encodable α; exact diff --git a/src/order/extension/well.lean b/src/order/extension/well.lean index 76c290e56f5d0..b684cf4c87f44 100644 --- a/src/order/extension/well.lean +++ b/src/order/extension/well.lean @@ -44,7 +44,7 @@ arbitrary well-order to serve as a tiebreak between two elements of same rank. noncomputable def well_order_extension : linear_order α := let l : linear_order α := is_well_order.linear_order well_ordering_rel in by exactI @linear_order.lift' α (ordinal ×ₗ α) _ - (λ a : α, (well_founded.rank.{u u} hwf a, a)) (λ _ _, congr_arg prod.snd) + (λ a : α, (well_founded.rank.{u} hwf a, a)) (λ _ _, congr_arg prod.snd) instance well_order_extension.is_well_founded_lt : is_well_founded α hwf.well_order_extension.lt := ⟨inv_image.wf _ $ prod.lex_wf ordinal.well_founded_lt.wf well_ordering_rel.is_well_order.wf⟩ diff --git a/src/order/grade.lean b/src/order/grade.lean index 9d693b59419ca..1f0c44cac0b23 100644 --- a/src/order/grade.lean +++ b/src/order/grade.lean @@ -9,6 +9,9 @@ import data.int.succ_pred /-! # Graded orders +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines graded orders, also known as ranked orders. A `𝕆`-graded order is an order `α` equipped with a distinguished "grade" function `α → 𝕆` which diff --git a/src/order/hom/bounded.lean b/src/order/hom/bounded.lean index 359ab4bfd14d0..cb58486e39e07 100644 --- a/src/order/hom/bounded.lean +++ b/src/order/hom/bounded.lean @@ -9,6 +9,9 @@ import order.bounded_order /-! # Bounded order homomorphisms +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines (bounded) order homomorphisms. We use the `fun_like` design, so each type of morphisms has a companion typeclass which is meant to diff --git a/src/order/hom/complete_lattice.lean b/src/order/hom/complete_lattice.lean index e37431e776b62..6f1a44ed63de6 100644 --- a/src/order/hom/complete_lattice.lean +++ b/src/order/hom/complete_lattice.lean @@ -627,17 +627,17 @@ See also `complete_lattice_hom.set_preimage`. -/ map_rel_iff' := λ s t, ⟨λ h, by simpa using @monotone_image _ _ e.symm _ _ h, λ h, monotone_image h⟩ } -section +variables [complete_lattice α] (x : α × α) -variables (α) [complete_lattice α] +/-- The map `(a, b) ↦ a ⊔ b` as a `Sup_hom`. -/ +def sup_Sup_hom : Sup_hom (α × α) α := +{ to_fun := λ x, x.1 ⊔ x.2, + map_Sup' := λ s, by simp_rw [prod.fst_Sup, prod.snd_Sup, Sup_image, supr_sup_eq] } /-- The map `(a, b) ↦ a ⊓ b` as an `Inf_hom`. -/ -@[simps] def inf_Inf_hom : Inf_hom (α × α) α := +def inf_Inf_hom : Inf_hom (α × α) α := { to_fun := λ x, x.1 ⊓ x.2, - map_Inf' := λ s, by simp_rw [prod.fst_Inf, prod.snd_Inf, Inf_image, infi_inf_eq], } + map_Inf' := λ s, by simp_rw [prod.fst_Inf, prod.snd_Inf, Inf_image, infi_inf_eq] } -/-- The map `(a, b) ↦ a ⊔ b` as a `Sup_hom`. -/ -@[simps] def sup_Sup_hom : Sup_hom (α × α) α := -{ to_fun := λ x, x.1 ⊔ x.2, map_Sup' := (inf_Inf_hom αᵒᵈ).map_Inf' } - -end +@[simp, norm_cast] lemma sup_Sup_hom_apply : sup_Sup_hom x = x.1 ⊔ x.2 := rfl +@[simp, norm_cast] lemma inf_Inf_hom_apply : inf_Inf_hom x = x.1 ⊓ x.2 := rfl diff --git a/src/order/sup_indep.lean b/src/order/sup_indep.lean index 70c982f5ccad4..487c9e623553d 100644 --- a/src/order/sup_indep.lean +++ b/src/order/sup_indep.lean @@ -10,6 +10,9 @@ import data.fintype.basic /-! # Supremum independence +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + In this file, we define supremum independence of indexed sets. An indexed family `f : ι → α` is sup-independent if, for all `a`, `f a` and the supremum of the rest are disjoint. diff --git a/src/order/upper_lower.lean b/src/order/upper_lower.lean index 61b3c5a6c2b40..2bcc1e90c8a6b 100644 --- a/src/order/upper_lower.lean +++ b/src/order/upper_lower.lean @@ -757,13 +757,15 @@ def upper_closure (s : set α) : upper_set α := def lower_closure (s : set α) : lower_set α := ⟨{x | ∃ a ∈ s, x ≤ a}, λ x y h, Exists₂.imp $ λ a _, h.trans⟩ --- We do not tag those two as `simp` to respect the abstraction. -@[norm_cast] lemma coe_upper_closure (s : set α) : ↑(upper_closure s) = {x | ∃ a ∈ s, a ≤ x} := rfl -@[norm_cast] lemma coe_lower_closure (s : set α) : ↑(lower_closure s) = {x | ∃ a ∈ s, x ≤ a} := rfl - @[simp] lemma mem_upper_closure : x ∈ upper_closure s ↔ ∃ a ∈ s, a ≤ x := iff.rfl @[simp] lemma mem_lower_closure : x ∈ lower_closure s ↔ ∃ a ∈ s, x ≤ a := iff.rfl +-- We do not tag those two as `simp` to respect the abstraction. +@[norm_cast] lemma coe_upper_closure (s : set α) : ↑(upper_closure s) = ⋃ a ∈ s, Ici a := +by { ext, simp } +@[norm_cast] lemma coe_lower_closure (s : set α) : ↑(lower_closure s) = ⋃ a ∈ s, Iic a := +by { ext, simp } + lemma subset_upper_closure : s ⊆ upper_closure s := λ x hx, ⟨x, hx, le_rfl⟩ lemma subset_lower_closure : s ⊆ lower_closure s := λ x hx, ⟨x, hx, le_rfl⟩ diff --git a/src/ring_theory/adjoin_root.lean b/src/ring_theory/adjoin_root.lean index 4d8c81e2496b2..6fb17d1e38b18 100644 --- a/src/ring_theory/adjoin_root.lean +++ b/src/ring_theory/adjoin_root.lean @@ -334,8 +334,17 @@ basis.of_equiv_fun exact (degree_eq_nat_degree $ hg.ne_zero).symm ▸ degree_sum_fin_lt _ }, end} --- This was moved after the definition to prevent a timeout -attribute [simps] power_basis_aux' +/-- This lemma could be autogenerated by `@[simps]` but unfortunately that would require +unfolding that causes a timeout. -/ +@[simp] lemma power_basis_aux'_repr_symm_apply (hg : g.monic) (c : fin g.nat_degree →₀ R) : + (power_basis_aux' hg).repr.symm c = mk g (∑ (i : fin _), monomial i (c i)) := rfl + +/-- This lemma could be autogenerated by `@[simps]` but unfortunately that would require +unfolding that causes a timeout. -/ +@[simp] theorem power_basis_aux'_repr_apply_to_fun (hg : g.monic) (f : adjoin_root g) + (i : fin g.nat_degree) : + (power_basis_aux' hg).repr f i = (mod_by_monic_hom hg f).coeff ↑i := +rfl /-- The power basis `1, root g, ..., root g ^ (d - 1)` for `adjoin_root g`, where `g` is a monic polynomial of degree `d`. -/ diff --git a/src/ring_theory/congruence.lean b/src/ring_theory/congruence.lean index a79ed4a5a51b9..f0c339ea28545 100644 --- a/src/ring_theory/congruence.lean +++ b/src/ring_theory/congruence.lean @@ -12,6 +12,9 @@ import group_theory.congruence /-! # Congruence relations on rings +> THIS FILE IS SYNCHRONIZED WITH MATHLIB4. +> Any changes to this file require a corresponding PR to mathlib4. + This file defines congruence relations on rings, which extend `con` and `add_con` on monoids and additive monoids. diff --git a/src/set_theory/ordinal/arithmetic.lean b/src/set_theory/ordinal/arithmetic.lean index 810d62bb70b87..6057463f2d35b 100644 --- a/src/set_theory/ordinal/arithmetic.lean +++ b/src/set_theory/ordinal/arithmetic.lean @@ -54,10 +54,10 @@ open function cardinal set equiv order open_locale classical cardinal ordinal universes u v w -variables {α : Type*} {β : Type*} {γ : Type*} - {r : α → α → Prop} {s : β → β → Prop} {t : γ → γ → Prop} namespace ordinal +variables {α : Type*} {β : Type*} {γ : Type*} + {r : α → α → Prop} {s : β → β → Prop} {t : γ → γ → Prop} /-! ### Further properties of addition on ordinals -/ @@ -2332,17 +2332,18 @@ meta def positivity_opow : expr → tactic strictness end tactic +variables {α : Type u} {r : α → α → Prop} {a b : α} + namespace acc -variables {a b : α} /-- The rank of an element `a` accessible under a relation `r` is defined inductively as the smallest ordinal greater than the ranks of all elements below it (i.e. elements `b` such that `r b a`). -/ -noncomputable def rank (h : acc r a) : ordinal := -acc.rec_on h $ λ a h ih, ordinal.sup $ λ b : {b // r b a}, order.succ $ ih b b.2 +noncomputable def rank (h : acc r a) : ordinal.{u} := +acc.rec_on h $ λ a h ih, ordinal.sup.{u u} $ λ b : {b // r b a}, order.succ $ ih b b.2 lemma rank_eq (h : acc r a) : - h.rank = ordinal.sup (λ b : {b // r b a}, order.succ (h.inv b.2).rank) := + h.rank = ordinal.sup.{u u} (λ b : {b // r b a}, order.succ (h.inv b.2).rank) := by { change (acc.intro a $ λ _, h.inv).rank = _, refl } /-- if `r a b` then the rank of `a` is less than the rank of `b`. -/ @@ -2352,15 +2353,15 @@ lemma rank_lt_of_rel (hb : acc r b) (h : r a b) : (hb.inv h).rank < hb.rank := end acc namespace well_founded -variables (hwf : well_founded r) {a b : α} +variables (hwf : well_founded r) include hwf /-- The rank of an element `a` under a well-founded relation `r` is defined inductively as the smallest ordinal greater than the ranks of all elements below it (i.e. elements `b` such that `r b a`). -/ -noncomputable def rank (a : α) : ordinal := (hwf.apply a).rank +noncomputable def rank (a : α) : ordinal.{u} := (hwf.apply a).rank -lemma rank_eq : hwf.rank a = ordinal.sup (λ b : {b // r b a}, order.succ $ hwf.rank b) := +lemma rank_eq : hwf.rank a = ordinal.sup.{u u} (λ b : {b // r b a}, order.succ $ hwf.rank b) := by { rw [rank, acc.rank_eq], refl } lemma rank_lt_of_rel (h : r a b) : hwf.rank a < hwf.rank b := acc.rank_lt_of_rel _ h diff --git a/src/tactic/ring.lean b/src/tactic/ring.lean index 59995a1c7e3d9..4dcb10f74db63 100644 --- a/src/tactic/ring.lean +++ b/src/tactic/ring.lean @@ -499,16 +499,21 @@ meta def eval (norm_atom : expr → tactic (expr × expr)) : expr → ring_m (ho p' ← ic_lift $ λ ic, ic.mk_app ``unfold_div [e₁, e₂, e', p], return (e', p')) (eval_norm_atom norm_atom e) -| e@`(@has_pow.pow _ _ %%P %%e₁ %%e₂) := do - (e₂', p₂) ← lift $ norm_num.derive e₂ <|> refl_conv e₂, - match e₂'.to_nat, P with - | some k, `(monoid.has_pow) := do - (e₁', p₁) ← eval e₁, - (e', p') ← eval_pow e₁' (e₂, k), - p ← ic_lift $ λ ic, ic.mk_app ``subst_into_pow [e₁, e₂, e₁', e₂', e', p₁, p₂, p'], - return (e', p) - | _, _ := eval_norm_atom norm_atom e - end +| e@`(@has_pow.pow _ _ %%inst %%e₁ %%e₂) := mcond + (succeeds (do + inst' ← ic_lift $ λ ic, ic.mk_app ``monoid.has_pow [], + lift $ is_def_eq inst inst')) + (do + (e₂', p₂) ← lift $ norm_num.derive e₂ <|> refl_conv e₂, + match e₂'.to_nat with + | some k := do + (e₁', p₁) ← eval e₁, + (e', p') ← eval_pow e₁' (e₂, k), + p ← ic_lift $ λ ic, ic.mk_app ``subst_into_pow [e₁, e₂, e₁', e₂', e', p₁, p₂, p'], + return (e', p) + | _ := eval_norm_atom norm_atom e + end) + (eval_norm_atom norm_atom e) | e := match e.to_nat with | some n := (const e n).refl_conv | none := eval_norm_atom norm_atom e diff --git a/src/topology/algebra/infinite_sum.lean b/src/topology/algebra/infinite_sum.lean index 2ba8e2083616a..7eaa1e5dc1354 100644 --- a/src/topology/algebra/infinite_sum.lean +++ b/src/topology/algebra/infinite_sum.lean @@ -405,6 +405,35 @@ lemma has_sum.sigma_of_has_sum [t3_space α] {γ : β → Type*} {f : (Σ b:β, has_sum f a := by simpa [(hf'.has_sum.sigma hf).unique ha] using hf'.has_sum +/-- Version of `has_sum.update` for `add_comm_monoid` rather than `add_comm_group`. +Rather than showing that `f.update` has a specific sum in terms of `has_sum`, +it gives a relationship between the sums of `f` and `f.update` given that both exist. -/ +lemma has_sum.update' {α β : Type*} [topological_space α] [add_comm_monoid α] [t2_space α] + [has_continuous_add α] {f : β → α} {a a' : α} (hf : has_sum f a) + (b : β) (x : α) (hf' : has_sum (f.update b x) a') : a + x = a' + f b := +begin + have : ∀ b', f b' + ite (b' = b) x 0 = f.update b x b' + ite (b' = b) (f b) 0, + { intro b', + split_ifs with hb', + { simpa only [function.update_apply, hb', eq_self_iff_true] using add_comm (f b) x }, + { simp only [function.update_apply, hb', if_false] } }, + have h := hf.add ((has_sum_ite_eq b x)), + simp_rw this at h, + exact has_sum.unique h (hf'.add (has_sum_ite_eq b (f b))) +end + +/-- Version of `has_sum_ite_sub_has_sum` for `add_comm_monoid` rather than `add_comm_group`. +Rather than showing that the `ite` expression has a specific sum in terms of `has_sum`, +it gives a relationship between the sums of `f` and `ite (n = b) 0 (f n)` given that both exist. -/ +lemma eq_add_of_has_sum_ite {α β : Type*} [topological_space α] [add_comm_monoid α] + [t2_space α] [has_continuous_add α] {f : β → α} {a : α} (hf : has_sum f a) (b : β) (a' : α) + (hf' : has_sum (λ n, ite (n = b) 0 (f n)) a') : a = a' + f b := +begin + refine (add_zero a).symm.trans (hf.update' b 0 _), + convert hf', + exact funext (f.update_apply b 0), +end + end has_sum section tsum @@ -446,6 +475,13 @@ lemma tsum_eq_sum {f : β → α} {s : finset β} (hf : ∀b∉s, f b = 0) : ∑' b, f b = ∑ b in s, f b := (has_sum_sum_of_ne_finset_zero hf).tsum_eq +lemma sum_eq_tsum_indicator (f : β → α) (s : finset β) : + ∑ x in s, f x = ∑' x, set.indicator ↑s f x := +have ∀ x ∉ s, set.indicator ↑s f x = 0, +from λ x hx, set.indicator_apply_eq_zero.2 (λ hx', (hx $ finset.mem_coe.1 hx').elim), +(finset.sum_congr rfl (λ x hx, (set.indicator_apply_eq_self.2 $ + λ hx', (hx' $ finset.mem_coe.2 hx).elim).symm)).trans (tsum_eq_sum this).symm + lemma tsum_congr {α β : Type*} [add_comm_monoid α] [topological_space α] {f g : β → α} (hfg : ∀ b, f b = g b) : ∑' b, f b = ∑' b, g b := congr_arg tsum (funext hfg) @@ -569,6 +605,19 @@ lemma tsum_sum {f : γ → β → α} {s : finset γ} (hf : ∀i∈s, summable ( ∑'b, ∑ i in s, f i b = ∑ i in s, ∑'b, f i b := (has_sum_sum $ assume i hi, (hf i hi).has_sum).tsum_eq +/-- Version of `tsum_eq_add_tsum_ite` for `add_comm_monoid` rather than `add_comm_group`. +Requires a different convergence assumption involving `function.update`. -/ +lemma tsum_eq_add_tsum_ite' {f : β → α} (b : β) (hf : summable (f.update b 0)) : + ∑' x, f x = f b + ∑' x, ite (x = b) 0 (f x) := +calc ∑' x, f x = ∑' x, ((ite (x = b) (f x) 0) + (f.update b 0 x)) : + tsum_congr (λ n, by split_ifs; simp [function.update_apply, h]) + ... = ∑' x, ite (x = b) (f x) 0 + ∑' x, f.update b 0 x : + tsum_add ⟨ite (b = b) (f b) 0, has_sum_single b (λ b hb, if_neg hb)⟩ (hf) + ... = (ite (b = b) (f b) 0) + ∑' x, f.update b 0 x : + by { congr, exact (tsum_eq_single b (λ b' hb', if_neg hb')) } + ... = f b + ∑' x, ite (x = b) 0 (f x) : + by simp only [function.update, eq_self_iff_true, if_true, eq_rec_constant, dite_eq_ite] + variables [add_comm_monoid δ] [topological_space δ] [t3_space δ] [has_continuous_add δ] lemma tsum_sigma' {γ : β → Type*} {f : (Σb:β, γ b) → δ} (h₁ : ∀b, summable (λc, f ⟨b, c⟩)) @@ -798,7 +847,7 @@ lemma set.finite.summable_compl_iff {s : set β} (hs : s.finite) : summable (f ∘ coe : sᶜ → α) ↔ summable f := (hs.summable f).summable_compl_iff -lemma has_sum_ite_eq_extract [decidable_eq β] (hf : has_sum f a) (b : β) : +lemma has_sum_ite_sub_has_sum [decidable_eq β] (hf : has_sum f a) (b : β) : has_sum (λ n, ite (n = b) 0 (f n)) (a - f b) := begin convert hf.update b 0 using 1, @@ -824,12 +873,12 @@ lemma sum_add_tsum_compl {s : finset β} (hf : summable f) : ((s.has_sum f).add_compl (s.summable_compl_iff.2 hf).has_sum).tsum_eq.symm /-- Let `f : β → α` be a sequence with summable series and let `b ∈ β` be an index. -Lemma `tsum_ite_eq_extract` writes `Σ f n` as the sum of `f b` plus the series of the +Lemma `tsum_eq_add_tsum_ite` writes `Σ f n` as the sum of `f b` plus the series of the remaining terms. -/ -lemma tsum_ite_eq_extract [decidable_eq β] (hf : summable f) (b : β) : +lemma tsum_eq_add_tsum_ite [decidable_eq β] (hf : summable f) (b : β) : ∑' n, f n = f b + ∑' n, ite (n = b) 0 (f n) := begin - rw (has_sum_ite_eq_extract hf.has_sum b).tsum_eq, + rw (has_sum_ite_sub_has_sum hf.has_sum b).tsum_eq, exact (add_sub_cancel'_right _ _).symm, end diff --git a/src/topology/algebra/module/basic.lean b/src/topology/algebra/module/basic.lean index 014bc026c3e3b..e13655b155398 100644 --- a/src/topology/algebra/module/basic.lean +++ b/src/topology/algebra/module/basic.lean @@ -262,6 +262,27 @@ end end closure +section pi + +lemma linear_map.continuous_on_pi {ι : Type*} {R : Type*} {M : Type*} [finite ι] [semiring R] + [topological_space R] [add_comm_monoid M] [module R M] [topological_space M] + [has_continuous_add M] [has_continuous_smul R M] (f : (ι → R) →ₗ[R] M) : + continuous f := +begin + casesI nonempty_fintype ι, + classical, + -- for the proof, write `f` in the standard basis, and use that each coordinate is a continuous + -- function. + have : (f : (ι → R) → M) = + (λx, ∑ i : ι, x i • (f (λ j, if i = j then 1 else 0))), + by { ext x, exact f.pi_apply_eq_sum_univ x }, + rw this, + refine continuous_finset_sum _ (λi hi, _), + exact (continuous_apply i).smul continuous_const +end + +end pi + /-- Continuous linear maps between modules. We only put the type classes that are necessary for the definition, although in applications `M` and `M₂` will be topological modules over the topological ring `R`. -/ diff --git a/src/topology/algebra/module/finite_dimension.lean b/src/topology/algebra/module/finite_dimension.lean index fece5d836a471..312648ad2a9ca 100644 --- a/src/topology/algebra/module/finite_dimension.lean +++ b/src/topology/algebra/module/finite_dimension.lean @@ -46,29 +46,6 @@ noncomputable theory open set finite_dimensional topological_space filter open_locale big_operators -section semiring - -variables {ι 𝕜 F : Type*} [finite ι] [semiring 𝕜] [topological_space 𝕜] - [add_comm_monoid F] [module 𝕜 F] [topological_space F] - [has_continuous_add F] [has_continuous_smul 𝕜 F] - -/-- A linear map on `ι → 𝕜` (where `ι` is finite) is continuous -/ -lemma linear_map.continuous_on_pi (f : (ι → 𝕜) →ₗ[𝕜] F) : continuous f := -begin - casesI nonempty_fintype ι, - classical, - -- for the proof, write `f` in the standard basis, and use that each coordinate is a continuous - -- function. - have : (f : (ι → 𝕜) → F) = - (λx, ∑ i : ι, x i • (f (λ j, if i = j then 1 else 0))), - by { ext x, exact f.pi_apply_eq_sum_univ x }, - rw this, - refine continuous_finset_sum _ (λi hi, _), - exact (continuous_apply i).smul continuous_const -end - -end semiring - section field variables {𝕜 E F : Type*} [field 𝕜] [topological_space 𝕜] [add_comm_group E] [module 𝕜 E] diff --git a/src/topology/basic.lean b/src/topology/basic.lean index 101947cac8602..e3784ae39ea07 100644 --- a/src/topology/basic.lean +++ b/src/topology/basic.lean @@ -11,7 +11,7 @@ import order.filter.lift /-! # Basic theory of topological spaces. -The main definition is the type class `topological space α` which endows a type `α` with a topology. +The main definition is the type class `topological_space α` which endows a type `α` with a topology. Then `set α` gets predicates `is_open`, `is_closed` and functions `interior`, `closure` and `frontier`. Each point `x` of `α` gets a neighborhood filter `𝓝 x`. A filter `F` on `α` has `x` as a cluster point if `cluster_pt x F : 𝓝 x ⊓ F ≠ ⊥`. A map `f : ι → α` clusters at `x` diff --git a/src/topology/instances/ennreal.lean b/src/topology/instances/ennreal.lean index 8d5af6c912c16..5bf8b42b77410 100644 --- a/src/topology/instances/ennreal.lean +++ b/src/topology/instances/ennreal.lean @@ -964,6 +964,9 @@ begin exact tsum_bUnion_le _ _ _ end +lemma tsum_eq_add_tsum_ite {f : β → ℝ≥0∞} (b : β) : ∑' x, f x = f b + ∑' x, ite (x = b) 0 (f x) := +tsum_eq_add_tsum_ite' b ennreal.summable + lemma tsum_add_one_eq_top {f : ℕ → ℝ≥0∞} (hf : ∑' n, f n = ∞) (hf0 : f 0 ≠ ∞) : ∑' n, f (n + 1) = ∞ := begin @@ -1205,6 +1208,14 @@ lemma tsum_pos {g : α → ℝ≥0} (hg : summable g) (i : α) (hi : 0 < g i) : 0 < ∑' b, g b := by { rw ← tsum_zero, exact tsum_lt_tsum (λ a, zero_le _) hi hg } +lemma tsum_eq_add_tsum_ite {f : α → ℝ≥0} (hf : summable f) (i : α) : + ∑' x, f x = f i + ∑' x, ite (x = i) 0 (f x) := +begin + refine tsum_eq_add_tsum_ite' i (nnreal.summable_of_le (λ i', _) hf), + rw [function.update_apply], + split_ifs; simp only [zero_le', le_rfl] +end + end nnreal namespace ennreal diff --git a/src/topology/metric_space/metrizable.lean b/src/topology/metric_space/metrizable.lean index d158d8e7691aa..250b0a3467d31 100644 --- a/src/topology/metric_space/metrizable.lean +++ b/src/topology/metric_space/metrizable.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Yury Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury Kudryashov -/ +import analysis.specific_limits.basic import topology.urysohns_lemma import topology.continuous_function.bounded import topology.uniform_space.cauchy diff --git a/src/topology/order/lower_topology.lean b/src/topology/order/lower_topology.lean new file mode 100644 index 0000000000000..0412ba2aa4092 --- /dev/null +++ b/src/topology/order/lower_topology.lean @@ -0,0 +1,230 @@ +/- +Copyright (c) 2023 Christopher Hoskin. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Christopher Hoskin +-/ +import topology.order.lattice + +/-! +# Lower topology + +This file introduces the lower topology on a preorder as the topology generated by the complements +of the closed intervals to infinity. + +## Main statements + +- `lower_topology.t0_space` - the lower topology on a partial order is T₀ +- `is_topological_basis.is_topological_basis` - the complements of the upper closures of finite + subsets form a basis for the lower topology +- `lower_topology.to_has_continuous_inf` - the inf map is continuous with respect to the lower + topology + +## Implementation notes + +A type synonym `with_lower_topology` is introduced and for a preorder `α`, `with_lower_topology α` +is made an instance of `topological_space` by the topology generated by the complements of the +closed intervals to infinity. + +We define a mixin class `lower_topology` for the class of types which are both a preorder and a +topology and where the topology is generated by the complements of the closed intervals to infinity. +It is shown that `with_lower_topology α` is an instance of `lower_topology`. + +## Motivation + +The lower topology is used with the `Scott` topology to define the Lawson topology. The restriction +of the lower topology to the spectrum of a complete lattice coincides with the hull-kernel topology. + +## References + +* [Gierz et al, *A Compendium of Continuous Lattices*][GierzEtAl1980] + +## Tags + +lower topology, preorder +-/ + +variables (α β : Type*) + +open set topological_space + +/-- +Type synonym for a preorder equipped with the lower topology +-/ +def with_lower_topology := α + +variables {α β} + +namespace with_lower_topology + +/-- `to_lower` is the identity function to the `with_lower_topology` of a type. -/ +@[pattern] def to_lower : α ≃ with_lower_topology α := equiv.refl _ + +/-- `of_lower` is the identity function from the `with_lower_topology` of a type. -/ +@[pattern] def of_lower : with_lower_topology α ≃ α := equiv.refl _ + +@[simp] lemma to_with_lower_topology_symm_eq : (@to_lower α).symm = of_lower := rfl +@[simp] lemma of_with_lower_topology_symm_eq : (@of_lower α).symm = to_lower := rfl +@[simp] lemma to_lower_of_lower (a : with_lower_topology α) : to_lower (of_lower a) = a := rfl +@[simp] lemma of_lower_to_lower (a : α) : of_lower (to_lower a) = a := rfl +@[simp] lemma to_lower_inj {a b : α} : to_lower a = to_lower b ↔ a = b := iff.rfl +@[simp] lemma of_lower_inj {a b : with_lower_topology α} : of_lower a = of_lower b ↔ a = b := +iff.rfl + +/-- A recursor for `with_lower_topology`. Use as `induction x using with_lower_topology.rec`. -/ +protected def rec {β : with_lower_topology α → Sort*} + (h : Π a, β (to_lower a)) : Π a, β a := λ a, h (of_lower a) + +instance [nonempty α] : nonempty (with_lower_topology α) := ‹nonempty α› +instance [inhabited α] : inhabited (with_lower_topology α) := ‹inhabited α› + +variables [preorder α] + +instance : preorder (with_lower_topology α) := ‹preorder α› + +instance : topological_space (with_lower_topology α) := generate_from {s | ∃ a, (Ici a)ᶜ = s} + +lemma is_open_preimage_of_lower (S : set α) : + is_open (with_lower_topology.of_lower ⁻¹' S) ↔ + (generate_from {s : set α | ∃ (a : α), (Ici a)ᶜ = s}).is_open S := iff.rfl + +lemma is_open_def (T : set (with_lower_topology α)) : + is_open T ↔ (generate_from {s : set α | ∃ (a : α), (Ici a)ᶜ = s}).is_open + (with_lower_topology.to_lower ⁻¹' T) := iff.rfl + +end with_lower_topology + +/-- +The lower topology is the topology generated by the complements of the closed intervals to infinity. +-/ +class lower_topology (α : Type*) [t : topological_space α] [preorder α] : Prop := +(topology_eq_lower_topology [] : t = generate_from {s | ∃ a, (Ici a)ᶜ = s}) + +instance [preorder α] : lower_topology (with_lower_topology α) := ⟨rfl⟩ + +namespace lower_topology + +/-- The complements of the upper closures of finite sets are a collection of lower sets +which form a basis for the lower topology. -/ +def lower_basis (α : Type*) [preorder α] := +{s : set α | ∃ t : set α, t.finite ∧ (upper_closure t : set α)ᶜ = s} + +section preorder +variables [preorder α] [topological_space α] [lower_topology α] {s : set α} + +/-- If `α` is equipped with the lower topology, then it is homeomorphic to `with_lower_topology α`. +-/ +def with_lower_topology_homeomorph : with_lower_topology α ≃ₜ α := +{ continuous_to_fun := by { convert continuous_id, apply topology_eq_lower_topology }, + continuous_inv_fun := by { convert ← continuous_id, apply topology_eq_lower_topology }, + ..with_lower_topology.of_lower } + +lemma is_open_iff_generate_Ici_compl : is_open s ↔ generate_open {t | ∃ a, (Ici a)ᶜ = t} s := +by rw topology_eq_lower_topology α; refl + +/-- Left-closed right-infinite intervals [a, ∞) are closed in the lower topology. -/ +lemma is_closed_Ici (a : α) : is_closed (Ici a) := +is_open_compl_iff.1 $ is_open_iff_generate_Ici_compl.2 $ generate_open.basic _ ⟨a, rfl⟩ + +/-- The upper closure of a finite set is closed in the lower topology. -/ +lemma is_closed_upper_closure (h : s.finite) : is_closed (upper_closure s : set α) := +begin + simp only [← upper_set.infi_Ici, upper_set.coe_infi], + exact is_closed_bUnion h (λ a h₁, is_closed_Ici a), +end + +/-- Every set open in the lower topology is a lower set. -/ +lemma is_lower_set_of_is_open (h : is_open s) : is_lower_set s := +begin + rw is_open_iff_generate_Ici_compl at h, + induction h, + case generate_open.basic : u h { obtain ⟨a, rfl⟩ := h, exact (is_upper_set_Ici a).compl }, + case univ : { exact is_lower_set_univ }, + case inter : u v hu1 hv1 hu2 hv2 { exact hu2.inter hv2 }, + case sUnion : _ _ ih { exact is_lower_set_sUnion ih }, +end + +lemma is_upper_set_of_is_closed (h : is_closed s) : is_upper_set s := +is_lower_set_compl.1 $ is_lower_set_of_is_open h.is_open_compl + +/-- +The closure of a singleton `{a}` in the lower topology is the left-closed right-infinite interval +[a, ∞). +-/ +@[simp] lemma closure_singleton (a : α) : closure {a} = Ici a := +subset_antisymm (closure_minimal (λ b h, h.ge) $ is_closed_Ici a) $ + (is_upper_set_of_is_closed is_closed_closure).Ici_subset $ subset_closure rfl + +protected lemma is_topological_basis : + is_topological_basis (lower_basis α) := +begin + convert is_topological_basis_of_subbasis (topology_eq_lower_topology α), + simp_rw [lower_basis, coe_upper_closure, compl_Union], + ext s, + split, + { rintro ⟨F, hF, rfl⟩, + refine ⟨(λ a, (Ici a)ᶜ) '' F, ⟨hF.image _, image_subset_iff.2 $ λ _ _, ⟨_, rfl⟩⟩, _⟩, + rw sInter_image }, + { rintro ⟨F, ⟨hF, hs⟩, rfl⟩, + haveI := hF.to_subtype, + rw [subset_def, subtype.forall'] at hs, + choose f hf using hs, + exact ⟨_, finite_range f, by simp_rw [bInter_range, hf, sInter_eq_Inter]⟩ } +end + +end preorder + +section partial_order +variables [partial_order α] [topological_space α] [lower_topology α] + +/-- +The lower topology on a partial order is T₀. +-/ +@[priority 90] -- see Note [lower instance priority] +instance : t0_space α := +(t0_space_iff_inseparable α).2 $ λ x y h, Ici_injective $ + by simpa only [inseparable_iff_closure_eq, closure_singleton] using h + +end partial_order +end lower_topology + +instance [preorder α] [topological_space α] [lower_topology α] [order_bot α] + [preorder β] [topological_space β] [lower_topology β] [order_bot β] : lower_topology (α × β) := +{ topology_eq_lower_topology := + begin + refine le_antisymm (le_generate_from _) _, + { rintro _ ⟨x, rfl⟩, + exact ((lower_topology.is_closed_Ici _).prod $ + lower_topology.is_closed_Ici _).is_open_compl }, + rw [(lower_topology.is_topological_basis.prod + lower_topology.is_topological_basis).eq_generate_from, + le_generate_from_iff_subset_is_open, image2_subset_iff], + rintro _ ⟨s, hs, rfl⟩ _ ⟨t, ht, rfl⟩, + dsimp, + simp_rw [coe_upper_closure, compl_Union, prod_eq, preimage_Inter, preimage_compl], + -- Note: `refine` doesn't work here because it tries using `prod.topological_space`. + apply (is_open_bInter hs $ λ a _, _).inter (is_open_bInter ht $ λ b _, _), + { exact generate_open.basic _ ⟨(a, ⊥), by simp [Ici_prod_eq, prod_univ]⟩ }, + { exact generate_open.basic _ ⟨(⊥, b), by simp [Ici_prod_eq, univ_prod]⟩ }, + all_goals { apply_instance }, + end } + +section complete_lattice +variables [complete_lattice α] [complete_lattice β] [topological_space α] [lower_topology α] + [topological_space β] [lower_topology β] + +lemma Inf_hom.continuous (f : Inf_hom α β) : continuous f := +begin + convert continuous_generated_from _, + { exact lower_topology.topology_eq_lower_topology β }, + rintro _ ⟨b, rfl⟩, + rw [preimage_compl, is_open_compl_iff], + convert lower_topology.is_closed_Ici (Inf $ f ⁻¹' Ici b), + refine subset_antisymm (λ a, Inf_le) (λ a ha, le_trans _ $ order_hom_class.mono f ha), + simp [map_Inf], +end + +@[priority 90] -- see Note [lower instance priority] +instance lower_topology.to_has_continuous_inf : has_continuous_inf α := +⟨(inf_Inf_hom : Inf_hom (α × α) α).continuous⟩ + +end complete_lattice diff --git a/src/topology/tietze_extension.lean b/src/topology/tietze_extension.lean index def8be075ac0c..abaea016c04ac 100644 --- a/src/topology/tietze_extension.lean +++ b/src/topology/tietze_extension.lean @@ -3,6 +3,7 @@ Copyright (c) 2021 Yury G. Kudryashov. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Yury G. Kudryashov -/ +import analysis.specific_limits.basic import data.set.intervals.iso_Ioo import topology.algebra.order.monotone_continuity import topology.urysohns_bounded diff --git a/test/ring.lean b/test/ring.lean index 82c99e1c65feb..f0d15a8954313 100644 --- a/test/ring.lean +++ b/test/ring.lean @@ -95,3 +95,10 @@ begin success_if_fail {{ ring_nf {recursive := ff} }}, ring_nf end + +-- instances do not have to syntactically be `monoid.has_pow` +example {R} [comm_semiring R] (x : ℕ → R) : x ^ 2 = x * x := by ring + +-- even if there's an instance we don't recognize, we treat it as an atom +example {R} [field R] (x : ℕ → R) : + (x ^ (2 : ℤ)) ^ 2 = (x ^ (2 : ℤ)) * (x ^ (2 : ℤ)) := by ring