diff --git a/Mathlib/Algebra/BigOperators/Pi.lean b/Mathlib/Algebra/BigOperators/Pi.lean index 3e6be401f3b0e..0ee8d785dcaa0 100644 --- a/Mathlib/Algebra/BigOperators/Pi.lean +++ b/Mathlib/Algebra/BigOperators/Pi.lean @@ -70,7 +70,7 @@ theorem prod_mk_prod {α β γ : Type _} [CommMonoid α] [CommMonoid β] (s : Fi haveI := Classical.decEq γ Finset.induction_on s rfl (by simp (config := { contextual := true }) [Prod.ext_iff]) #align prod_mk_prod prod_mk_prod -#align sum_mk_sum prod_mk_sum +#align prod_mk_sum prod_mk_sum section Single diff --git a/Mathlib/Algebra/BigOperators/Ring.lean b/Mathlib/Algebra/BigOperators/Ring.lean index 1b2b9b17f8136..41aa030796fe2 100644 --- a/Mathlib/Algebra/BigOperators/Ring.lean +++ b/Mathlib/Algebra/BigOperators/Ring.lean @@ -257,6 +257,7 @@ theorem prod_powerset_insert [DecidableEq α] [CommMonoid β] {s : Finset α} {x rw [← H₃₂] exact ne_insert_of_not_mem _ _ (not_mem_of_mem_powerset_of_not_mem h₁ h) #align finset.prod_powerset_insert Finset.prod_powerset_insert +#align finset.sum_powerset_insert Finset.sum_powerset_insert /-- A product over `powerset s` is equal to the double product over sets of subsets of `s` with `card s = k`, for `k = 1, ..., card s`. -/ @@ -267,6 +268,7 @@ theorem prod_powerset [CommMonoid β] (s : Finset α) (f : Finset α → β) : (∏ t in powerset s, f t) = ∏ j in range (card s + 1), ∏ t in powersetLen j s, f t := by rw [powerset_card_disjUnionᵢ, prod_disjUnionᵢ] #align finset.prod_powerset Finset.prod_powerset +#align finset.sum_powerset Finset.sum_powerset theorem sum_range_succ_mul_sum_range_succ [NonUnitalNonAssocSemiring β] (n k : ℕ) (f g : ℕ → β) : ((∑ i in range (n + 1), f i) * ∑ i in range (k + 1), g i) = diff --git a/Mathlib/Algebra/CharZero/Lemmas.lean b/Mathlib/Algebra/CharZero/Lemmas.lean index d7649f135f7bd..5ba1a124bbba9 100644 --- a/Mathlib/Algebra/CharZero/Lemmas.lean +++ b/Mathlib/Algebra/CharZero/Lemmas.lean @@ -34,6 +34,7 @@ variable {R : Type _} [AddMonoidWithOne R] [CharZero R] def castEmbedding : ℕ ↪ R := ⟨Nat.cast, cast_injective⟩ #align nat.cast_embedding Nat.castEmbedding +#align nat.cast_embedding_apply Nat.castEmbedding_apply @[simp] theorem cast_pow_eq_one {R : Type _} [Semiring R] [CharZero R] (q : ℕ) (n : ℕ) (hn : n ≠ 0) : diff --git a/Mathlib/Algebra/CovariantAndContravariant.lean b/Mathlib/Algebra/CovariantAndContravariant.lean index 9a84c3cdc9ec0..6816175b49cfa 100644 --- a/Mathlib/Algebra/CovariantAndContravariant.lean +++ b/Mathlib/Algebra/CovariantAndContravariant.lean @@ -81,6 +81,7 @@ action of a Type on another one and a relation on the acted-upon Type. See the `CovariantClass` doc-string for its meaning. -/ def Covariant : Prop := ∀ (m) {n₁ n₂}, r n₁ n₂ → r (μ m n₁) (μ m n₂) +#align covariant Covariant /-- `Contravariant` is useful to formulate succintly statements about the interactions between an action of a Type on another one and a relation on the acted-upon Type. @@ -88,6 +89,7 @@ action of a Type on another one and a relation on the acted-upon Type. See the `ContravariantClass` doc-string for its meaning. -/ def Contravariant : Prop := ∀ (m) {n₁ n₂}, r (μ m n₁) (μ m n₂) → r n₁ n₂ +#align contravariant Contravariant /-- Given an action `μ` of a Type `M` on a Type `N` and a relation `r` on `N`, informally, the `CovariantClass` says that "the action `μ` preserves the relation `r`." @@ -104,6 +106,7 @@ class CovariantClass : Prop where /-- For all `m ∈ M` and all elements `n₁, n₂ ∈ N`, if the relation `r` holds for the pair `(n₁, n₂)`, then, the relation `r` also holds for the pair `(μ m n₁, μ m n₂)` -/ protected elim : Covariant M N μ r +#align covariant_class CovariantClass /-- Given an action `μ` of a Type `M` on a Type `N` and a relation `r` on `N`, informally, the `ContravariantClass` says that "if the result of the action `μ` on a pair satisfies the @@ -122,10 +125,12 @@ class ContravariantClass : Prop where pair `(μ m n₁, μ m n₂)` obtained from `(n₁, n₂)` by acting upon it by `m`, then, the relation `r` also holds for the pair `(n₁, n₂)`. -/ protected elim : Contravariant M N μ r +#align contravariant_class ContravariantClass theorem rel_iff_cov [CovariantClass M N μ r] [ContravariantClass M N μ r] (m : M) {a b : N} : r (μ m a) (μ m b) ↔ r a b := ⟨ContravariantClass.elim _, CovariantClass.elim _⟩ +#align rel_iff_cov rel_iff_cov section flip @@ -133,9 +138,11 @@ variable {M N μ r} theorem Covariant.flip (h : Covariant M N μ r) : Covariant M N μ (flip r) := fun a _ _ hbc ↦ h a hbc +#align covariant.flip Covariant.flip theorem Contravariant.flip (h : Contravariant M N μ r) : Contravariant M N μ (flip r) := fun a _ _ hbc ↦ h a hbc +#align contravariant.flip Contravariant.flip end flip @@ -145,6 +152,7 @@ variable {M N μ r} [CovariantClass M N μ r] theorem act_rel_act_of_rel (m : M) {a b : N} (ab : r a b) : r (μ m a) (μ m b) := CovariantClass.elim _ ab +#align act_rel_act_of_rel act_rel_act_of_rel @[to_additive] theorem Group.covariant_iff_contravariant [Group N] : @@ -154,6 +162,8 @@ theorem Group.covariant_iff_contravariant [Group N] : exact h a⁻¹ bc · rw [← inv_mul_cancel_left a b, ← inv_mul_cancel_left a c] at bc exact h a⁻¹ bc +#align group.covariant_iff_contravariant Group.covariant_iff_contravariant +#align add_group.covariant_iff_contravariant AddGroup.covariant_iff_contravariant @[to_additive] instance (priority := 100) Group.covconv [Group N] [CovariantClass N N (· * ·) r] : @@ -168,6 +178,8 @@ theorem Group.covariant_swap_iff_contravariant_swap [Group N] : exact h a⁻¹ bc · rw [← mul_inv_cancel_right b a, ← mul_inv_cancel_right c a] at bc exact h a⁻¹ bc +#align group.covariant_swap_iff_contravariant_swap Group.covariant_swap_iff_contravariant_swap +#align add_group.covariant_swap_iff_contravariant_swap AddGroup.covariant_swap_iff_contravariant_swap @[to_additive] @@ -183,9 +195,11 @@ variable [IsTrans N r] (m n : M) {a b c d : N} -- Lemmas with 3 elements. theorem act_rel_of_rel_of_act_rel (ab : r a b) (rl : r (μ m b) c) : r (μ m a) c := _root_.trans (act_rel_act_of_rel m ab) rl +#align act_rel_of_rel_of_act_rel act_rel_of_rel_of_act_rel theorem rel_act_of_rel_of_rel_act (ab : r a b) (rr : r c (μ m a)) : r c (μ m b) := _root_.trans rr (act_rel_act_of_rel _ ab) +#align rel_act_of_rel_of_rel_act rel_act_of_rel_of_rel_act end Trans @@ -199,6 +213,7 @@ variable {M N μ r} {mu : N → N → N} [IsTrans N r] [i : CovariantClass N N m theorem act_rel_act_of_rel_of_rel (ab : r a b) (cd : r c d) : r (mu a c) (mu b d) := _root_.trans (@act_rel_act_of_rel _ _ (swap mu) r _ c _ _ ab) (act_rel_act_of_rel b cd) +#align act_rel_act_of_rel_of_rel act_rel_act_of_rel_of_rel end MEqN @@ -208,6 +223,7 @@ variable {M N μ r} [ContravariantClass M N μ r] theorem rel_of_act_rel_act (m : M) {a b : N} (ab : r (μ m a) (μ m b)) : r a b := ContravariantClass.elim _ ab +#align rel_of_act_rel_act rel_of_act_rel_act section Trans @@ -217,10 +233,12 @@ variable [IsTrans N r] (m n : M) {a b c d : N} theorem act_rel_of_act_rel_of_rel_act_rel (ab : r (μ m a) b) (rl : r (μ m b) (μ m c)) : r (μ m a) c := _root_.trans ab (rel_of_act_rel_act m rl) +#align act_rel_of_act_rel_of_rel_act_rel act_rel_of_act_rel_of_rel_act_rel theorem rel_act_of_act_rel_act_of_rel_act (ab : r (μ m a) (μ m b)) (rr : r b (μ m c)) : r a (μ m c) := _root_.trans (rel_of_act_rel_act m ab) rr +#align rel_act_of_act_rel_act_of_rel_act rel_act_of_act_rel_act_of_rel_act end Trans @@ -235,28 +253,33 @@ variable {f : N → α} /-- The partial application of a constant to a covariant operator is monotone. -/ theorem Covariant.monotone_of_const [CovariantClass M N μ (· ≤ ·)] (m : M) : Monotone (μ m) := fun _ _ ha ↦ CovariantClass.elim m ha +#align covariant.monotone_of_const Covariant.monotone_of_const /-- A monotone function remains monotone when composed with the partial application of a covariant operator. E.g., `∀ (m : ℕ), Monotone f → Monotone (λ n, f (m + n))`. -/ theorem Monotone.covariant_of_const [CovariantClass M N μ (· ≤ ·)] (hf : Monotone f) (m : M) : Monotone fun n ↦ f (μ m n) := fun _ _ x ↦ hf (Covariant.monotone_of_const m x) +#align monotone.covariant_of_const Monotone.covariant_of_const /-- Same as `Monotone.covariant_of_const`, but with the constant on the other side of the operator. E.g., `∀ (m : ℕ), monotone f → monotone (λ n, f (n + m))`. -/ theorem Monotone.covariant_of_const' {μ : N → N → N} [CovariantClass N N (swap μ) (· ≤ ·)] (hf : Monotone f) (m : N) : Monotone fun n ↦ f (μ n m) := fun _ _ x ↦ hf (@Covariant.monotone_of_const _ _ (swap μ) _ _ m _ _ x) +#align monotone.covariant_of_const' Monotone.covariant_of_const' /-- Dual of `Monotone.covariant_of_const` -/ theorem Antitone.covariant_of_const [CovariantClass M N μ (· ≤ ·)] (hf : Antitone f) (m : M) : Antitone fun n ↦ f (μ m n) := hf.comp_monotone <| Covariant.monotone_of_const m +#align antitone.covariant_of_const Antitone.covariant_of_const /-- Dual of `Monotone.covariant_of_const'` -/ theorem Antitone.covariant_of_const' {μ : N → N → N} [CovariantClass N N (swap μ) (· ≤ ·)] (hf : Antitone f) (m : N) : Antitone fun n ↦ f (μ n m) := hf.comp_monotone <| @Covariant.monotone_of_const _ _ (swap μ) _ _ m +#align antitone.covariant_of_const' Antitone.covariant_of_const' end Monotone @@ -266,21 +289,25 @@ theorem covariant_le_of_covariant_lt [PartialOrder N] : rcases le_iff_eq_or_lt.mp bc with (rfl | bc) · exact rfl.le · exact (h _ bc).le +#align covariant_le_of_covariant_lt covariant_le_of_covariant_lt theorem contravariant_lt_of_contravariant_le [PartialOrder N] : Contravariant M N μ (· ≤ ·) → Contravariant M N μ (· < ·) := by refine fun h a b c bc ↦ lt_iff_le_and_ne.mpr ⟨h a bc.le, ?_⟩ rintro rfl; exact lt_irrefl _ bc +#align contravariant_lt_of_contravariant_le contravariant_lt_of_contravariant_le theorem covariant_le_iff_contravariant_lt [LinearOrder N] : Covariant M N μ (· ≤ ·) ↔ Contravariant M N μ (· < ·) := ⟨fun h _ _ _ bc ↦ not_le.mp fun k ↦ not_le.mpr bc (h _ k), fun h _ _ _ bc ↦ not_lt.mp fun k ↦ not_lt.mpr bc (h _ k)⟩ +#align covariant_le_iff_contravariant_lt covariant_le_iff_contravariant_lt theorem covariant_lt_iff_contravariant_le [LinearOrder N] : Covariant M N μ (· < ·) ↔ Contravariant M N μ (· ≤ ·) := ⟨fun h _ _ _ bc ↦ not_lt.mp fun k ↦ not_lt.mpr bc (h _ k), fun h _ _ _ bc ↦ not_le.mp fun k ↦ not_le.mpr bc (h _ k)⟩ +#align covariant_lt_iff_contravariant_le covariant_lt_iff_contravariant_le -- Porting note: `covariant_flip_mul_iff` used to use the `IsSymmOp` typeclass from Lean 3 core. -- To avoid it, we prove the relevant lemma here. @@ -291,10 +318,14 @@ lemma flip_mul [CommSemigroup N] : (flip (· * ·) : N → N → N) = (· * ·) @[to_additive] theorem covariant_flip_mul_iff [CommSemigroup N] : Covariant N N (flip (· * ·)) r ↔ Covariant N N (· * ·) r := by rw [flip_mul] +#align covariant_flip_mul_iff covariant_flip_mul_iff +#align covariant_flip_add_iff covariant_flip_add_iff @[to_additive] theorem contravariant_flip_mul_iff [CommSemigroup N] : Contravariant N N (flip (· * ·)) r ↔ Contravariant N N (· * ·) r := by rw [flip_mul] +#align contravariant_flip_mul_iff contravariant_flip_mul_iff +#align contravariant_flip_add_iff contravariant_flip_add_iff @[to_additive] instance contravariant_mul_lt_of_covariant_mul_le [Mul N] [LinearOrder N] diff --git a/Mathlib/Algebra/Divisibility/Basic.lean b/Mathlib/Algebra/Divisibility/Basic.lean index 5dd6246a2f0cb..b640aba98934e 100644 --- a/Mathlib/Algebra/Divisibility/Basic.lean +++ b/Mathlib/Algebra/Divisibility/Basic.lean @@ -53,6 +53,7 @@ theorem Dvd.intro (c : α) (h : a * c = b) : a ∣ b := #align dvd.intro Dvd.intro alias Dvd.intro ← dvd_of_mul_right_eq +#align dvd_of_mul_right_eq dvd_of_mul_right_eq theorem exists_eq_mul_right_of_dvd (h : a ∣ b) : ∃ c, b = a * c := h @@ -140,6 +141,7 @@ theorem Dvd.intro_left (c : α) (h : c * a = b) : a ∣ b := #align dvd.intro_left Dvd.intro_left alias Dvd.intro_left ← dvd_of_mul_left_eq +#align dvd_of_mul_left_eq dvd_of_mul_left_eq theorem exists_eq_mul_left_of_dvd (h : a ∣ b) : ∃ c, b = c * a := Dvd.elim h fun c => fun H1 : b = a * c => Exists.intro c (Eq.trans H1 (mul_comm a c)) diff --git a/Mathlib/Algebra/Divisibility/Units.lean b/Mathlib/Algebra/Divisibility/Units.lean index d60d124d736d8..c932d5e92a3c8 100644 --- a/Mathlib/Algebra/Divisibility/Units.lean +++ b/Mathlib/Algebra/Divisibility/Units.lean @@ -29,12 +29,14 @@ variable [Monoid α] {a b : α} {u : αˣ} divide any element of the monoid. -/ theorem coe_dvd : ↑u ∣ a := ⟨↑u⁻¹ * a, by simp⟩ +#align units.coe_dvd Units.coe_dvd /-- In a monoid, an element `a` divides an element `b` iff `a` divides all associates of `b`. -/ theorem dvd_mul_right : a ∣ b * u ↔ a ∣ b := Iff.intro (fun ⟨c, Eq⟩ ↦ ⟨c * ↑u⁻¹, by rw [← mul_assoc, ← Eq, Units.mul_inv_cancel_right]⟩) fun ⟨c, Eq⟩ ↦ Eq.symm ▸ (_root_.dvd_mul_right _ _).mul_right _ +#align units.dvd_mul_right Units.dvd_mul_right /-- In a monoid, an element `a` divides an element `b` iff all associates of `a` divide `b`. -/ theorem mul_right_dvd : a * u ∣ b ↔ a ∣ b := @@ -77,17 +79,20 @@ variable [Monoid α] {a b u : α} (hu : IsUnit u) theorem dvd : u ∣ a := by rcases hu with ⟨u, rfl⟩ apply Units.coe_dvd +#align is_unit.dvd IsUnit.dvd @[simp] theorem dvd_mul_right : a ∣ b * u ↔ a ∣ b := by rcases hu with ⟨u, rfl⟩ apply Units.dvd_mul_right +#align is_unit.dvd_mul_right IsUnit.dvd_mul_right /-- In a monoid, an element a divides an element b iff all associates of `a` divide `b`.-/ @[simp] theorem mul_right_dvd : a * u ∣ b ↔ a ∣ b := by rcases hu with ⟨u, rfl⟩ apply Units.mul_right_dvd +#align is_unit.mul_right_dvd IsUnit.mul_right_dvd end Monoid @@ -101,6 +106,7 @@ variable [CommMonoid α] (a b u : α) (hu : IsUnit u) theorem dvd_mul_left : a ∣ u * b ↔ a ∣ b := by rcases hu with ⟨u, rfl⟩ apply Units.dvd_mul_left +#align is_unit.dvd_mul_left IsUnit.dvd_mul_left /-- In a commutative monoid, an element `a` divides an element `b` iff all left associates of `a` divide `b`.-/ @@ -108,6 +114,7 @@ theorem dvd_mul_left : a ∣ u * b ↔ a ∣ b := by theorem mul_left_dvd : u * a ∣ b ↔ a ∣ b := by rcases hu with ⟨u, rfl⟩ apply Units.mul_left_dvd +#align is_unit.mul_left_dvd IsUnit.mul_left_dvd end CommMonoid diff --git a/Mathlib/Algebra/Field/Basic.lean b/Mathlib/Algebra/Field/Basic.lean index 098e3b24e6836..0287961cb7a7d 100644 --- a/Mathlib/Algebra/Field/Basic.lean +++ b/Mathlib/Algebra/Field/Basic.lean @@ -30,36 +30,46 @@ section DivisionSemiring variable [DivisionSemiring α] {a b c : α} theorem add_div (a b c : α) : (a + b) / c = a / c + b / c := by simp_rw [div_eq_mul_inv, add_mul] +#align add_div add_div @[field_simps] theorem div_add_div_same (a b c : α) : a / c + b / c = (a + b) / c := (add_div _ _ _).symm +#align div_add_div_same div_add_div_same theorem same_add_div (h : b ≠ 0) : (b + a) / b = 1 + a / b := by rw [← div_self h, add_div] +#align same_add_div same_add_div theorem div_add_same (h : b ≠ 0) : (a + b) / b = a / b + 1 := by rw [← div_self h, add_div] +#align div_add_same div_add_same theorem one_add_div (h : b ≠ 0) : 1 + a / b = (b + a) / b := (same_add_div h).symm +#align one_add_div one_add_div theorem div_add_one (h : b ≠ 0) : a / b + 1 = (a + b) / b := (div_add_same h).symm +#align div_add_one div_add_one theorem one_div_mul_add_mul_one_div_eq_one_div_add_one_div (ha : a ≠ 0) (hb : b ≠ 0) : 1 / a * (a + b) * (1 / b) = 1 / a + 1 / b := by rw [mul_add, one_div_mul_cancel ha, add_mul, one_mul, mul_assoc, mul_one_div_cancel hb, mul_one, add_comm] +#align one_div_mul_add_mul_one_div_eq_one_div_add_one_div one_div_mul_add_mul_one_div_eq_one_div_add_one_div theorem add_div_eq_mul_add_div (a b : α) (hc : c ≠ 0) : a + b / c = (a * c + b) / c := (eq_div_iff_mul_eq hc).2 <| by rw [right_distrib, div_mul_cancel _ hc] +#align add_div_eq_mul_add_div add_div_eq_mul_add_div @[field_simps] theorem add_div' (a b c : α) (hc : c ≠ 0) : b + a / c = (b * c + a) / c := by rw [add_div, mul_div_cancel _ hc] +#align add_div' add_div' @[field_simps] theorem div_add' (a b c : α) (hc : c ≠ 0) : a / c + b = (a + b * c) / c := by rwa [add_comm, add_div', add_comm] +#align div_add' div_add' end DivisionSemiring @@ -70,6 +80,7 @@ variable [DivisionMonoid K] [HasDistribNeg K] {a b : K} theorem one_div_neg_one_eq_neg_one : (1 : K) / -1 = -1 := have : -1 * -1 = (1 : K) := by rw [neg_mul_neg, one_mul] Eq.symm (eq_one_div_of_mul_eq_one_right this) +#align one_div_neg_one_eq_neg_one one_div_neg_one_eq_neg_one theorem one_div_neg_eq_neg_one_div (a : K) : 1 / -a = -(1 / a) := calc @@ -77,6 +88,7 @@ theorem one_div_neg_eq_neg_one_div (a : K) : 1 / -a = -(1 / a) := _ = 1 / a * (1 / -1) := by rw [one_div_mul_one_div_rev] _ = 1 / a * -1 := by rw [one_div_neg_one_eq_neg_one] _ = -(1 / a) := by rw [mul_neg, mul_one] +#align one_div_neg_eq_neg_one_div one_div_neg_eq_neg_one_div theorem div_neg_eq_neg_div (a b : K) : b / -a = -(b / a) := calc @@ -84,20 +96,27 @@ theorem div_neg_eq_neg_div (a b : K) : b / -a = -(b / a) := _ = b * -(1 / a) := by rw [one_div_neg_eq_neg_one_div] _ = -(b * (1 / a)) := by rw [neg_mul_eq_mul_neg] _ = -(b / a) := by rw [mul_one_div] +#align div_neg_eq_neg_div div_neg_eq_neg_div theorem neg_div (a b : K) : -b / a = -(b / a) := by rw [neg_eq_neg_one_mul, mul_div_assoc, ← neg_eq_neg_one_mul] +#align neg_div neg_div @[field_simps] theorem neg_div' (a b : K) : -(b / a) = -b / a := by simp [neg_div] +#align neg_div' neg_div' theorem neg_div_neg_eq (a b : K) : -a / -b = a / b := by rw [div_neg_eq_neg_div, neg_div, neg_neg] +#align neg_div_neg_eq neg_div_neg_eq theorem neg_inv : -a⁻¹ = (-a)⁻¹ := by rw [inv_eq_one_div, inv_eq_one_div, div_neg_eq_neg_div] +#align neg_inv neg_inv theorem div_neg (a : K) : a / -b = -(a / b) := by rw [← div_neg_eq_neg_div] +#align div_neg div_neg theorem inv_neg : (-a)⁻¹ = -a⁻¹ := by rw [neg_inv] +#align inv_neg inv_neg theorem inv_neg_one : (-1 : K)⁻¹ = -1 := by rw [← neg_inv, inv_one] @@ -109,36 +128,46 @@ variable [DivisionRing K] {a b : K} @[simp] theorem div_neg_self {a : K} (h : a ≠ 0) : a / -a = -1 := by rw [div_neg_eq_neg_div, div_self h] +#align div_neg_self div_neg_self @[simp] theorem neg_div_self {a : K} (h : a ≠ 0) : -a / a = -1 := by rw [neg_div, div_self h] +#align neg_div_self neg_div_self theorem div_sub_div_same (a b c : K) : a / c - b / c = (a - b) / c := by rw [sub_eq_add_neg, ← neg_div, div_add_div_same, sub_eq_add_neg] +#align div_sub_div_same div_sub_div_same theorem same_sub_div {a b : K} (h : b ≠ 0) : (b - a) / b = 1 - a / b := by simpa only [← @div_self _ _ b h] using (div_sub_div_same b a b).symm +#align same_sub_div same_sub_div theorem one_sub_div {a b : K} (h : b ≠ 0) : 1 - a / b = (b - a) / b := (same_sub_div h).symm +#align one_sub_div one_sub_div theorem div_sub_same {a b : K} (h : b ≠ 0) : (a - b) / b = a / b - 1 := by simpa only [← @div_self _ _ b h] using (div_sub_div_same a b b).symm +#align div_sub_same div_sub_same theorem div_sub_one {a b : K} (h : b ≠ 0) : a / b - 1 = (a - b) / b := (div_sub_same h).symm +#align div_sub_one div_sub_one theorem sub_div (a b c : K) : (a - b) / c = a / c - b / c := (div_sub_div_same _ _ _).symm +#align sub_div sub_div /-- See `inv_sub_inv` for the more convenient version when `K` is commutative. -/ theorem inv_sub_inv' {a b : K} (ha : a ≠ 0) (hb : b ≠ 0) : a⁻¹ - b⁻¹ = a⁻¹ * (b - a) * b⁻¹ := by rw [mul_sub, sub_mul, mul_inv_cancel_right₀ hb, inv_mul_cancel ha, one_mul] +#align inv_sub_inv' inv_sub_inv' theorem one_div_mul_sub_mul_one_div_eq_one_div_add_one_div (ha : a ≠ 0) (hb : b ≠ 0) : 1 / a * (b - a) * (1 / b) = 1 / a - 1 / b := by rw [mul_sub_left_distrib (1 / a), one_div_mul_cancel ha, mul_sub_right_distrib, one_mul, mul_assoc, mul_one_div_cancel hb, mul_one] +#align one_div_mul_sub_mul_one_div_eq_one_div_add_one_div one_div_mul_sub_mul_one_div_eq_one_div_add_one_div -- see Note [lower instance priority] instance (priority := 100) DivisionRing.isDomain : IsDomain K := @@ -154,12 +183,15 @@ variable [Semifield α] {a b c d : α} theorem div_add_div (a : α) (c : α) (hb : b ≠ 0) (hd : d ≠ 0) : a / b + c / d = (a * d + b * c) / (b * d) := by rw [← mul_div_mul_right _ b hd, ← mul_div_mul_left c d hb, div_add_div_same] +#align div_add_div div_add_div theorem one_div_add_one_div (ha : a ≠ 0) (hb : b ≠ 0) : 1 / a + 1 / b = (a + b) / (a * b) := by rw [div_add_div _ _ ha hb, one_mul, mul_one, add_comm] +#align one_div_add_one_div one_div_add_one_div theorem inv_add_inv (ha : a ≠ 0) (hb : b ≠ 0) : a⁻¹ + b⁻¹ = (a + b) / (a * b) := by rw [inv_eq_one_div, inv_eq_one_div, one_div_add_one_div ha hb] +#align inv_add_inv inv_add_inv end Semifield @@ -175,17 +207,21 @@ theorem div_sub_div (a : K) {b : K} (c : K) {d : K} (hb : b ≠ 0) (hd : d ≠ 0 simp [sub_eq_add_neg] rw [neg_eq_neg_one_mul, ← mul_div_assoc, div_add_div _ _ hb hd, ← mul_assoc, mul_comm b, mul_assoc, ← neg_eq_neg_one_mul] +#align div_sub_div div_sub_div theorem inv_sub_inv {a b : K} (ha : a ≠ 0) (hb : b ≠ 0) : a⁻¹ - b⁻¹ = (b - a) / (a * b) := by rw [inv_eq_one_div, inv_eq_one_div, div_sub_div _ _ ha hb, one_mul, mul_one] +#align inv_sub_inv inv_sub_inv @[field_simps] theorem sub_div' (a b c : K) (hc : c ≠ 0) : b - a / c = (b * c - a) / c := by simpa using div_sub_div b a one_ne_zero hc +#align sub_div' sub_div' @[field_simps] theorem div_sub' (a b c : K) (hc : c ≠ 0) : a / c - b = (a - c * b) / c := by simpa using div_sub_div a b hc one_ne_zero +#align div_sub' div_sub' -- see Note [lower instance priority] instance (priority := 100) Field.isDomain : IsDomain K := @@ -321,10 +357,12 @@ instance [h : Field α] : Field αᵒᵈ := @[simp] theorem toDual_rat_cast [RatCast α] (n : ℚ) : toDual (n : α) = n := rfl +#align to_dual_rat_cast toDual_rat_cast @[simp] theorem ofDual_rat_cast [RatCast α] (n : ℚ) : (ofDual n : α) = n := rfl +#align of_dual_rat_cast ofDual_rat_cast /-! ### Lexicographic order -/ @@ -346,7 +384,9 @@ instance [h : Field α] : Field (Lex α) := @[simp] theorem toLex_rat_cast [RatCast α] (n : ℚ) : toLex (n : α) = n := rfl +#align to_lex_rat_cast toLex_rat_cast @[simp] theorem ofLex_rat_cast [RatCast α] (n : ℚ) : (ofLex n : α) = n := rfl +#align of_lex_rat_cast ofLex_rat_cast diff --git a/Mathlib/Algebra/FreeMonoid/Basic.lean b/Mathlib/Algebra/FreeMonoid/Basic.lean index e052c601e3a92..863c49b7513b9 100644 --- a/Mathlib/Algebra/FreeMonoid/Basic.lean +++ b/Mathlib/Algebra/FreeMonoid/Basic.lean @@ -213,11 +213,15 @@ The purpose is to make `FreeAddMonoid.lift_eval_of` true by `rfl`."] def prodAux {M} [Monoid M] : List M → M | [] => 1 | (x :: xs) => List.foldl (· * ·) x xs +#align free_monoid.prod_aux FreeMonoid.prodAux +#align free_add_monoid.sum_aux FreeAddMonoid.sumAux @[to_additive] lemma prodAux_eq : ∀ l : List M, FreeMonoid.prodAux l = l.prod | [] => rfl | (_ :: xs) => congr_arg (fun x => List.foldl (· * ·) x xs) (one_mul _).symm +#align free_monoid.prod_aux_eq FreeMonoid.prodAux_eq +#align free_add_monoid.sum_aux_eq FreeAddMonoid.sumAux_eq /-- Equivalence between maps `α → M` and monoid homomorphisms `FreeMonoid α →* M`. -/ @[to_additive "Equivalence between maps `α → A` and additive monoid homomorphisms diff --git a/Mathlib/Algebra/GCDMonoid/Basic.lean b/Mathlib/Algebra/GCDMonoid/Basic.lean index ecbfe378f8213..804b7cc01fb5c 100644 --- a/Mathlib/Algebra/GCDMonoid/Basic.lean +++ b/Mathlib/Algebra/GCDMonoid/Basic.lean @@ -970,6 +970,8 @@ def associatesEquivOfUniqueUnits : Associates α ≃* α where right_inv _ := (Associates.out_mk _).trans <| normalize_eq _ map_mul' := Associates.out_mul #align associates_equiv_of_unique_units associatesEquivOfUniqueUnits +#align associates_equiv_of_unique_units_symm_apply associatesEquivOfUniqueUnits_symmApply +#align associates_equiv_of_unique_units_apply associatesEquivOfUniqueUnits_apply end UniqueUnit diff --git a/Mathlib/Algebra/GradedMonoid.lean b/Mathlib/Algebra/GradedMonoid.lean index b2a4c95e184ba..91c7037854e03 100644 --- a/Mathlib/Algebra/GradedMonoid.lean +++ b/Mathlib/Algebra/GradedMonoid.lean @@ -458,6 +458,7 @@ instance Monoid.gMonoid [AddMonoid ι] [Monoid R] : GradedMonoid.GMonoid fun _ : gnpow_zero' := fun _ => Sigma.ext (zero_nsmul _) (heq_of_eq (Monoid.npow_zero _)) gnpow_succ' := fun _ ⟨_, _⟩ => Sigma.ext (succ_nsmul _ _) (heq_of_eq (Monoid.npow_succ _ _)) } #align monoid.gmonoid Monoid.gMonoid +#align monoid.gmonoid_gnpow Monoid.gMonoid_gnpow /-- If all grades are the same type and themselves form a commutative monoid, then there is a trivial grading structure. -/ diff --git a/Mathlib/Algebra/Group/Basic.lean b/Mathlib/Algebra/Group/Basic.lean index 1fc22d40d06f4..b1aa6b5ca33c1 100644 --- a/Mathlib/Algebra/Group/Basic.lean +++ b/Mathlib/Algebra/Group/Basic.lean @@ -35,6 +35,8 @@ is equal to a addition on the left by `x + y`."] theorem comp_mul_left [Semigroup α] (x y : α) : (x * ·) ∘ (y * ·) = (x * y * ·) := by ext z simp [mul_assoc] +#align comp_mul_left comp_mul_left +#align comp_add_left comp_add_left /-- Composing two multiplications on the right by `y` and `x` is equal to a multiplication on the right by `y * x`. @@ -44,6 +46,8 @@ is equal to a addition on the right by `y + x`."] theorem comp_mul_right [Semigroup α] (x y : α) : (· * x) ∘ (· * y) = (· * (y * x)) := by ext z simp [mul_assoc] +#align comp_mul_right comp_mul_right +#align comp_add_right comp_add_right end Semigroup @@ -55,23 +59,33 @@ variable {M : Type u} [MulOneClass M] theorem ite_mul_one {P : Prop} [Decidable P] {a b : M} : ite P (a * b) 1 = ite P a 1 * ite P b 1 := by by_cases h:P <;> simp [h] +#align ite_mul_one ite_mul_one +#align ite_add_zero ite_add_zero @[to_additive] theorem ite_one_mul {P : Prop} [Decidable P] {a b : M} : ite P 1 (a * b) = ite P 1 a * ite P 1 b := by by_cases h:P <;> simp [h] +#align ite_one_mul ite_one_mul +#align ite_zero_add ite_zero_add @[to_additive] theorem eq_one_iff_eq_one_of_mul_eq_one {a b : M} (h : a * b = 1) : a = 1 ↔ b = 1 := by constructor <;> (rintro rfl; simpa using h) +#align eq_one_iff_eq_one_of_mul_eq_one eq_one_iff_eq_one_of_mul_eq_one +#align eq_zero_iff_eq_zero_of_add_eq_zero eq_zero_iff_eq_zero_of_add_eq_zero @[to_additive] theorem one_mul_eq_id : (· * ·) (1 : M) = id := funext one_mul +#align one_mul_eq_id one_mul_eq_id +#align zero_add_eq_id zero_add_eq_id @[to_additive] theorem mul_one_eq_id : (· * (1 : M)) = id := funext mul_one +#align mul_one_eq_id mul_one_eq_id +#align add_zero_eq_id add_zero_eq_id end MulOneClass @@ -82,22 +96,32 @@ variable [CommSemigroup G] @[to_additive] theorem mul_left_comm : ∀ a b c : G, a * (b * c) = b * (a * c) := left_comm Mul.mul mul_comm mul_assoc +#align mul_left_comm mul_left_comm +#align add_left_comm add_left_comm @[to_additive] theorem mul_right_comm : ∀ a b c : G, a * b * c = a * c * b := right_comm Mul.mul mul_comm mul_assoc +#align mul_right_comm mul_right_comm +#align add_right_comm add_right_comm @[to_additive] theorem mul_mul_mul_comm (a b c d : G) : a * b * (c * d) = a * c * (b * d) := by simp only [mul_left_comm, mul_assoc] +#align mul_mul_mul_comm mul_mul_mul_comm +#align add_add_add_comm add_add_add_comm @[to_additive] theorem mul_rotate (a b c : G) : a * b * c = b * c * a := by simp only [mul_left_comm, mul_comm] +#align mul_rotate mul_rotate +#align add_rotate add_rotate @[to_additive] theorem mul_rotate' (a b c : G) : a * (b * c) = b * (c * a) := by simp only [mul_left_comm, mul_comm] +#align mul_rotate' mul_rotate' +#align add_rotate' add_rotate' end CommSemigroup @@ -145,6 +169,8 @@ variable {M : Type u} [CommMonoid M] {x y z : M} @[to_additive] theorem inv_unique (hy : x * y = 1) (hz : x * z = 1) : y = z := left_inv_eq_right_inv (Trans.trans (mul_comm _ _) hy) hz +#align inv_unique inv_unique +#align neg_unique neg_unique end CommMonoid @@ -156,10 +182,14 @@ variable {M : Type u} [LeftCancelMonoid M] {a b : M} theorem mul_right_eq_self : a * b = a ↔ b = 1 := calc a * b = a ↔ a * b = a * 1 := by rw [mul_one] _ ↔ b = 1 := mul_left_cancel_iff +#align mul_right_eq_self mul_right_eq_self +#align add_right_eq_self add_right_eq_self @[to_additive (attr := simp)] theorem self_eq_mul_right : a = a * b ↔ b = 1 := eq_comm.trans mul_right_eq_self +#align self_eq_mul_right self_eq_mul_right +#align self_eq_add_right self_eq_add_right end LeftCancelMonoid @@ -171,10 +201,14 @@ variable {M : Type u} [RightCancelMonoid M] {a b : M} theorem mul_left_eq_self : a * b = b ↔ a = 1 := calc a * b = b ↔ a * b = 1 * b := by rw [one_mul] _ ↔ a = 1 := mul_right_cancel_iff +#align mul_left_eq_self mul_left_eq_self +#align add_left_eq_self add_left_eq_self @[to_additive (attr := simp)] theorem self_eq_mul_left : b = a * b ↔ a = 1 := eq_comm.trans mul_left_eq_self +#align self_eq_mul_left self_eq_mul_left +#align self_eq_add_left self_eq_add_left end RightCancelMonoid @@ -185,35 +219,51 @@ variable [InvolutiveInv G] {a b : G} @[to_additive (attr := simp)] theorem inv_involutive : Function.Involutive (Inv.inv : G → G) := inv_inv +#align inv_involutive inv_involutive +#align neg_involutive neg_involutive @[to_additive (attr := simp)] theorem inv_surjective : Function.Surjective (Inv.inv : G → G) := inv_involutive.surjective +#align inv_surjective inv_surjective +#align neg_surjective neg_surjective @[to_additive] theorem inv_injective : Function.Injective (Inv.inv : G → G) := inv_involutive.injective +#align inv_injective inv_injective +#align neg_injective neg_injective @[to_additive (attr := simp)] theorem inv_inj {a b : G} : a⁻¹ = b⁻¹ ↔ a = b := inv_injective.eq_iff +#align inv_inj inv_inj +#align neg_inj neg_inj @[to_additive] theorem eq_inv_of_eq_inv (h : a = b⁻¹) : b = a⁻¹ := by simp [h] +#align eq_inv_of_eq_inv eq_inv_of_eq_inv +#align eq_neg_of_eq_neg eq_neg_of_eq_neg @[to_additive] theorem eq_inv_iff_eq_inv : a = b⁻¹ ↔ b = a⁻¹ := ⟨eq_inv_of_eq_inv, eq_inv_of_eq_inv⟩ +#align eq_inv_iff_eq_inv eq_inv_iff_eq_inv +#align eq_neg_iff_eq_neg eq_neg_iff_eq_neg @[to_additive] theorem inv_eq_iff_inv_eq : a⁻¹ = b ↔ b⁻¹ = a := eq_comm.trans <| eq_inv_iff_eq_inv.trans eq_comm +#align inv_eq_iff_inv_eq inv_eq_iff_inv_eq +#align neg_eq_iff_neg_eq neg_eq_iff_neg_eq variable (G) @[to_additive] theorem inv_comp_inv : Inv.inv ∘ Inv.inv = @id G := inv_involutive.comp_self +#align inv_comp_inv inv_comp_inv +#align neg_comp_neg neg_comp_neg @[to_additive] theorem leftInverse_inv : LeftInverse (fun a : G ↦ a⁻¹) fun a ↦ a⁻¹ := @@ -235,28 +285,42 @@ variable [DivInvMonoid G] {a b c : G} @[to_additive, field_simps] -- The attributes are out of order on purpose theorem inv_eq_one_div (x : G) : x⁻¹ = 1 / x := by rw [div_eq_mul_inv, one_mul] +#align inv_eq_one_div inv_eq_one_div +#align neg_eq_zero_sub neg_eq_zero_sub @[to_additive] theorem mul_one_div (x y : G) : x * (1 / y) = x / y := by rw [div_eq_mul_inv, one_mul, div_eq_mul_inv] +#align mul_one_div mul_one_div +#align add_zero_sub add_zero_sub @[to_additive] theorem mul_div_assoc (a b c : G) : a * b / c = a * (b / c) := by rw [div_eq_mul_inv, div_eq_mul_inv, mul_assoc _ _ _] +#align mul_div_assoc mul_div_assoc +#align add_sub_assoc add_sub_assoc @[to_additive, field_simps] -- The attributes are out of order on purpose theorem mul_div_assoc' (a b c : G) : a * (b / c) = a * b / c := (mul_div_assoc _ _ _).symm +#align mul_div_assoc' mul_div_assoc' +#align add_sub_assoc' add_sub_assoc' @[to_additive (attr := simp)] theorem one_div (a : G) : 1 / a = a⁻¹ := (inv_eq_one_div a).symm +#align one_div one_div +#align zero_sub zero_sub @[to_additive] theorem mul_div (a b c : G) : a * (b / c) = a * b / c := by simp only [mul_assoc, div_eq_mul_inv] +#align mul_div mul_div +#align add_sub add_sub @[to_additive] theorem div_eq_mul_one_div (a b : G) : a / b = a * (1 / b) := by rw [div_eq_mul_inv, one_div] +#align div_eq_mul_one_div div_eq_mul_one_div +#align sub_eq_add_zero_sub sub_eq_add_zero_sub end DivInvMonoid @@ -266,10 +330,14 @@ variable [DivInvOneMonoid G] @[to_additive (attr := simp)] theorem div_one (a : G) : a / 1 = a := by simp [div_eq_mul_inv] +#align div_one div_one +#align sub_zero sub_zero @[to_additive] theorem one_div_one : (1 : G) / 1 = 1 := div_one _ +#align one_div_one one_div_one +#align zero_sub_zero zero_sub_zero end DivInvOneMonoid @@ -282,47 +350,71 @@ attribute [local simp] mul_assoc div_eq_mul_inv @[to_additive] theorem inv_eq_of_mul_eq_one_left (h : a * b = 1) : b⁻¹ = a := by rw [← inv_eq_of_mul_eq_one_right h, inv_inv] +#align inv_eq_of_mul_eq_one_left inv_eq_of_mul_eq_one_left +#align neg_eq_of_add_eq_zero_left neg_eq_of_add_eq_zero_left @[to_additive] theorem eq_inv_of_mul_eq_one_left (h : a * b = 1) : a = b⁻¹ := (inv_eq_of_mul_eq_one_left h).symm +#align eq_inv_of_mul_eq_one_left eq_inv_of_mul_eq_one_left +#align eq_neg_of_add_eq_zero_left eq_neg_of_add_eq_zero_left @[to_additive] theorem eq_inv_of_mul_eq_one_right (h : a * b = 1) : b = a⁻¹ := (inv_eq_of_mul_eq_one_right h).symm +#align eq_inv_of_mul_eq_one_right eq_inv_of_mul_eq_one_right +#align eq_neg_of_add_eq_zero_right eq_neg_of_add_eq_zero_right @[to_additive] theorem eq_one_div_of_mul_eq_one_left (h : b * a = 1) : b = 1 / a := by rw [eq_inv_of_mul_eq_one_left h, one_div] +#align eq_one_div_of_mul_eq_one_left eq_one_div_of_mul_eq_one_left +#align eq_zero_sub_of_add_eq_zero_left eq_zero_sub_of_add_eq_zero_left @[to_additive] theorem eq_one_div_of_mul_eq_one_right (h : a * b = 1) : b = 1 / a := by rw [eq_inv_of_mul_eq_one_right h, one_div] +#align eq_one_div_of_mul_eq_one_right eq_one_div_of_mul_eq_one_right +#align eq_zero_sub_of_add_eq_zero_right eq_zero_sub_of_add_eq_zero_right @[to_additive] theorem eq_of_div_eq_one (h : a / b = 1) : a = b := inv_injective <| inv_eq_of_mul_eq_one_right <| by rwa [← div_eq_mul_inv] +#align eq_of_div_eq_one eq_of_div_eq_one +#align eq_of_sub_eq_zero eq_of_sub_eq_zero @[to_additive] theorem div_ne_one_of_ne : a ≠ b → a / b ≠ 1 := mt eq_of_div_eq_one +#align div_ne_one_of_ne div_ne_one_of_ne +#align sub_ne_zero_of_ne sub_ne_zero_of_ne variable (a b c) @[to_additive] theorem one_div_mul_one_div_rev : 1 / a * (1 / b) = 1 / (b * a) := by simp +#align one_div_mul_one_div_rev one_div_mul_one_div_rev +#align zero_sub_add_zero_sub_rev zero_sub_add_zero_sub_rev @[to_additive] theorem inv_div_left : a⁻¹ / b = (b * a)⁻¹ := by simp +#align inv_div_left inv_div_left +#align neg_sub_left neg_sub_left @[to_additive (attr := simp)] theorem inv_div : (a / b)⁻¹ = b / a := by simp +#align inv_div inv_div +#align neg_sub neg_sub @[to_additive] theorem one_div_div : 1 / (a / b) = b / a := by simp +#align one_div_div one_div_div +#align zero_sub_sub zero_sub_sub @[to_additive] theorem one_div_one_div : 1 / (1 / a) = a := by simp +#align one_div_one_div one_div_one_div +#align zero_sub_zero_sub zero_sub_zero_sub @[to_additive SubtractionMonoid.toSubNegZeroMonoid] instance (priority := 100) DivisionMonoid.toDivInvOneMonoid : DivInvOneMonoid α := @@ -334,30 +426,44 @@ variable {a b c} @[to_additive (attr := simp)] theorem inv_eq_one : a⁻¹ = 1 ↔ a = 1 := inv_injective.eq_iff' inv_one +#align inv_eq_one inv_eq_one +#align neg_eq_zero neg_eq_zero @[to_additive (attr := simp)] theorem one_eq_inv : 1 = a⁻¹ ↔ a = 1 := eq_comm.trans inv_eq_one +#align one_eq_inv one_eq_inv +#align zero_eq_neg zero_eq_neg @[to_additive] theorem inv_ne_one : a⁻¹ ≠ 1 ↔ a ≠ 1 := inv_eq_one.not +#align inv_ne_one inv_ne_one +#align neg_ne_zero neg_ne_zero @[to_additive] theorem eq_of_one_div_eq_one_div (h : 1 / a = 1 / b) : a = b := by rw [← one_div_one_div a, h, one_div_one_div] +#align eq_of_one_div_eq_one_div eq_of_one_div_eq_one_div +#align eq_of_zero_sub_eq_zero_sub eq_of_zero_sub_eq_zero_sub variable (a b c) @[to_additive, field_simps] -- The attributes are out of order on purpose theorem div_div_eq_mul_div : a / (b / c) = a * c / b := by simp +#align div_div_eq_mul_div div_div_eq_mul_div +#align sub_sub_eq_add_sub sub_sub_eq_add_sub @[to_additive (attr := simp)] theorem div_inv_eq_mul : a / b⁻¹ = a * b := by simp +#align div_inv_eq_mul div_inv_eq_mul +#align sub_neg_eq_add sub_neg_eq_add @[to_additive] theorem div_mul_eq_div_div_swap : a / (b * c) = a / c / b := by simp only [mul_assoc, mul_inv_rev, div_eq_mul_inv] +#align div_mul_eq_div_div_swap div_mul_eq_div_div_swap +#align sub_add_eq_sub_sub_swap sub_add_eq_sub_sub_swap end DivisionMonoid @@ -366,6 +472,7 @@ section SubtractionMonoid set_option linter.deprecated false lemma bit0_neg [SubtractionMonoid α] (a : α) : bit0 (-a) = -bit0 a := (neg_add_rev _ _).symm +#align bit0_neg bit0_neg end SubtractionMonoid @@ -377,69 +484,113 @@ attribute [local simp] mul_assoc mul_comm mul_left_comm div_eq_mul_inv @[to_additive neg_add] theorem mul_inv : (a * b)⁻¹ = a⁻¹ * b⁻¹ := by simp +#align mul_inv mul_inv +#align neg_add neg_add @[to_additive] theorem inv_div' : (a / b)⁻¹ = a⁻¹ / b⁻¹ := by simp +#align inv_div' inv_div' +#align neg_sub' neg_sub' @[to_additive] theorem div_eq_inv_mul : a / b = b⁻¹ * a := by simp +#align div_eq_inv_mul div_eq_inv_mul +#align sub_eq_neg_add sub_eq_neg_add @[to_additive] theorem inv_mul_eq_div : a⁻¹ * b = b / a := by simp +#align inv_mul_eq_div inv_mul_eq_div +#align neg_add_eq_sub neg_add_eq_sub @[to_additive] theorem inv_mul' : (a * b)⁻¹ = a⁻¹ / b := by simp +#align inv_mul' inv_mul' +#align neg_add' neg_add' @[to_additive] theorem inv_div_inv : a⁻¹ / b⁻¹ = b / a := by simp +#align inv_div_inv inv_div_inv +#align neg_sub_neg neg_sub_neg @[to_additive] theorem inv_inv_div_inv : (a⁻¹ / b⁻¹)⁻¹ = a / b := by simp +#align inv_inv_div_inv inv_inv_div_inv +#align neg_neg_sub_neg neg_neg_sub_neg @[to_additive] theorem one_div_mul_one_div : 1 / a * (1 / b) = 1 / (a * b) := by simp +#align one_div_mul_one_div one_div_mul_one_div +#align zero_sub_add_zero_sub zero_sub_add_zero_sub @[to_additive] theorem div_right_comm : a / b / c = a / c / b := by simp +#align div_right_comm div_right_comm +#align sub_right_comm sub_right_comm @[to_additive, field_simps] theorem div_div : a / b / c = a / (b * c) := by simp +#align div_div div_div +#align sub_sub sub_sub @[to_additive] theorem div_mul : a / b * c = a / (b / c) := by simp +#align div_mul div_mul +#align sub_add sub_add @[to_additive] theorem mul_div_left_comm : a * (b / c) = b * (a / c) := by simp +#align mul_div_left_comm mul_div_left_comm +#align add_sub_left_comm add_sub_left_comm @[to_additive] theorem mul_div_right_comm : a * b / c = a / c * b := by simp +#align mul_div_right_comm mul_div_right_comm +#align add_sub_right_comm add_sub_right_comm @[to_additive] theorem div_mul_eq_div_div : a / (b * c) = a / b / c := by simp +#align div_mul_eq_div_div div_mul_eq_div_div +#align sub_add_eq_sub_sub sub_add_eq_sub_sub @[to_additive, field_simps] theorem div_mul_eq_mul_div : a / b * c = a * c / b := by simp +#align div_mul_eq_mul_div div_mul_eq_mul_div +#align sub_add_eq_add_sub sub_add_eq_add_sub @[to_additive] theorem mul_comm_div : a / b * c = a * (c / b) := by simp +#align mul_comm_div mul_comm_div +#align add_comm_sub add_comm_sub @[to_additive] theorem div_mul_comm : a / b * c = c / b * a := by simp +#align div_mul_comm div_mul_comm +#align sub_add_comm sub_add_comm @[to_additive] theorem div_mul_eq_div_mul_one_div : a / (b * c) = a / b * (1 / c) := by simp +#align div_mul_eq_div_mul_one_div div_mul_eq_div_mul_one_div +#align sub_add_eq_sub_add_zero_sub sub_add_eq_sub_add_zero_sub @[to_additive] theorem div_div_div_eq : a / b / (c / d) = a * d / (b * c) := by simp +#align div_div_div_eq div_div_div_eq +#align sub_sub_sub_eq sub_sub_sub_eq @[to_additive] theorem div_div_div_comm : a / b / (c / d) = a / c / (b / d) := by simp +#align div_div_div_comm div_div_div_comm +#align sub_sub_sub_comm sub_sub_sub_comm @[to_additive] theorem div_mul_div_comm : a / b * (c / d) = a * c / (b * d) := by simp +#align div_mul_div_comm div_mul_div_comm +#align sub_add_sub_comm sub_add_sub_comm @[to_additive] theorem mul_div_mul_comm : a * b / (c * d) = a / c * (b / d) := by simp +#align mul_div_mul_comm mul_div_mul_comm +#align add_sub_add_comm add_sub_add_comm end DivisionCommMonoid @@ -449,157 +600,241 @@ variable [Group G] {a b c d : G} @[to_additive (attr := simp)] theorem div_eq_inv_self : a / b = b⁻¹ ↔ a = 1 := by rw [div_eq_mul_inv, mul_left_eq_self] +#align div_eq_inv_self div_eq_inv_self +#align sub_eq_neg_self sub_eq_neg_self @[to_additive] theorem mul_left_surjective (a : G) : Function.Surjective ((· * ·) a) := fun x ↦ ⟨a⁻¹ * x, mul_inv_cancel_left a x⟩ +#align mul_left_surjective mul_left_surjective +#align add_left_surjective add_left_surjective @[to_additive] theorem mul_right_surjective (a : G) : Function.Surjective fun x ↦ x * a := fun x ↦ ⟨x * a⁻¹, inv_mul_cancel_right x a⟩ +#align mul_right_surjective mul_right_surjective +#align add_right_surjective add_right_surjective @[to_additive] theorem eq_mul_inv_of_mul_eq (h : a * c = b) : a = b * c⁻¹ := by simp [h.symm] +#align eq_mul_inv_of_mul_eq eq_mul_inv_of_mul_eq +#align eq_add_neg_of_add_eq eq_add_neg_of_add_eq @[to_additive] theorem eq_inv_mul_of_mul_eq (h : b * a = c) : a = b⁻¹ * c := by simp [h.symm] +#align eq_inv_mul_of_mul_eq eq_inv_mul_of_mul_eq +#align eq_neg_add_of_add_eq eq_neg_add_of_add_eq @[to_additive] theorem inv_mul_eq_of_eq_mul (h : b = a * c) : a⁻¹ * b = c := by simp [h] +#align inv_mul_eq_of_eq_mul inv_mul_eq_of_eq_mul +#align neg_add_eq_of_eq_add neg_add_eq_of_eq_add @[to_additive] theorem mul_inv_eq_of_eq_mul (h : a = c * b) : a * b⁻¹ = c := by simp [h] +#align mul_inv_eq_of_eq_mul mul_inv_eq_of_eq_mul +#align add_neg_eq_of_eq_add add_neg_eq_of_eq_add @[to_additive] theorem eq_mul_of_mul_inv_eq (h : a * c⁻¹ = b) : a = b * c := by simp [h.symm] +#align eq_mul_of_mul_inv_eq eq_mul_of_mul_inv_eq +#align eq_add_of_add_neg_eq eq_add_of_add_neg_eq @[to_additive] theorem eq_mul_of_inv_mul_eq (h : b⁻¹ * a = c) : a = b * c := by simp [h.symm, mul_inv_cancel_left] +#align eq_mul_of_inv_mul_eq eq_mul_of_inv_mul_eq +#align eq_add_of_neg_add_eq eq_add_of_neg_add_eq @[to_additive] theorem mul_eq_of_eq_inv_mul (h : b = a⁻¹ * c) : a * b = c := by rw [h, mul_inv_cancel_left] +#align mul_eq_of_eq_inv_mul mul_eq_of_eq_inv_mul +#align add_eq_of_eq_neg_add add_eq_of_eq_neg_add @[to_additive] theorem mul_eq_of_eq_mul_inv (h : a = c * b⁻¹) : a * b = c := by simp [h] +#align mul_eq_of_eq_mul_inv mul_eq_of_eq_mul_inv +#align add_eq_of_eq_add_neg add_eq_of_eq_add_neg @[to_additive] theorem mul_eq_one_iff_eq_inv : a * b = 1 ↔ a = b⁻¹ := ⟨eq_inv_of_mul_eq_one_left, fun h ↦ by rw [h, mul_left_inv]⟩ +#align mul_eq_one_iff_eq_inv mul_eq_one_iff_eq_inv +#align add_eq_zero_iff_eq_neg add_eq_zero_iff_eq_neg @[to_additive] theorem mul_eq_one_iff_inv_eq : a * b = 1 ↔ a⁻¹ = b := by rw [mul_eq_one_iff_eq_inv, eq_inv_iff_eq_inv, eq_comm] +#align mul_eq_one_iff_inv_eq mul_eq_one_iff_inv_eq +#align add_eq_zero_iff_neg_eq add_eq_zero_iff_neg_eq @[to_additive] theorem eq_inv_iff_mul_eq_one : a = b⁻¹ ↔ a * b = 1 := mul_eq_one_iff_eq_inv.symm +#align eq_inv_iff_mul_eq_one eq_inv_iff_mul_eq_one +#align eq_neg_iff_add_eq_zero eq_neg_iff_add_eq_zero @[to_additive] theorem inv_eq_iff_mul_eq_one : a⁻¹ = b ↔ a * b = 1 := mul_eq_one_iff_inv_eq.symm +#align inv_eq_iff_mul_eq_one inv_eq_iff_mul_eq_one +#align neg_eq_iff_add_eq_zero neg_eq_iff_add_eq_zero @[to_additive] theorem eq_mul_inv_iff_mul_eq : a = b * c⁻¹ ↔ a * c = b := ⟨fun h ↦ by rw [h, inv_mul_cancel_right], fun h ↦ by rw [← h, mul_inv_cancel_right]⟩ +#align eq_mul_inv_iff_mul_eq eq_mul_inv_iff_mul_eq +#align eq_add_neg_iff_add_eq eq_add_neg_iff_add_eq @[to_additive] theorem eq_inv_mul_iff_mul_eq : a = b⁻¹ * c ↔ b * a = c := ⟨fun h ↦ by rw [h, mul_inv_cancel_left], fun h ↦ by rw [← h, inv_mul_cancel_left]⟩ +#align eq_inv_mul_iff_mul_eq eq_inv_mul_iff_mul_eq +#align eq_neg_add_iff_add_eq eq_neg_add_iff_add_eq @[to_additive] theorem inv_mul_eq_iff_eq_mul : a⁻¹ * b = c ↔ b = a * c := ⟨fun h ↦ by rw [← h, mul_inv_cancel_left], fun h ↦ by rw [h, inv_mul_cancel_left]⟩ +#align inv_mul_eq_iff_eq_mul inv_mul_eq_iff_eq_mul +#align neg_add_eq_iff_eq_add neg_add_eq_iff_eq_add @[to_additive] theorem mul_inv_eq_iff_eq_mul : a * b⁻¹ = c ↔ a = c * b := ⟨fun h ↦ by rw [← h, inv_mul_cancel_right], fun h ↦ by rw [h, mul_inv_cancel_right]⟩ +#align mul_inv_eq_iff_eq_mul mul_inv_eq_iff_eq_mul +#align add_neg_eq_iff_eq_add add_neg_eq_iff_eq_add @[to_additive] theorem mul_inv_eq_one : a * b⁻¹ = 1 ↔ a = b := by rw [mul_eq_one_iff_eq_inv, inv_inv] +#align mul_inv_eq_one mul_inv_eq_one +#align add_neg_eq_zero add_neg_eq_zero @[to_additive] theorem inv_mul_eq_one : a⁻¹ * b = 1 ↔ a = b := by rw [mul_eq_one_iff_eq_inv, inv_inj] +#align inv_mul_eq_one inv_mul_eq_one +#align neg_add_eq_zero neg_add_eq_zero @[to_additive] theorem div_left_injective : Function.Injective fun a ↦ a / b := by -- FIXME this could be by `simpa`, but it fails. This is probably a bug in `simpa`. simp only [div_eq_mul_inv] exact fun a a' h ↦ mul_left_injective b⁻¹ h +#align div_left_injective div_left_injective +#align sub_left_injective sub_left_injective @[to_additive] theorem div_right_injective : Function.Injective fun a ↦ b / a := by -- FIXME see above simp only [div_eq_mul_inv] exact fun a a' h ↦ inv_injective (mul_right_injective b h) +#align div_right_injective div_right_injective +#align sub_right_injective sub_right_injective @[to_additive (attr := simp) sub_add_cancel] theorem div_mul_cancel' (a b : G) : a / b * b = a := by rw [div_eq_mul_inv, inv_mul_cancel_right a b] +#align div_mul_cancel' div_mul_cancel' +#align sub_add_cancel sub_add_cancel @[to_additive (attr := simp) sub_self] theorem div_self' (a : G) : a / a = 1 := by rw [div_eq_mul_inv, mul_right_inv a] +#align div_self' div_self' +#align sub_self sub_self @[to_additive (attr := simp) add_sub_cancel] theorem mul_div_cancel'' (a b : G) : a * b / b = a := by rw [div_eq_mul_inv, mul_inv_cancel_right a b] +#align mul_div_cancel'' mul_div_cancel'' +#align add_sub_cancel add_sub_cancel @[to_additive (attr := simp)] theorem mul_div_mul_right_eq_div (a b c : G) : a * c / (b * c) = a / b := by rw [div_mul_eq_div_div_swap]; simp only [mul_left_inj, eq_self_iff_true, mul_div_cancel''] +#align mul_div_mul_right_eq_div mul_div_mul_right_eq_div +#align add_sub_add_right_eq_sub add_sub_add_right_eq_sub @[to_additive eq_sub_of_add_eq] theorem eq_div_of_mul_eq' (h : a * c = b) : a = b / c := by simp [← h] +#align eq_div_of_mul_eq' eq_div_of_mul_eq' +#align eq_sub_of_add_eq eq_sub_of_add_eq @[to_additive sub_eq_of_eq_add] theorem div_eq_of_eq_mul'' (h : a = c * b) : a / b = c := by simp [h] +#align div_eq_of_eq_mul'' div_eq_of_eq_mul'' +#align sub_eq_of_eq_add sub_eq_of_eq_add @[to_additive] theorem eq_mul_of_div_eq (h : a / c = b) : a = b * c := by simp [← h] +#align eq_mul_of_div_eq eq_mul_of_div_eq +#align eq_add_of_sub_eq eq_add_of_sub_eq @[to_additive] theorem mul_eq_of_eq_div (h : a = c / b) : a * b = c := by simp [h] +#align mul_eq_of_eq_div mul_eq_of_eq_div +#align add_eq_of_eq_sub add_eq_of_eq_sub @[to_additive (attr := simp)] theorem div_right_inj : a / b = a / c ↔ b = c := div_right_injective.eq_iff +#align div_right_inj div_right_inj +#align sub_right_inj sub_right_inj @[to_additive (attr := simp)] theorem div_left_inj : b / a = c / a ↔ b = c := by rw [div_eq_mul_inv, div_eq_mul_inv] exact mul_left_inj _ +#align div_left_inj div_left_inj +#align sub_left_inj sub_left_inj @[to_additive (attr := simp) sub_add_sub_cancel] theorem div_mul_div_cancel' (a b c : G) : a / b * (b / c) = a / c := by rw [← mul_div_assoc, div_mul_cancel'] +#align div_mul_div_cancel' div_mul_div_cancel' +#align sub_add_sub_cancel sub_add_sub_cancel @[to_additive (attr := simp) sub_sub_sub_cancel_right] theorem div_div_div_cancel_right' (a b c : G) : a / c / (b / c) = a / b := by rw [← inv_div c b, div_inv_eq_mul, div_mul_div_cancel'] +#align div_div_div_cancel_right' div_div_div_cancel_right' +#align sub_sub_sub_cancel_right sub_sub_sub_cancel_right @[to_additive] theorem div_eq_one : a / b = 1 ↔ a = b := ⟨eq_of_div_eq_one, fun h ↦ by rw [h, div_self']⟩ +#align div_eq_one div_eq_one +#align sub_eq_zero sub_eq_zero alias div_eq_one ↔ _ div_eq_one_of_eq +#align div_eq_one_of_eq div_eq_one_of_eq alias sub_eq_zero ↔ _ sub_eq_zero_of_eq +#align sub_eq_zero_of_eq sub_eq_zero_of_eq @[to_additive] theorem div_ne_one : a / b ≠ 1 ↔ a ≠ b := not_congr div_eq_one +#align div_ne_one div_ne_one +#align sub_ne_zero sub_ne_zero @[to_additive (attr := simp)] theorem div_eq_self : a / b = a ↔ b = 1 := by rw [div_eq_mul_inv, mul_right_eq_self, inv_eq_one] +#align div_eq_self div_eq_self +#align sub_eq_self sub_eq_self @[to_additive eq_sub_iff_add_eq] theorem eq_div_iff_mul_eq' : a = b / c ↔ a * c = b := by rw [div_eq_mul_inv, eq_mul_inv_iff_mul_eq] +#align eq_div_iff_mul_eq' eq_div_iff_mul_eq' +#align eq_sub_iff_add_eq eq_sub_iff_add_eq @[to_additive] theorem div_eq_iff_eq_mul : a / b = c ↔ a = c * b := by rw [div_eq_mul_inv, mul_inv_eq_iff_eq_mul] +#align div_eq_iff_eq_mul div_eq_iff_eq_mul +#align sub_eq_iff_eq_add sub_eq_iff_eq_add @[to_additive] theorem eq_iff_eq_of_div_eq_div (H : a / b = c / d) : a = b ↔ c = d := by rw [← div_eq_one, H, div_eq_one] +#align eq_iff_eq_of_div_eq_div eq_iff_eq_of_div_eq_div +#align eq_iff_eq_of_sub_eq_sub eq_iff_eq_of_sub_eq_sub @[to_additive] theorem leftInverse_div_mul_left (c : G) : Function.LeftInverse (fun x ↦ x / c) fun x ↦ x * c := @@ -638,6 +873,8 @@ theorem exists_npow_eq_one_of_zpow_eq_one {n : ℤ} (hn : n ≠ 0) {x : G} (h : rfl · rw [zpow_negSucc, inv_eq_one] at h refine' ⟨n + 1, n.succ_pos, h⟩ +#align exists_npow_eq_one_of_zpow_eq_one exists_npow_eq_one_of_zpow_eq_one +#align exists_nsmul_eq_zero_of_zsmul_eq_zero exists_nsmul_eq_zero_of_zsmul_eq_zero end Group @@ -650,51 +887,79 @@ attribute [local simp] mul_assoc mul_comm mul_left_comm div_eq_mul_inv @[to_additive] theorem div_eq_of_eq_mul' {a b c : G} (h : a = b * c) : a / b = c := by rw [h, div_eq_mul_inv, mul_comm, inv_mul_cancel_left] +#align div_eq_of_eq_mul' div_eq_of_eq_mul' +#align sub_eq_of_eq_add' sub_eq_of_eq_add' @[to_additive (attr := simp)] theorem mul_div_mul_left_eq_div (a b c : G) : c * a / (c * b) = a / b := by rw [div_eq_mul_inv, mul_inv_rev, mul_comm b⁻¹ c⁻¹, mul_comm c a, mul_assoc, ←mul_assoc c, mul_right_inv, one_mul, div_eq_mul_inv] +#align mul_div_mul_left_eq_div mul_div_mul_left_eq_div +#align add_sub_add_left_eq_sub add_sub_add_left_eq_sub @[to_additive eq_sub_of_add_eq'] theorem eq_div_of_mul_eq'' (h : c * a = b) : a = b / c := by simp [h.symm] +#align eq_div_of_mul_eq'' eq_div_of_mul_eq'' +#align eq_sub_of_add_eq' eq_sub_of_add_eq' @[to_additive] theorem eq_mul_of_div_eq' (h : a / b = c) : a = b * c := by simp [h.symm] +#align eq_mul_of_div_eq' eq_mul_of_div_eq' +#align eq_add_of_sub_eq' eq_add_of_sub_eq' @[to_additive] theorem mul_eq_of_eq_div' (h : b = c / a) : a * b = c := by simp [h] rw [mul_comm c, mul_inv_cancel_left] +#align mul_eq_of_eq_div' mul_eq_of_eq_div' +#align add_eq_of_eq_sub' add_eq_of_eq_sub' @[to_additive sub_sub_self] theorem div_div_self' (a b : G) : a / (a / b) = b := by simpa using mul_inv_cancel_left a b +#align div_div_self' div_div_self' +#align sub_sub_self sub_sub_self @[to_additive] theorem div_eq_div_mul_div (a b c : G) : a / b = c / b * (a / c) := by simp [mul_left_comm c] +#align div_eq_div_mul_div div_eq_div_mul_div +#align sub_eq_sub_add_sub sub_eq_sub_add_sub @[to_additive (attr := simp)] theorem div_div_cancel (a b : G) : a / (a / b) = b := div_div_self' a b +#align div_div_cancel div_div_cancel +#align sub_sub_cancel sub_sub_cancel @[to_additive (attr := simp)] theorem div_div_cancel_left (a b : G) : a / b / a = b⁻¹ := by simp +#align div_div_cancel_left div_div_cancel_left +#align sub_sub_cancel_left sub_sub_cancel_left @[to_additive eq_sub_iff_add_eq'] theorem eq_div_iff_mul_eq'' : a = b / c ↔ c * a = b := by rw [eq_div_iff_mul_eq', mul_comm] +#align eq_div_iff_mul_eq'' eq_div_iff_mul_eq'' +#align eq_sub_iff_add_eq' eq_sub_iff_add_eq' @[to_additive] theorem div_eq_iff_eq_mul' : a / b = c ↔ a = b * c := by rw [div_eq_iff_eq_mul, mul_comm] +#align div_eq_iff_eq_mul' div_eq_iff_eq_mul' +#align sub_eq_iff_eq_add' sub_eq_iff_eq_add' @[to_additive (attr := simp) add_sub_cancel'] theorem mul_div_cancel''' (a b : G) : a * b / a = b := by rw [div_eq_inv_mul, inv_mul_cancel_left] +#align mul_div_cancel''' mul_div_cancel''' +#align add_sub_cancel' add_sub_cancel' @[to_additive (attr := simp)] theorem mul_div_cancel'_right (a b : G) : a * (b / a) = b := by rw [← mul_div_assoc, mul_div_cancel'''] +#align mul_div_cancel'_right mul_div_cancel'_right +#align add_sub_cancel'_right add_sub_cancel'_right @[to_additive (attr := simp) sub_add_cancel'] theorem div_mul_cancel'' (a b : G) : a / (a * b) = b⁻¹ := by rw [← inv_div, mul_div_cancel'''] +#align div_mul_cancel'' div_mul_cancel'' +#align sub_add_cancel' sub_add_cancel' -- This lemma is in the `simp` set under the name `mul_inv_cancel_comm_assoc`, -- along with the additive version `add_neg_cancel_comm_assoc`, @@ -702,35 +967,51 @@ theorem div_mul_cancel'' (a b : G) : a / (a * b) = b⁻¹ := by rw [← inv_div, @[to_additive] theorem mul_mul_inv_cancel'_right (a b : G) : a * (b * a⁻¹) = b := by rw [← div_eq_mul_inv, mul_div_cancel'_right a b] +#align mul_mul_inv_cancel'_right mul_mul_inv_cancel'_right +#align add_add_neg_cancel'_right add_add_neg_cancel'_right @[to_additive (attr := simp)] theorem mul_mul_div_cancel (a b c : G) : a * c * (b / c) = a * b := by rw [mul_assoc, mul_div_cancel'_right] +#align mul_mul_div_cancel mul_mul_div_cancel +#align add_add_sub_cancel add_add_sub_cancel @[to_additive (attr := simp)] theorem div_mul_mul_cancel (a b c : G) : a / c * (b * c) = a * b := by rw [mul_left_comm, div_mul_cancel', mul_comm] +#align div_mul_mul_cancel div_mul_mul_cancel +#align sub_add_add_cancel sub_add_add_cancel @[to_additive (attr := simp) sub_add_sub_cancel'] theorem div_mul_div_cancel'' (a b c : G) : a / b * (c / a) = c / b := by rw [mul_comm]; apply div_mul_div_cancel' +#align div_mul_div_cancel'' div_mul_div_cancel'' +#align sub_add_sub_cancel' sub_add_sub_cancel' @[to_additive (attr := simp)] theorem mul_div_div_cancel (a b c : G) : a * b / (a / c) = b * c := by rw [← div_mul, mul_div_cancel'''] +#align mul_div_div_cancel mul_div_div_cancel +#align add_sub_sub_cancel add_sub_sub_cancel @[to_additive (attr := simp)] theorem div_div_div_cancel_left (a b c : G) : c / a / (c / b) = b / a := by rw [← inv_div b c, div_inv_eq_mul, mul_comm, div_mul_div_cancel'] +#align div_div_div_cancel_left div_div_div_cancel_left +#align sub_sub_sub_cancel_left sub_sub_sub_cancel_left @[to_additive] theorem div_eq_div_iff_mul_eq_mul : a / b = c / d ↔ a * d = c * b := by rw [div_eq_iff_eq_mul, div_mul_eq_mul_div, eq_comm, div_eq_iff_eq_mul'] simp only [mul_comm, eq_comm] +#align div_eq_div_iff_mul_eq_mul div_eq_div_iff_mul_eq_mul +#align sub_eq_sub_iff_add_eq_add sub_eq_sub_iff_add_eq_add @[to_additive] theorem div_eq_div_iff_div_eq_div : a / b = c / d ↔ a / c = b / d := by rw [div_eq_iff_eq_mul, div_mul_eq_mul_div, div_eq_iff_eq_mul', mul_div_assoc] +#align div_eq_div_iff_div_eq_div div_eq_div_iff_div_eq_div +#align sub_eq_sub_iff_sub_eq_sub sub_eq_sub_iff_sub_eq_sub end CommGroup diff --git a/Mathlib/Algebra/Group/Commute.lean b/Mathlib/Algebra/Group/Commute.lean index 744198896ab57..94f03cf0ce7cc 100644 --- a/Mathlib/Algebra/Group/Commute.lean +++ b/Mathlib/Algebra/Group/Commute.lean @@ -196,6 +196,7 @@ theorem pow_self (a : M) (n : ℕ) : Commute (a ^ n) a := (Commute.refl a).pow_left n #align add_commute.nsmul_self AddCommute.nsmul_selfₓ -- `MulOneClass.toHasMul` vs. `MulOneClass.toMul` +#align commute.pow_self Commute.pow_self -- porting note: `simpNF` told me to remove the `simp` attribute @[to_additive] diff --git a/Mathlib/Algebra/Group/Defs.lean b/Mathlib/Algebra/Group/Defs.lean index 54305307d9abc..a421bd5ca9c66 100644 --- a/Mathlib/Algebra/Group/Defs.lean +++ b/Mathlib/Algebra/Group/Defs.lean @@ -121,10 +121,14 @@ variable [Mul G] /-- `leftMul g` denotes left multiplication by `g` -/ @[to_additive "`left_add g` denotes left addition by `g`"] def leftMul : G → G → G := fun g : G ↦ fun x : G ↦ g * x +#align left_mul leftMul +#align left_add leftAdd /-- `rightMul g` denotes right multiplication by `g` -/ @[to_additive "`right_add g` denotes right addition by `g`"] def rightMul : G → G → G := fun g : G ↦ fun x : G ↦ x * g +#align right_mul rightMul +#align right_add rightAdd /-- A mixin for left cancellative multiplication. -/ class IsLeftCancelMul (G : Type u) [Mul G] : Prop where @@ -136,11 +140,15 @@ class IsRightCancelMul (G : Type u) [Mul G] : Prop where protected mul_right_cancel : ∀ a b c : G, a * b = c * b → a = c /-- A mixin for cancellative multiplication. -/ class IsCancelMul (G : Type u) [Mul G] extends IsLeftCancelMul G, IsRightCancelMul G : Prop +#align is_cancel_mul IsCancelMul +#align is_right_cancel_mul IsRightCancelMul +#align is_left_cancel_mul IsLeftCancelMul /-- A mixin for left cancellative addition. -/ class IsLeftCancelAdd (G : Type u) [Add G] : Prop where /-- Addition is left cancellative. -/ protected add_left_cancel : ∀ a b c : G, a + b = a + c → b = c +#align is_left_cancel_add IsLeftCancelAdd attribute [to_additive IsLeftCancelAdd] IsLeftCancelMul @@ -148,11 +156,13 @@ attribute [to_additive IsLeftCancelAdd] IsLeftCancelMul class IsRightCancelAdd (G : Type u) [Add G] : Prop where /-- Addition is right cancellative. -/ protected add_right_cancel : ∀ a b c : G, a + b = c + b → a = c +#align is_right_cancel_add IsRightCancelAdd attribute [to_additive IsRightCancelAdd] IsRightCancelMul /-- A mixin for cancellative addition. -/ class IsCancelAdd (G : Type u) [Add G] extends IsLeftCancelAdd G, IsRightCancelAdd G : Prop +#align is_cancel_add IsCancelAdd attribute [to_additive IsCancelAdd] IsCancelMul @@ -163,21 +173,31 @@ variable [IsLeftCancelMul G] {a b c : G} @[to_additive] theorem mul_left_cancel : a * b = a * c → b = c := IsLeftCancelMul.mul_left_cancel a b c +#align mul_left_cancel mul_left_cancel +#align add_left_cancel add_left_cancel @[to_additive] theorem mul_left_cancel_iff : a * b = a * c ↔ b = c := ⟨mul_left_cancel, congr_arg _⟩ +#align mul_left_cancel_iff mul_left_cancel_iff +#align add_left_cancel_iff add_left_cancel_iff @[to_additive] theorem mul_right_injective (a : G) : Function.Injective ((· * ·) a) := fun _ _ ↦ mul_left_cancel +#align mul_right_injective mul_right_injective +#align add_right_injective add_right_injective @[to_additive (attr := simp)] theorem mul_right_inj (a : G) {b c : G} : a * b = a * c ↔ b = c := (mul_right_injective a).eq_iff +#align mul_right_inj mul_right_inj +#align add_right_inj add_right_inj @[to_additive] theorem mul_ne_mul_right (a : G) {b c : G} : a * b ≠ a * c ↔ b ≠ c := (mul_right_injective a).ne_iff +#align mul_ne_mul_right mul_ne_mul_right +#align add_ne_add_right add_ne_add_right end IsLeftCancelMul @@ -188,21 +208,31 @@ variable [IsRightCancelMul G] {a b c : G} @[to_additive] theorem mul_right_cancel : a * b = c * b → a = c := IsRightCancelMul.mul_right_cancel a b c +#align mul_right_cancel mul_right_cancel +#align add_right_cancel add_right_cancel @[to_additive] theorem mul_right_cancel_iff : b * a = c * a ↔ b = c := ⟨mul_right_cancel, congr_arg (· * a)⟩ +#align mul_right_cancel_iff mul_right_cancel_iff +#align add_right_cancel_iff add_right_cancel_iff @[to_additive] theorem mul_left_injective (a : G) : Function.Injective (· * a) := fun _ _ ↦ mul_right_cancel +#align mul_left_injective mul_left_injective +#align add_left_injective add_left_injective @[to_additive (attr := simp)] theorem mul_left_inj (a : G) {b c : G} : b * a = c * a ↔ b = c := (mul_left_injective a).eq_iff +#align mul_left_inj mul_left_inj +#align add_left_inj add_left_inj @[to_additive] theorem mul_ne_mul_left (a : G) {b c : G} : b * a ≠ c * a ↔ b ≠ c := (mul_left_injective a).ne_iff +#align mul_ne_mul_left mul_ne_mul_left +#align add_ne_add_left add_ne_add_left end IsRightCancelMul @@ -213,12 +243,18 @@ end Mul class Semigroup (G : Type u) extends Mul G where /-- Multiplication is associative -/ mul_assoc : ∀ a b c : G, a * b * c = a * (b * c) +#align semigroup Semigroup +#align semigroup.ext Semigroup.ext +#align semigroup.ext_iff Semigroup.ext_iff /-- An additive semigroup is a type with an associative `(+)`. -/ @[ext] class AddSemigroup (G : Type u) extends Add G where /-- Addition is associative -/ add_assoc : ∀ a b c : G, a + b + c = a + (b + c) +#align add_semigroup AddSemigroup +#align add_semigroup.ext AddSemigroup.ext +#align add_semigroup.ext_iff AddSemigroup.ext_iff attribute [to_additive] Semigroup @@ -229,6 +265,8 @@ variable [Semigroup G] @[to_additive] theorem mul_assoc : ∀ a b c : G, a * b * c = a * (b * c) := Semigroup.mul_assoc +#align mul_assoc mul_assoc +#align add_assoc add_assoc @[to_additive] instance Semigroup.to_isAssociative : IsAssociative G (· * ·) := @@ -243,12 +281,18 @@ end Semigroup class CommSemigroup (G : Type u) extends Semigroup G where /-- Multiplication is commutative in a commutative semigroup. -/ mul_comm : ∀ a b : G, a * b = b * a +#align comm_semigroup CommSemigroup +#align comm_semigroup.ext_iff CommSemigroup.ext_iff +#align comm_semigroup.ext CommSemigroup.ext /-- A commutative additive semigroup is a type with an associative commutative `(+)`. -/ @[ext] class AddCommSemigroup (G : Type u) extends AddSemigroup G where /-- Addition is commutative in an additive commutative semigroup. -/ add_comm : ∀ a b : G, a + b = b + a +#align add_comm_semigroup AddCommSemigroup +#align add_comm_semigroup.ext AddCommSemigroup.ext +#align add_comm_semigroup.ext_iff AddCommSemigroup.ext_iff attribute [to_additive] CommSemigroup @@ -259,6 +303,8 @@ variable [CommSemigroup G] @[to_additive] theorem mul_comm : ∀ a b : G, a * b = b * a := CommSemigroup.mul_comm +#align mul_comm mul_comm +#align add_comm add_comm @[to_additive] instance CommSemigroup.to_isCommutative : IsCommutative G (· * ·) := @@ -316,12 +362,18 @@ end CommSemigroup @[ext] class LeftCancelSemigroup (G : Type u) extends Semigroup G where mul_left_cancel : ∀ a b c : G, a * b = a * c → b = c +#align left_cancel_semigroup LeftCancelSemigroup +#align left_cancel_semigroup.ext_iff LeftCancelSemigroup.ext_iff +#align left_cancel_semigroup.ext LeftCancelSemigroup.ext /-- An `AddLeftCancelSemigroup` is an additive semigroup such that `a + b = a + c` implies `b = c`. -/ @[ext] class AddLeftCancelSemigroup (G : Type u) extends AddSemigroup G where add_left_cancel : ∀ a b c : G, a + b = a + c → b = c +#align add_left_cancel_semigroup AddLeftCancelSemigroup +#align add_left_cancel_semigroup.ext AddLeftCancelSemigroup.ext +#align add_left_cancel_semigroup.ext_iff AddLeftCancelSemigroup.ext_iff attribute [to_additive] LeftCancelSemigroup @@ -338,12 +390,18 @@ instance (priority := 100) LeftCancelSemigroup.toIsLeftCancelMul (G : Type u) @[ext] class RightCancelSemigroup (G : Type u) extends Semigroup G where mul_right_cancel : ∀ a b c : G, a * b = c * b → a = c +#align right_cancel_semigroup RightCancelSemigroup +#align right_cancel_semigroup.ext_iff RightCancelSemigroup.ext_iff +#align right_cancel_semigroup.ext RightCancelSemigroup.ext /-- An `AddRightCancelSemigroup` is an additive semigroup such that `a + b = c + b` implies `a = c`. -/ @[ext] class AddRightCancelSemigroup (G : Type u) extends AddSemigroup G where add_right_cancel : ∀ a b c : G, a + b = c + b → a = c +#align add_right_cancel_semigroup AddRightCancelSemigroup +#align add_right_cancel_semigroup.ext_iff AddRightCancelSemigroup.ext_iff +#align add_right_cancel_semigroup.ext AddRightCancelSemigroup.ext attribute [to_additive] RightCancelSemigroup @@ -363,6 +421,7 @@ class MulOneClass (M : Type u) extends One M, Mul M where one_mul : ∀ a : M, 1 * a = a /-- One is a right neutral element for multiplication -/ mul_one : ∀ a : M, a * 1 = a +#align mul_one_class MulOneClass /-- Typeclass for expressing that a type `M` with addition and a zero satisfies `0 + a = a` and `a + 0 = a` for all `a : M`. -/ @@ -371,6 +430,7 @@ class AddZeroClass (M : Type u) extends Zero M, Add M where zero_add : ∀ a : M, 0 + a = a /-- Zero is a right neutral element for addition -/ add_zero : ∀ a : M, a + 0 = a +#align add_zero_class AddZeroClass attribute [to_additive] MulOneClass @@ -381,6 +441,8 @@ theorem MulOneClass.ext {M : Type u} : ∀ ⦃m₁ m₂ : MulOneClass M⦄, m₁ -- congr suffices one₁ = one₂ by cases this; rfl exact (one_mul₂ one₁).symm.trans (mul_one₁ one₂) +#align mul_one_class.ext MulOneClass.ext +#align add_zero_class.ext AddZeroClass.ext section MulOneClass @@ -389,10 +451,14 @@ variable {M : Type u} [MulOneClass M] @[to_additive (attr := simp)] theorem one_mul : ∀ a : M, 1 * a = a := MulOneClass.one_mul +#align one_mul one_mul +#align zero_add zero_add @[to_additive (attr := simp)] theorem mul_one : ∀ a : M, a * 1 = a := MulOneClass.mul_one +#align mul_one mul_one +#align add_zero add_zero end MulOneClass @@ -407,12 +473,14 @@ Use instead `a ^ n`, which has better definitional behavior. -/ def npowRec [One M] [Mul M] : ℕ → M → M | 0, _ => 1 | n + 1, a => a * npowRec n a +#align npow_rec npowRec /-- The fundamental scalar multiplication in an additive monoid. `nsmulRec n a = a+a+...+a` n times. Use instead `n • a`, which has better definitional behavior. -/ def nsmulRec [Zero M] [Add M] : ℕ → M → M | 0, _ => 0 | n + 1, a => a + nsmulRec n a +#align nsmul_rec nsmulRec attribute [to_additive] npowRec @@ -509,6 +577,7 @@ class AddMonoid (M : Type u) extends AddSemigroup M, AddZeroClass M where nsmul_zero : ∀ x, nsmul 0 x = 0 := by intros; rfl /-- Multiplication by `(n + 1 : ℕ)` behaves as expected. -/ nsmul_succ : ∀ (n : ℕ) (x), nsmul (n + 1) x = x + nsmul n x := by intros; rfl +#align add_monoid AddMonoid #align add_monoid.nsmul_zero' AddMonoid.nsmul_zero #align add_monoid.nsmul_succ' AddMonoid.nsmul_succ @@ -522,6 +591,7 @@ class Monoid (M : Type u) extends Semigroup M, MulOneClass M where npow_zero : ∀ x, npow 0 x = 1 := by intros; rfl /-- Raising to the power `(n + 1 : ℕ)` behaves as expected. -/ npow_succ : ∀ (n : ℕ) (x), npow (n + 1) x = x * npow n x := by intros; rfl +#align monoid Monoid #align monoid.npow_zero' Monoid.npow_zero #align monoid.npow_succ' Monoid.npow_succ @@ -546,15 +616,21 @@ variable {M : Type _} [Monoid M] @[to_additive (attr := simp) nsmul_eq_smul] theorem npow_eq_pow (n : ℕ) (x : M) : Monoid.npow n x = x ^ n := rfl +#align npow_eq_pow npow_eq_pow +#align nsmul_eq_smul nsmul_eq_smul -- the attributes are intentionally out of order. `zero_smul` proves `zero_nsmul`. @[to_additive zero_nsmul, simp] theorem pow_zero (a : M) : a ^ 0 = 1 := Monoid.npow_zero _ +#align pow_zero pow_zero +#align zero_nsmul zero_nsmul @[to_additive succ_nsmul] theorem pow_succ (a : M) (n : ℕ) : a ^ (n + 1) = a * a ^ n := Monoid.npow_succ n a +#align pow_succ pow_succ +#align succ_nsmul succ_nsmul end @@ -565,15 +641,19 @@ variable {M : Type u} [Monoid M] @[to_additive] theorem left_inv_eq_right_inv {a b c : M} (hba : b * a = 1) (hac : a * c = 1) : b = c := by rw [← one_mul c, ← hba, mul_assoc, hac, mul_one b] +#align left_inv_eq_right_inv left_inv_eq_right_inv +#align left_neg_eq_right_neg left_neg_eq_right_neg end Monoid /-- An additive commutative monoid is an additive monoid with commutative `(+)`. -/ class AddCommMonoid (M : Type u) extends AddMonoid M, AddCommSemigroup M +#align add_comm_monoid AddCommMonoid /-- A commutative monoid is a monoid with commutative `(*)`. -/ @[to_additive] class CommMonoid (M : Type u) extends Monoid M, CommSemigroup M +#align comm_monoid CommMonoid attribute [to_additive] CommMonoid.toCommSemigroup @@ -583,10 +663,12 @@ section LeftCancelMonoid Main examples are `ℕ` and groups. This is the right typeclass for many sum lemmas, as having a zero is useful to define the sum over the empty set, so `AddLeftCancelSemigroup` is not enough. -/ class AddLeftCancelMonoid (M : Type u) extends AddLeftCancelSemigroup M, AddMonoid M +#align add_left_cancel_monoid AddLeftCancelMonoid /-- A monoid in which multiplication is left-cancellative. -/ @[to_additive] class LeftCancelMonoid (M : Type u) extends LeftCancelSemigroup M, Monoid M +#align left_cancel_monoid LeftCancelMonoid attribute [to_additive] LeftCancelMonoid.toMonoid @@ -598,10 +680,12 @@ section RightCancelMonoid Main examples are `ℕ` and groups. This is the right typeclass for many sum lemmas, as having a zero is useful to define the sum over the empty set, so `AddRightCancelSemigroup` is not enough. -/ class AddRightCancelMonoid (M : Type u) extends AddRightCancelSemigroup M, AddMonoid M +#align add_right_cancel_monoid AddRightCancelMonoid /-- A monoid in which multiplication is right-cancellative. -/ @[to_additive] class RightCancelMonoid (M : Type u) extends RightCancelSemigroup M, Monoid M +#align right_cancel_monoid RightCancelMonoid attribute [to_additive] RightCancelMonoid.toMonoid @@ -613,19 +697,23 @@ section CancelMonoid Main examples are `ℕ` and groups. This is the right typeclass for many sum lemmas, as having a zero is useful to define the sum over the empty set, so `AddRightCancelMonoid` is not enough. -/ class AddCancelMonoid (M : Type u) extends AddLeftCancelMonoid M, AddRightCancelMonoid M +#align add_cancel_monoid AddCancelMonoid /-- A monoid in which multiplication is cancellative. -/ @[to_additive] class CancelMonoid (M : Type u) extends LeftCancelMonoid M, RightCancelMonoid M +#align cancel_monoid CancelMonoid attribute [to_additive] CancelMonoid.toRightCancelMonoid /-- Commutative version of `AddCancelMonoid`. -/ class AddCancelCommMonoid (M : Type u) extends AddLeftCancelMonoid M, AddCommMonoid M +#align add_cancel_comm_monoid AddCancelCommMonoid /-- Commutative version of `CancelMonoid`. -/ @[to_additive] class CancelCommMonoid (M : Type u) extends LeftCancelMonoid M, CommMonoid M +#align cancel_comm_monoid CancelCommMonoid attribute [to_additive] CancelCommMonoid.toCommMonoid @@ -653,12 +741,14 @@ Use instead `a ^ n`, which has better definitional behavior. -/ def zpowRec {M : Type _} [One M] [Mul M] [Inv M] : ℤ → M → M | Int.ofNat n, a => npowRec n a | Int.negSucc n, a => (npowRec n.succ a)⁻¹ +#align zpow_rec zpowRec /-- The fundamental scalar multiplication in an additive group. `zpowRec n a = a+a+...+a` n times, for integer `n`. Use instead `n • a`, which has better definitional behavior. -/ def zsmulRec {M : Type _} [Zero M] [Add M] [Neg M] : ℤ → M → M | Int.ofNat n, a => nsmulRec n a | Int.negSucc n, a => -nsmulRec n.succ a +#align zsmul_rec zsmulRec attribute [to_additive] zpowRec @@ -682,6 +772,8 @@ variable [InvolutiveInv G] @[to_additive (attr := simp)] theorem inv_inv (a : G) : a⁻¹⁻¹ = a := InvolutiveInv.inv_inv _ +#align inv_inv inv_inv +#align neg_neg neg_neg end InvolutiveInv @@ -749,6 +841,7 @@ class DivInvMonoid (G : Type u) extends Monoid G, Inv G, Div G where intros; rfl /-- `a ^ -(n + 1) = (a ^ (n + 1))⁻¹` -/ zpow_neg' (n : ℕ) (a : G) : zpow (Int.negSucc n) a = (zpow n.succ a)⁻¹ := by intros; rfl +#align div_inv_monoid DivInvMonoid /-- A `SubNegMonoid` is an `AddMonoid` with unary `-` and binary `-` operations satisfying `sub_eq_add_neg : ∀ a b, a - b = a + -b`. @@ -775,6 +868,7 @@ class SubNegMonoid (G : Type u) extends AddMonoid G, Neg G, Sub G where zsmul_succ' (n : ℕ) (a : G) : zsmul (Int.ofNat n.succ) a = a + zsmul (Int.ofNat n) a := by intros; rfl zsmul_neg' (n : ℕ) (a : G) : zsmul (Int.negSucc n) a = -zsmul n.succ a := by intros; rfl +#align sub_neg_monoid SubNegMonoid attribute [to_additive SubNegMonoid] DivInvMonoid @@ -795,9 +889,13 @@ variable [DivInvMonoid G] {a b : G} @[to_additive (attr := simp) zsmul_eq_smul] theorem zpow_eq_pow (n : ℤ) (x : G) : DivInvMonoid.zpow n x = x ^ n := rfl +#align zsmul_eq_smul zsmul_eq_smul +#align zpow_eq_pow zpow_eq_pow @[to_additive (attr := simp) zero_zsmul] theorem zpow_zero (a : G) : a ^ (0 : ℤ) = 1 := DivInvMonoid.zpow_zero' a +#align zpow_zero zpow_zero +#align zero_zsmul zero_zsmul @[to_additive (attr := norm_cast) ofNat_zsmul] theorem zpow_ofNat (a : G) : ∀ n : ℕ, a ^ (n : ℤ) = a ^ n @@ -831,8 +929,11 @@ This is a duplicate of `DivInvMonoid.div_eq_mul_inv` ensuring that the types unf This is a duplicate of `SubNegMonoid.sub_eq_mul_neg` ensuring that the types unfold better."] theorem div_eq_mul_inv (a b : G) : a / b = a * b⁻¹ := DivInvMonoid.div_eq_mul_inv _ _ +#align div_eq_mul_inv div_eq_mul_inv +#align sub_eq_add_neg sub_eq_add_neg alias div_eq_mul_inv ← division_def +#align division_def division_def end DivInvMonoid @@ -841,18 +942,22 @@ section InvOneClass /-- Typeclass for expressing that `-0 = 0`. -/ class NegZeroClass (G : Type _) extends Zero G, Neg G where neg_zero : -(0 : G) = 0 +#align neg_zero_class NegZeroClass /-- A `SubNegMonoid` where `-0 = 0`. -/ class SubNegZeroMonoid (G : Type _) extends SubNegMonoid G, NegZeroClass G +#align sub_neg_zero_monoid SubNegZeroMonoid /-- Typeclass for expressing that `1⁻¹ = 1`. -/ @[to_additive] class InvOneClass (G : Type _) extends One G, Inv G where inv_one : (1 : G)⁻¹ = 1 +#align inv_one_class InvOneClass /-- A `DivInvMonoid` where `1⁻¹ = 1`. -/ @[to_additive SubNegZeroMonoid] class DivInvOneMonoid (G : Type _) extends DivInvMonoid G, InvOneClass G +#align div_inv_one_monoid DivInvOneMonoid -- FIXME: `to_additive` is not operating on the second parent. (#660) attribute [to_additive] DivInvOneMonoid.toInvOneClass @@ -862,6 +967,8 @@ variable [InvOneClass G] @[to_additive (attr := simp)] theorem inv_one : (1 : G)⁻¹ = 1 := InvOneClass.inv_one +#align inv_one inv_one +#align neg_zero neg_zero end InvOneClass @@ -872,6 +979,7 @@ class SubtractionMonoid (G : Type u) extends SubNegMonoid G, InvolutiveNeg G whe /- Despite the asymmetry of `neg_eq_of_add`, the symmetric version is true thanks to the involutivity of negation. -/ neg_eq_of_add (a b : G) : a + b = 0 → -a = b +#align subtraction_monoid SubtractionMonoid /-- A `DivisionMonoid` is a `DivInvMonoid` with involutive inversion and such that `(a * b)⁻¹ = b⁻¹ * a⁻¹` and `a * b = 1 → a⁻¹ = b`. @@ -883,6 +991,7 @@ class DivisionMonoid (G : Type u) extends DivInvMonoid G, InvolutiveInv G where /- Despite the asymmetry of `inv_eq_of_mul`, the symmetric version is true thanks to the involutivity of inversion. -/ inv_eq_of_mul (a b : G) : a * b = 1 → a⁻¹ = b +#align division_monoid DivisionMonoid attribute [to_additive] DivisionMonoid.toInvolutiveInv @@ -893,21 +1002,27 @@ variable [DivisionMonoid G] {a b : G} @[to_additive (attr := simp) neg_add_rev] theorem mul_inv_rev (a b : G) : (a * b)⁻¹ = b⁻¹ * a⁻¹ := DivisionMonoid.mul_inv_rev _ _ +#align mul_inv_rev mul_inv_rev +#align neg_add_rev neg_add_rev @[to_additive] theorem inv_eq_of_mul_eq_one_right : a * b = 1 → a⁻¹ = b := DivisionMonoid.inv_eq_of_mul _ _ +#align inv_eq_of_mul_eq_one_right inv_eq_of_mul_eq_one_right +#align neg_eq_of_add_eq_zero_right neg_eq_of_add_eq_zero_right end DivisionMonoid /-- Commutative `SubtractionMonoid`. -/ class SubtractionCommMonoid (G : Type u) extends SubtractionMonoid G, AddCommMonoid G +#align subtraction_comm_monoid SubtractionCommMonoid /-- Commutative `DivisionMonoid`. This is the immediate common ancestor of `CommGroup` and `CommGroupWithZero`. -/ @[to_additive SubtractionCommMonoid] class DivisionCommMonoid (G : Type u) extends DivisionMonoid G, CommMonoid G +#align division_comm_monoid DivisionCommMonoid attribute [to_additive] DivisionCommMonoid.toCommMonoid @@ -918,6 +1033,7 @@ with a default so that `a / b = a * b⁻¹` holds by definition. -/ class Group (G : Type u) extends DivInvMonoid G where mul_left_inv : ∀ a : G, a⁻¹ * a = 1 +#align group Group /-- An `AddGroup` is an `AddMonoid` with a unary `-` satisfying `-a + a = 0`. @@ -926,6 +1042,7 @@ with a default so that `a - b = a + -b` holds by definition. -/ class AddGroup (A : Type u) extends SubNegMonoid A where add_left_neg : ∀ a : A, -a + a = 0 +#align add_group AddGroup attribute [to_additive] Group @@ -936,10 +1053,14 @@ variable [Group G] {a b c : G} @[to_additive (attr := simp)] theorem mul_left_inv : ∀ a : G, a⁻¹ * a = 1 := Group.mul_left_inv +#align mul_left_inv mul_left_inv +#align add_left_neg add_left_neg @[to_additive] theorem inv_mul_self (a : G) : a⁻¹ * a = 1 := mul_left_inv a +#align inv_mul_self inv_mul_self +#align neg_add_self neg_add_self @[to_additive] private theorem inv_eq_of_mul (h : a * b = 1) : a⁻¹ = b := @@ -948,26 +1069,38 @@ private theorem inv_eq_of_mul (h : a * b = 1) : a⁻¹ = b := @[to_additive (attr := simp)] theorem mul_right_inv (a : G) : a * a⁻¹ = 1 := by rw [← mul_left_inv a⁻¹, inv_eq_of_mul (mul_left_inv a)] +#align mul_right_inv mul_right_inv +#align add_right_neg add_right_neg @[to_additive] theorem mul_inv_self (a : G) : a * a⁻¹ = 1 := mul_right_inv a +#align mul_inv_self mul_inv_self +#align add_neg_self add_neg_self @[to_additive (attr := simp)] theorem inv_mul_cancel_left (a b : G) : a⁻¹ * (a * b) = b := by rw [← mul_assoc, mul_left_inv, one_mul] +#align inv_mul_cancel_left inv_mul_cancel_left +#align neg_add_cancel_left neg_add_cancel_left @[to_additive (attr := simp)] theorem mul_inv_cancel_left (a b : G) : a * (a⁻¹ * b) = b := by rw [← mul_assoc, mul_right_inv, one_mul] +#align mul_inv_cancel_left mul_inv_cancel_left +#align add_neg_cancel_left add_neg_cancel_left @[to_additive (attr := simp)] theorem mul_inv_cancel_right (a b : G) : a * b * b⁻¹ = a := by rw [mul_assoc, mul_right_inv, mul_one] +#align mul_inv_cancel_right mul_inv_cancel_right +#align add_neg_cancel_right add_neg_cancel_right @[to_additive (attr := simp)] theorem inv_mul_cancel_right (a b : G) : a * b⁻¹ * b = a := by rw [mul_assoc, mul_left_inv, mul_one] +#align inv_mul_cancel_right inv_mul_cancel_right +#align neg_add_cancel_right neg_add_cancel_right @[to_additive AddGroup.toSubtractionMonoid] instance (priority := 100) Group.toDivisionMonoid : DivisionMonoid G := @@ -993,10 +1126,12 @@ theorem Group.toDivInvMonoid_injective {G : Type _} : /-- An additive commutative group is an additive group with commutative `(+)`. -/ class AddCommGroup (G : Type u) extends AddGroup G, AddCommMonoid G +#align add_comm_group AddCommGroup /-- A commutative group is a group with commutative `(*)`. -/ @[to_additive] class CommGroup (G : Type u) extends Group G, CommMonoid G +#align comm_group CommGroup attribute [to_additive] CommGroup.toCommMonoid diff --git a/Mathlib/Algebra/Group/InjSurj.lean b/Mathlib/Algebra/Group/InjSurj.lean index de4177a6a0acf..ca8f2856f8cbc 100644 --- a/Mathlib/Algebra/Group/InjSurj.lean +++ b/Mathlib/Algebra/Group/InjSurj.lean @@ -167,6 +167,7 @@ protected def cancelMonoid [CancelMonoid M₂] (f : M₁ → M₂) (hf : Injecti CancelMonoid M₁ := { hf.leftCancelMonoid f one mul npow, hf.rightCancelMonoid f one mul npow with } #align function.injective.add_cancel_monoid Function.Injective.addCancelMonoid +#align function.injective.cancel_monoid Function.Injective.cancelMonoid /-- A type endowed with `1` and `*` is a commutative monoid, if it admits an injective map that preserves `1` and `*` to a commutative monoid. See note [reducible non-instances]. -/ diff --git a/Mathlib/Algebra/Group/Opposite.lean b/Mathlib/Algebra/Group/Opposite.lean index 07c3f77d2ae9a..209772896a804 100644 --- a/Mathlib/Algebra/Group/Opposite.lean +++ b/Mathlib/Algebra/Group/Opposite.lean @@ -238,6 +238,7 @@ attribute [nolint simpComm] AddOpposite.commute_unop def opAddEquiv [Add α] : α ≃+ αᵐᵒᵖ := { opEquiv with map_add' := fun _ _ => rfl } #align mul_opposite.op_add_equiv MulOpposite.opAddEquiv +#align mul_opposite.op_add_equiv_symm_apply MulOpposite.opAddEquiv_symmApply @[simp] theorem opAddEquiv_toEquiv [Add α] : (opAddEquiv : α ≃+ αᵐᵒᵖ).toEquiv = opEquiv := rfl @@ -305,6 +306,7 @@ variable {α} def opMulEquiv [Mul α] : α ≃* αᵃᵒᵖ := { opEquiv with map_mul' := fun _ _ => rfl } #align add_opposite.op_mul_equiv AddOpposite.opMulEquiv +#align add_opposite.op_mul_equiv_symm_apply AddOpposite.opMulEquiv_symmApply @[simp] theorem opMulEquiv_toEquiv [Mul α] : (opMulEquiv : α ≃* αᵃᵒᵖ).toEquiv = opEquiv := @@ -325,6 +327,7 @@ def MulEquiv.inv' (G : Type _) [DivisionMonoid G] : G ≃* Gᵐᵒᵖ := { (Equiv.inv G).trans opEquiv with map_mul' := fun x y => unop_injective <| mul_inv_rev x y } #align mul_equiv.inv' MulEquiv.inv' #align add_equiv.neg' AddEquiv.neg' +#align mul_equiv.inv'_symm_apply MulEquiv.inv'_symmApply /-- A semigroup homomorphism `f : M →ₙ* N` such that `f x` commutes with `f y` for all `x, y` defines a semigroup homomorphism to `Nᵐᵒᵖ`. -/ @@ -338,6 +341,7 @@ def MulHom.toOpposite {M N : Type _} [Mul M] [Mul N] (f : M →ₙ* N) map_mul' x y := by simp [(hf x y).eq] #align mul_hom.to_opposite MulHom.toOpposite #align add_hom.to_opposite AddHom.toOpposite +#align mul_hom.to_opposite_apply MulHom.toOpposite_apply /-- A semigroup homomorphism `f : M →ₙ* N` such that `f x` commutes with `f y` for all `x, y` defines a semigroup homomorphism from `Mᵐᵒᵖ`. -/ @@ -351,6 +355,7 @@ def MulHom.fromOpposite {M N : Type _} [Mul M] [Mul N] (f : M →ₙ* N) map_mul' _ _ := (f.map_mul _ _).trans (hf _ _).eq #align mul_hom.from_opposite MulHom.fromOpposite #align add_hom.from_opposite AddHom.fromOpposite +#align mul_hom.from_opposite_apply MulHom.fromOpposite_apply /-- A monoid homomorphism `f : M →* N` such that `f x` commutes with `f y` for all `x, y` defines a monoid homomorphism to `Nᵐᵒᵖ`. -/ @@ -365,6 +370,7 @@ def MonoidHom.toOpposite {M N : Type _} [MulOneClass M] [MulOneClass N] (f : M map_mul' x y := by simp [(hf x y).eq] #align monoid_hom.to_opposite MonoidHom.toOpposite #align add_monoid_hom.to_opposite AddMonoidHom.toOpposite +#align monoid_hom.to_opposite_apply MonoidHom.toOpposite_apply /-- A monoid homomorphism `f : M →* N` such that `f x` commutes with `f y` for all `x, y` defines a monoid homomorphism from `Mᵐᵒᵖ`. -/ @@ -379,6 +385,7 @@ def MonoidHom.fromOpposite {M N : Type _} [MulOneClass M] [MulOneClass N] (f : M map_mul' _ _ := (f.map_mul _ _).trans (hf _ _).eq #align monoid_hom.from_opposite MonoidHom.fromOpposite #align add_monoid_hom.from_opposite AddMonoidHom.fromOpposite +#align monoid_hom.from_opposite_apply MonoidHom.fromOpposite_apply /-- The units of the opposites are equivalent to the opposites of the units. -/ @[to_additive @@ -429,6 +436,8 @@ def MulHom.op {M N} [Mul M] [Mul N] : (M →ₙ* N) ≃ (Mᵐᵒᵖ →ₙ* Nᵐ simp #align mul_hom.op MulHom.op #align add_hom.op AddHom.op +#align mul_hom.op_symm_apply_apply MulHom.op_symm_apply_apply +#align mul_hom.op_apply_apply MulHom.op_apply_apply /-- The 'unopposite' of a semigroup homomorphism `Mᵐᵒᵖ →ₙ* Nᵐᵒᵖ`. Inverse to `MulHom.op`. -/ @[simp, @@ -461,6 +470,8 @@ def AddHom.mulOp {M N} [Add M] [Add N] : AddHom M N ≃ AddHom Mᵐᵒᵖ Nᵐ intro x simp #align add_hom.mul_op AddHom.mulOp +#align add_hom.mul_op_symm_apply_apply AddHom.mulOp_symm_apply_apply +#align add_hom.mul_op_apply_apply AddHom.mulOp_apply_apply /-- The 'unopposite' of an additive semigroup hom `αᵐᵒᵖ →+ βᵐᵒᵖ`. Inverse to `AddHom.mul_op`. -/ @@ -491,6 +502,8 @@ def MonoidHom.op {M N} [MulOneClass M] [MulOneClass N] : (M →* N) ≃ (Mᵐᵒ simp #align monoid_hom.op MonoidHom.op #align add_monoid_hom.op AddMonoidHom.op +#align monoid_hom.op_apply_apply MonoidHom.op_apply_apply +#align monoid_hom.op_symm_apply_apply MonoidHom.op_symm_apply_apply /-- The 'unopposite' of a monoid homomorphism `Mᵐᵒᵖ →* Nᵐᵒᵖ`. Inverse to `MonoidHom.op`. -/ @[simp, @@ -523,6 +536,8 @@ def AddMonoidHom.mulOp {M N} [AddZeroClass M] [AddZeroClass N] : (M →+ N) ≃ intro simp #align add_monoid_hom.mul_op AddMonoidHom.mulOp +#align add_monoid_hom.mul_op_symm_apply_apply AddMonoidHom.mulOp_symm_apply_apply +#align add_monoid_hom.mul_op_apply_apply AddMonoidHom.mulOp_apply_apply /-- The 'unopposite' of an additive monoid hom `αᵐᵒᵖ →+ βᵐᵒᵖ`. Inverse to `AddMonoidHom.mul_op`. -/ @@ -545,6 +560,8 @@ def AddEquiv.mulOp {α β} [Add α] [Add β] : α ≃+ β ≃ (αᵐᵒᵖ ≃+ intro rfl #align add_equiv.mul_op AddEquiv.mulOp +#align add_equiv.mul_op_apply AddEquiv.mulOp_apply +#align add_equiv.mul_op_symm_apply AddEquiv.mulOp_symm_apply /-- The 'unopposite' of an iso `αᵐᵒᵖ ≃+ βᵐᵒᵖ`. Inverse to `AddEquiv.mul_op`. -/ @[simp] @@ -575,6 +592,10 @@ def MulEquiv.op {α β} [Mul α] [Mul β] : α ≃* β ≃ (αᵐᵒᵖ ≃* β rfl #align mul_equiv.op MulEquiv.op #align add_equiv.op AddEquiv.op +#align mul_equiv.op_symm_apply_symm_apply MulEquiv.op_symm_apply_symmApply +#align mul_equiv.op_apply_apply MulEquiv.op_apply_apply +#align mul_equiv.op_apply_symm_apply MulEquiv.op_apply_symmApply +#align mul_equiv.op_symm_apply_apply MulEquiv.op_symm_apply_apply /-- The 'unopposite' of an iso `αᵐᵒᵖ ≃* βᵐᵒᵖ`. Inverse to `MulEquiv.op`. -/ @[to_additive (attr := simp) diff --git a/Mathlib/Algebra/Group/OrderSynonym.lean b/Mathlib/Algebra/Group/OrderSynonym.lean index 59c3013d17402..81c17f1603683 100644 --- a/Mathlib/Algebra/Group/OrderSynonym.lean +++ b/Mathlib/Algebra/Group/OrderSynonym.lean @@ -290,11 +290,13 @@ theorem ofLex_pow [Pow α β] (a : Lex α) (b : β) : ofLex (a ^ b) = ofLex a ^ theorem pow_toLex [Pow α β] (a : α) (b : β) : a ^ toLex b = a ^ b := rfl #align pow_to_lex pow_toLex -- Porting note: Duplicate to_additive name in Lean 3 => no lemma? +#align to_lex_smul' toLex_smul' @[to_additive (attr := simp) (reorder := 1 4) ofLex_smul'] theorem pow_ofLex [Pow α β] (a : α) (b : Lex β) : a ^ ofLex b = a ^ b := rfl #align pow_of_lex pow_ofLex -- Porting note: Duplicate to_additive name in Lean 3 => no lemma? +#align of_lex_smul' ofLex_smul' attribute [to_additive] instSMulOrderDual instSMulOrderDual' instSMulLex instSMulLex' attribute [to_additive (attr := simp)] diff --git a/Mathlib/Algebra/Group/Pi.lean b/Mathlib/Algebra/Group/Pi.lean index ed5a8cb1c41cb..140ac13da91a7 100644 --- a/Mathlib/Algebra/Group/Pi.lean +++ b/Mathlib/Algebra/Group/Pi.lean @@ -256,6 +256,7 @@ def Pi.mulHom {γ : Type w} [∀ i, Mul (f i)] [Mul γ] (g : ∀ i, γ →ₙ* f map_mul' x y := funext fun i => (g i).map_mul x y #align pi.mul_hom Pi.mulHom #align pi.add_hom Pi.addHom +#align pi.mul_hom_apply Pi.mulHom_apply @[to_additive] theorem Pi.mulHom_injective {γ : Type w} [Nonempty I] [∀ i, Mul (f i)] [Mul γ] (g : ∀ i, γ →ₙ* f i) @@ -278,6 +279,7 @@ def Pi.monoidHom {γ : Type w} [∀ i, MulOneClass (f i)] [MulOneClass γ] (g : map_one' := funext fun i => (g i).map_one } #align pi.monoid_hom Pi.monoidHom #align pi.add_monoid_hom Pi.addMonoidHom +#align pi.monoid_hom_apply Pi.monoidHom_apply @[to_additive] theorem Pi.monoidHom_injective {γ : Type w} [Nonempty I] [∀ i, MulOneClass (f i)] [MulOneClass γ] @@ -301,6 +303,7 @@ def Pi.evalMulHom (i : I) : (∀ i, f i) →ₙ* f i where map_mul' _ _ := Pi.mul_apply _ _ i #align pi.eval_mul_hom Pi.evalMulHom #align pi.eval_add_hom Pi.evalAddHom +#align pi.eval_mul_hom_apply Pi.evalMulHom_apply /-- `Function.const` as a `MulHom`. -/ @[to_additive "`Function.const` as an `AddHom`.", simps] @@ -310,6 +313,7 @@ def Pi.constMulHom (α β : Type _) [Mul β] : map_mul' _ _ := rfl #align pi.const_mul_hom Pi.constMulHom #align pi.const_add_hom Pi.constAddHom +#align pi.const_mul_hom_apply Pi.constMulHom_apply /-- Coercion of a `MulHom` into a function is itself a `MulHom`. @@ -325,6 +329,7 @@ def MulHom.coeFn (α β : Type _) [Mul α] [CommSemigroup β] : map_mul' _ _ := rfl #align mul_hom.coe_fn MulHom.coeFn #align add_hom.coe_fn AddHom.coeFn +#align mul_hom.coe_fn_apply MulHom.coeFn_apply /-- Semigroup homomorphism between the function spaces `I → α` and `I → β`, induced by a semigroup homomorphism `f` between `α` and `β`. -/ @@ -338,6 +343,7 @@ protected def MulHom.compLeft {α β : Type _} [Mul α] [Mul β] (f : α →ₙ* map_mul' _ _ := by ext; simp #align mul_hom.comp_left MulHom.compLeft #align add_hom.comp_left AddHom.compLeft +#align mul_hom.comp_left_apply MulHom.compLeft_apply end MulHom @@ -359,6 +365,7 @@ def Pi.evalMonoidHom (i : I) : map_mul' _ _ := Pi.mul_apply _ _ i #align pi.eval_monoid_hom Pi.evalMonoidHom #align pi.eval_add_monoid_hom Pi.evalAddMonoidHom +#align pi.eval_monoid_hom_apply Pi.evalMonoidHom_apply /-- `Function.const` as a `MonoidHom`. -/ @[to_additive "`Function.const` as an `AddMonoidHom`.", simps] @@ -369,6 +376,7 @@ def Pi.constMonoidHom (α β : Type _) [MulOneClass β] : map_mul' _ _ := rfl #align pi.const_monoid_hom Pi.constMonoidHom #align pi.const_add_monoid_hom Pi.constAddMonoidHom +#align pi.const_monoid_hom_apply Pi.constMonoidHom_apply /-- Coercion of a `MonoidHom` into a function is itself a `MonoidHom`. @@ -385,6 +393,7 @@ def MonoidHom.coeFn (α β : Type _) [MulOneClass α] [CommMonoid β] : map_mul' _ _ := rfl #align monoid_hom.coe_fn MonoidHom.coeFn #align add_monoid_hom.coe_fn AddMonoidHom.coeFn +#align monoid_hom.coe_fn_apply MonoidHom.coeFn_apply /-- Monoid homomorphism between the function spaces `I → α` and `I → β`, induced by a monoid homomorphism `f` between `α` and `β`. -/ @@ -399,6 +408,7 @@ protected def MonoidHom.compLeft {α β : Type _} [MulOneClass α] [MulOneClass map_mul' _ _ := by ext; simp #align monoid_hom.comp_left MonoidHom.compLeft #align add_monoid_hom.comp_left AddMonoidHom.compLeft +#align monoid_hom.comp_left_apply MonoidHom.compLeft_apply end MonoidHom @@ -464,6 +474,7 @@ def MulHom.single [∀ i, MulZeroClass <| f i] (i : I) : toFun := Pi.single i map_mul' := Pi.single_op₂ (fun _ => (· * ·)) (fun _ => zero_mul _) _ #align mul_hom.single MulHom.single +#align mul_hom.single_apply MulHom.single_apply variable {f} @@ -655,5 +666,6 @@ noncomputable def Function.ExtendByOne.hom [MulOneClass R] : map_mul' f g := by simpa using Function.extend_mul s f g 1 1 #align function.extend_by_one.hom Function.ExtendByOne.hom #align function.extend_by_zero.hom Function.ExtendByZero.hom +#align function.extend_by_one.hom_apply Function.ExtendByOne.hom_apply end Extend diff --git a/Mathlib/Algebra/Group/Prod.lean b/Mathlib/Algebra/Group/Prod.lean index c219fff04dc8d..67ae6f0654eb2 100644 --- a/Mathlib/Algebra/Group/Prod.lean +++ b/Mathlib/Algebra/Group/Prod.lean @@ -774,6 +774,7 @@ def embedProduct (α : Type _) [Monoid α] : map_mul' x y := by simp only [mul_inv_rev, op_mul, Units.val_mul, Prod.mk_mul_mk] #align units.embed_product Units.embedProduct #align add_units.embed_product AddUnits.embedProduct +#align units.embed_product_apply Units.embedProduct_apply @[to_additive] theorem embedProduct_injective (α : Type _) [Monoid α] : Function.Injective (embedProduct α) := @@ -798,6 +799,7 @@ def mulMulHom [CommSemigroup α] : map_mul' _ _ := mul_mul_mul_comm _ _ _ _ #align mul_mul_hom mulMulHom #align add_add_hom addAddHom +#align mul_mul_hom_apply mulMulHom_apply /-- Multiplication as a monoid homomorphism. -/ @[to_additive "Addition as an additive monoid homomorphism.", simps] @@ -805,12 +807,14 @@ def mulMonoidHom [CommMonoid α] : α × α →* α := { mulMulHom with map_one' := mul_one _ } #align mul_monoid_hom mulMonoidHom #align add_add_monoid_hom addAddMonoidHom +#align mul_monoid_hom_apply mulMonoidHom_apply /-- Multiplication as a multiplicative homomorphism with zero. -/ @[simps] def mulMonoidWithZeroHom [CommMonoidWithZero α] : α × α →*₀ α := { mulMonoidHom with map_zero' := mul_zero _ } #align mul_monoid_with_zero_hom mulMonoidWithZeroHom +#align mul_monoid_with_zero_hom_apply mulMonoidWithZeroHom_apply /-- Division as a monoid homomorphism. -/ @[to_additive "Subtraction as an additive monoid homomorphism.", simps] @@ -821,6 +825,7 @@ def divMonoidHom [DivisionCommMonoid α] : map_mul' _ _ := mul_div_mul_comm _ _ _ _ #align div_monoid_hom divMonoidHom #align sub_add_monoid_hom subAddMonoidHom +#align div_monoid_hom_apply divMonoidHom_apply /-- Division as a multiplicative homomorphism with zero. -/ @[simps] @@ -831,5 +836,6 @@ def divMonoidWithZeroHom [CommGroupWithZero α] : map_one' := div_one _ map_mul' _ _ := mul_div_mul_comm _ _ _ _ #align div_monoid_with_zero_hom divMonoidWithZeroHom +#align div_monoid_with_zero_hom_apply divMonoidWithZeroHom_apply end BundledMulDiv diff --git a/Mathlib/Algebra/Group/TypeTags.lean b/Mathlib/Algebra/Group/TypeTags.lean index 61828e21b1aaa..215588aa638fd 100644 --- a/Mathlib/Algebra/Group/TypeTags.lean +++ b/Mathlib/Algebra/Group/TypeTags.lean @@ -416,6 +416,8 @@ def AddMonoidHom.toMultiplicative [AddZeroClass α] [AddZeroClass β] : left_inv _ := rfl right_inv _ := rfl #align add_monoid_hom.to_multiplicative AddMonoidHom.toMultiplicative +#align add_monoid_hom.to_multiplicative_symm_apply_apply AddMonoidHom.toMultiplicative_symm_apply_apply +#align add_monoid_hom.to_multiplicative_apply_apply AddMonoidHom.toMultiplicative_apply_apply /-- Reinterpret `α →* β` as `Additive α →+ Additive β`. -/ @[simps] @@ -434,6 +436,8 @@ def MonoidHom.toAdditive [MulOneClass α] [MulOneClass β] : left_inv _ := rfl right_inv _ := rfl #align monoid_hom.to_additive MonoidHom.toAdditive +#align monoid_hom.to_additive_symm_apply_apply MonoidHom.toAdditive_symm_apply_apply +#align monoid_hom.to_additive_apply_apply MonoidHom.toAdditive_apply_apply /-- Reinterpret `Additive α →+ β` as `α →* Multiplicative β`. -/ @[simps] @@ -452,6 +456,8 @@ def AddMonoidHom.toMultiplicative' [MulOneClass α] [AddZeroClass β] : left_inv _ := rfl right_inv _ := rfl #align add_monoid_hom.to_multiplicative' AddMonoidHom.toMultiplicative' +#align add_monoid_hom.to_multiplicative'_apply_apply AddMonoidHom.toMultiplicative'_apply_apply +#align add_monoid_hom.to_multiplicative'_symm_apply_apply AddMonoidHom.toMultiplicative'_symm_apply_apply /-- Reinterpret `α →* Multiplicative β` as `Additive α →+ β`. -/ @[simps] @@ -459,6 +465,8 @@ def MonoidHom.toAdditive' [MulOneClass α] [AddZeroClass β] : (α →* Multiplicative β) ≃ (Additive α →+ β) := AddMonoidHom.toMultiplicative'.symm #align monoid_hom.to_additive' MonoidHom.toAdditive' +#align monoid_hom.to_additive'_symm_apply_apply MonoidHom.toAdditive'_symm_apply_apply +#align monoid_hom.to_additive'_apply_apply MonoidHom.toAdditive'_apply_apply /-- Reinterpret `α →+ Additive β` as `Multiplicative α →* β`. -/ @[simps] @@ -477,6 +485,8 @@ def AddMonoidHom.toMultiplicative'' [AddZeroClass α] [MulOneClass β] : left_inv _ := rfl right_inv _ := rfl #align add_monoid_hom.to_multiplicative'' AddMonoidHom.toMultiplicative'' +#align add_monoid_hom.to_multiplicative''_symm_apply_apply AddMonoidHom.toMultiplicative''_symm_apply_apply +#align add_monoid_hom.to_multiplicative''_apply_apply AddMonoidHom.toMultiplicative''_apply_apply /-- Reinterpret `Multiplicative α →* β` as `α →+ Additive β`. -/ @[simps] @@ -484,6 +494,8 @@ def MonoidHom.toAdditive'' [AddZeroClass α] [MulOneClass β] : (Multiplicative α →* β) ≃ (α →+ Additive β) := AddMonoidHom.toMultiplicative''.symm #align monoid_hom.to_additive'' MonoidHom.toAdditive'' +#align monoid_hom.to_additive''_symm_apply_apply MonoidHom.toAdditive''_symm_apply_apply +#align monoid_hom.to_additive''_apply_apply MonoidHom.toAdditive''_apply_apply /-- If `α` has some multiplicative structure and coerces to a function, then `Additive α` should also coerce to the same function. diff --git a/Mathlib/Algebra/Group/Units.lean b/Mathlib/Algebra/Group/Units.lean index 5199563be092d..87c0de2c9ba29 100644 --- a/Mathlib/Algebra/Group/Units.lean +++ b/Mathlib/Algebra/Group/Units.lean @@ -310,27 +310,39 @@ theorem inv_mul_cancel_right (a : α) (b : αˣ) : a * ↑b⁻¹ * b = a := by theorem mul_right_inj (a : αˣ) {b c : α} : (a : α) * b = a * c ↔ b = c := ⟨fun h => by simpa only [inv_mul_cancel_left] using congr_arg (fun x : α => ↑(a⁻¹ : αˣ) * x) h, congr_arg _⟩ +#align units.mul_right_inj Units.mul_right_inj +#align add_units.add_right_inj AddUnits.add_right_inj @[to_additive (attr := simp)] theorem mul_left_inj (a : αˣ) {b c : α} : b * a = c * a ↔ b = c := ⟨fun h => by simpa only [mul_inv_cancel_right] using congr_arg (fun x : α => x * ↑(a⁻¹ : αˣ)) h, congr_arg (· * a.val)⟩ +#align units.mul_left_inj Units.mul_left_inj +#align add_units.add_left_inj AddUnits.add_left_inj @[to_additive] theorem eq_mul_inv_iff_mul_eq {a b : α} : a = b * ↑c⁻¹ ↔ a * c = b := ⟨fun h => by rw [h, inv_mul_cancel_right], fun h => by rw [← h, mul_inv_cancel_right]⟩ +#align units.eq_mul_inv_iff_mul_eq Units.eq_mul_inv_iff_mul_eq +#align add_units.eq_add_neg_iff_add_eq AddUnits.eq_add_neg_iff_add_eq @[to_additive] theorem eq_inv_mul_iff_mul_eq {a c : α} : a = ↑b⁻¹ * c ↔ ↑b * a = c := ⟨fun h => by rw [h, mul_inv_cancel_left], fun h => by rw [← h, inv_mul_cancel_left]⟩ +#align units.eq_inv_mul_iff_mul_eq Units.eq_inv_mul_iff_mul_eq +#align add_units.eq_neg_add_iff_add_eq AddUnits.eq_neg_add_iff_add_eq @[to_additive] theorem inv_mul_eq_iff_eq_mul {b c : α} : ↑a⁻¹ * b = c ↔ b = a * c := ⟨fun h => by rw [← h, mul_inv_cancel_left], fun h => by rw [h, inv_mul_cancel_left]⟩ +#align units.inv_mul_eq_iff_eq_mul Units.inv_mul_eq_iff_eq_mul +#align add_units.neg_add_eq_iff_eq_add AddUnits.neg_add_eq_iff_eq_add @[to_additive] theorem mul_inv_eq_iff_eq_mul {a c : α} : a * ↑b⁻¹ = c ↔ a = c * b := ⟨fun h => by rw [← h, inv_mul_cancel_right], fun h => by rw [h, mul_inv_cancel_right]⟩ +#align units.mul_inv_eq_iff_eq_mul Units.mul_inv_eq_iff_eq_mul +#align add_units.add_neg_eq_iff_eq_add AddUnits.add_neg_eq_iff_eq_add -- Porting note: have to explicitly type annotate the 1 @[to_additive] @@ -338,6 +350,8 @@ protected theorem inv_eq_of_mul_eq_one_left {a : α} (h : a * u = 1) : ↑u⁻¹ calc ↑u⁻¹ = (1 : α) * ↑u⁻¹ := by rw [one_mul] _ = a := by rw [← h, mul_inv_cancel_right] +#align units.inv_eq_of_mul_eq_one_left Units.inv_eq_of_mul_eq_one_left +#align add_units.neg_eq_of_add_eq_zero_left AddUnits.neg_eq_of_add_eq_zero_left -- Porting note: have to explicitly type annotate the 1 @@ -346,33 +360,49 @@ protected theorem inv_eq_of_mul_eq_one_right {a : α} (h : ↑u * a = 1) : ↑u calc ↑u⁻¹ = ↑u⁻¹ * (1 : α) := by rw [mul_one] _ = a := by rw [← h, inv_mul_cancel_left] +#align units.inv_eq_of_mul_eq_one_right Units.inv_eq_of_mul_eq_one_right +#align add_units.neg_eq_of_add_eq_zero_right AddUnits.neg_eq_of_add_eq_zero_right @[to_additive] protected theorem eq_inv_of_mul_eq_one_left {a : α} (h : ↑u * a = 1) : a = ↑u⁻¹ := (Units.inv_eq_of_mul_eq_one_right h).symm +#align units.eq_inv_of_mul_eq_one_left Units.eq_inv_of_mul_eq_one_left +#align add_units.eq_neg_of_add_eq_zero_left AddUnits.eq_neg_of_add_eq_zero_left @[to_additive] protected theorem eq_inv_of_mul_eq_one_right {a : α} (h : a * u = 1) : a = ↑u⁻¹ := (Units.inv_eq_of_mul_eq_one_left h).symm +#align units.eq_inv_of_mul_eq_one_right Units.eq_inv_of_mul_eq_one_right +#align add_units.eq_neg_of_add_eq_zero_right AddUnits.eq_neg_of_add_eq_zero_right @[to_additive (attr := simp)] theorem mul_inv_eq_one {a : α} : a * ↑u⁻¹ = 1 ↔ a = u := ⟨inv_inv u ▸ Units.eq_inv_of_mul_eq_one_right, fun h => mul_inv_of_eq h.symm⟩ +#align units.mul_inv_eq_one Units.mul_inv_eq_one +#align add_units.add_neg_eq_zero AddUnits.add_neg_eq_zero @[to_additive (attr := simp)] theorem inv_mul_eq_one {a : α} : ↑u⁻¹ * a = 1 ↔ ↑u = a := ⟨inv_inv u ▸ Units.inv_eq_of_mul_eq_one_right, inv_mul_of_eq⟩ +#align units.inv_mul_eq_one Units.inv_mul_eq_one +#align add_units.neg_add_eq_zero AddUnits.neg_add_eq_zero @[to_additive] theorem mul_eq_one_iff_eq_inv {a : α} : a * u = 1 ↔ a = ↑u⁻¹ := by rw [← mul_inv_eq_one, inv_inv] +#align units.mul_eq_one_iff_eq_inv Units.mul_eq_one_iff_eq_inv +#align add_units.add_eq_zero_iff_eq_neg AddUnits.add_eq_zero_iff_eq_neg @[to_additive] theorem mul_eq_one_iff_inv_eq {a : α} : ↑u * a = 1 ↔ ↑u⁻¹ = a := by rw [← inv_mul_eq_one, inv_inv] +#align units.mul_eq_one_iff_inv_eq Units.mul_eq_one_iff_inv_eq +#align add_units.add_eq_zero_iff_neg_eq AddUnits.add_eq_zero_iff_neg_eq @[to_additive] theorem inv_unique {u₁ u₂ : αˣ} (h : (↑u₁ : α) = ↑u₂) : (↑u₁⁻¹ : α) = ↑u₂⁻¹ := Units.inv_eq_of_mul_eq_one_right <| by rw [h, u₂.mul_inv] +#align units.inv_unique Units.inv_unique +#align add_units.neg_unique AddUnits.neg_unique @[to_additive (attr := simp)] theorem val_inv_eq_inv_val {M : Type _} [DivisionMonoid M] (u : Units M) : ↑u⁻¹ = (u⁻¹ : M) := @@ -386,6 +416,8 @@ end Units "For `a, b` in an `AddCommMonoid` such that `a + b = 0`, makes an add_unit out of `a`."] def Units.mkOfMulEqOne [CommMonoid α] (a b : α) (hab : a * b = 1) : αˣ := ⟨a, b, hab, (mul_comm b a).trans hab⟩ +#align units.mk_of_mul_eq_one Units.mkOfMulEqOne +#align add_units.mk_of_add_eq_zero AddUnits.mkOfAddEqZero @[to_additive (attr := simp)] theorem Units.val_mkOfMulEqOne [CommMonoid α] {a b : α} (h : a * b = 1) : @@ -403,6 +435,7 @@ variable [Monoid α] {a b c : α} in `DivisionRing` it is not totalized at zero. -/ def divp (a : α) (u : Units α) : α := a * (u⁻¹ : αˣ) +#align divp divp @[inherit_doc] infixl:70 " /ₚ " => divp @@ -410,57 +443,71 @@ infixl:70 " /ₚ " => divp @[simp] theorem divp_self (u : αˣ) : (u : α) /ₚ u = 1 := Units.mul_inv _ +#align divp_self divp_self @[simp] theorem divp_one (a : α) : a /ₚ 1 = a := mul_one _ +#align divp_one divp_one theorem divp_assoc (a b : α) (u : αˣ) : a * b /ₚ u = a * (b /ₚ u) := mul_assoc _ _ _ +#align divp_assoc divp_assoc /-- `field_simp` needs the reverse direction of `divp_assoc` to move all `/ₚ` to the right. -/ @[field_simps] theorem divp_assoc' (x y : α) (u : αˣ) : x * (y /ₚ u) = x * y /ₚ u := (divp_assoc _ _ _).symm +#align divp_assoc' divp_assoc' @[simp] theorem divp_inv (u : αˣ) : a /ₚ u⁻¹ = a * u := rfl +#align divp_inv divp_inv @[simp] theorem divp_mul_cancel (a : α) (u : αˣ) : a /ₚ u * u = a := (mul_assoc _ _ _).trans <| by rw [Units.inv_mul, mul_one] +#align divp_mul_cancel divp_mul_cancel @[simp] theorem mul_divp_cancel (a : α) (u : αˣ) : a * u /ₚ u = a := (mul_assoc _ _ _).trans <| by rw [Units.mul_inv, mul_one] +#align mul_divp_cancel mul_divp_cancel @[simp] theorem divp_left_inj (u : αˣ) {a b : α} : a /ₚ u = b /ₚ u ↔ a = b := Units.mul_left_inj _ +#align divp_left_inj divp_left_inj @[field_simps] theorem divp_divp_eq_divp_mul (x : α) (u₁ u₂ : αˣ) : x /ₚ u₁ /ₚ u₂ = x /ₚ (u₂ * u₁) := by simp only [divp, mul_inv_rev, Units.val_mul, mul_assoc] +#align divp_divp_eq_divp_mul divp_divp_eq_divp_mul @[field_simps] theorem divp_eq_iff_mul_eq {x : α} {u : αˣ} {y : α} : x /ₚ u = y ↔ y * u = x := u.mul_left_inj.symm.trans <| by rw [divp_mul_cancel]; exact ⟨Eq.symm, Eq.symm⟩ +#align divp_eq_iff_mul_eq divp_eq_iff_mul_eq @[field_simps] theorem eq_divp_iff_mul_eq {x : α} {u : αˣ} {y : α} : x = y /ₚ u ↔ x * u = y := by rw [eq_comm, divp_eq_iff_mul_eq] +#align eq_divp_iff_mul_eq eq_divp_iff_mul_eq theorem divp_eq_one_iff_eq {a : α} {u : αˣ} : a /ₚ u = 1 ↔ a = u := (Units.mul_left_inj u).symm.trans <| by rw [divp_mul_cancel, one_mul] +#align divp_eq_one_iff_eq divp_eq_one_iff_eq @[simp] theorem one_divp (u : αˣ) : 1 /ₚ u = ↑u⁻¹ := one_mul _ +#align one_divp one_divp /-- Used for `field_simp` to deal with inverses of units. -/ @[field_simps] theorem inv_eq_one_divp (u : αˣ) : ↑u⁻¹ = 1 /ₚ u := by rw [one_divp] +#align inv_eq_one_divp inv_eq_one_divp /-- Used for `field_simp` to deal with inverses of units. This form of the lemma is essential since `field_simp` likes to use `inv_eq_one_div` to rewrite @@ -469,6 +516,7 @@ is essential since `field_simp` likes to use `inv_eq_one_div` to rewrite @[field_simps] theorem inv_eq_one_divp' (u : αˣ) : ((1 / u : αˣ) : α) = 1 /ₚ u := by rw [one_div, one_divp] +#align inv_eq_one_divp' inv_eq_one_divp' /-- `field_simp` moves division inside `αˣ` to the right, and this lemma lifts the calculation to `α`. @@ -487,16 +535,19 @@ variable [CommMonoid α] @[field_simps] theorem divp_mul_eq_mul_divp (x y : α) (u : αˣ) : x /ₚ u * y = x * y /ₚ u := by rw [divp, divp, mul_right_comm] +#align divp_mul_eq_mul_divp divp_mul_eq_mul_divp -- Theoretically redundant as `field_simp` lemma. @[field_simps] theorem divp_eq_divp_iff {x y : α} {ux uy : αˣ} : x /ₚ ux = y /ₚ uy ↔ x * uy = y * ux := by rw [divp_eq_iff_mul_eq, divp_mul_eq_mul_divp, divp_eq_iff_mul_eq] +#align divp_eq_divp_iff divp_eq_divp_iff -- Theoretically redundant as `field_simp` lemma. @[field_simps] theorem divp_mul_divp (x y : α) (ux uy : αˣ) : x /ₚ ux * (y /ₚ uy) = x * y /ₚ (ux * uy) := by rw [divp_mul_eq_mul_divp, divp_assoc', divp_divp_eq_divp_mul] +#align divp_mul_divp divp_mul_divp end CommMonoid @@ -518,6 +569,8 @@ The actual definition says that `a` is equal to some `u : Mˣ`, where where `AddUnits M` is a bundled version of `IsAddUnit`."] def IsUnit [Monoid M] (a : M) : Prop := ∃ u : Mˣ, (u : M) = a +#align is_unit IsUnit +#align is_add_unit IsAddUnit @[to_additive (attr := nontriviality)] theorem isUnit_of_subsingleton [Monoid M] [Subsingleton M] (a : M) : IsUnit a := @@ -558,11 +611,15 @@ theorem isUnit_of_mul_eq_one [CommMonoid M] (a b : M) (h : a * b = 1) : IsUnit a theorem IsUnit.exists_right_inv [Monoid M] {a : M} (h : IsUnit a) : ∃ b, a * b = 1 := by rcases h with ⟨⟨a, b, hab, _⟩, rfl⟩ exact ⟨b, hab⟩ +#align is_unit.exists_right_inv IsUnit.exists_right_inv +#align is_add_unit.exists_neg IsAddUnit.exists_neg @[to_additive IsAddUnit.exists_neg'] theorem IsUnit.exists_left_inv [Monoid M] {a : M} (h : IsUnit a) : ∃ b, b * a = 1 := by rcases h with ⟨⟨a, b, _, hba⟩, rfl⟩ exact ⟨b, hba⟩ +#align is_unit.exists_left_inv IsUnit.exists_left_inv +#align is_add_unit.exists_neg' IsAddUnit.exists_neg' @[to_additive] theorem isUnit_iff_exists_inv [CommMonoid M] {a : M} : IsUnit a ↔ ∃ b, a * b = 1 := @@ -582,6 +639,8 @@ theorem isUnit_iff_exists_inv' [CommMonoid M] {a : M} : IsUnit a ↔ ∃ b, b * theorem IsUnit.mul [Monoid M] {x y : M} : IsUnit x → IsUnit y → IsUnit (x * y) := by rintro ⟨x, rfl⟩ ⟨y, rfl⟩ exact ⟨x * y, Units.val_mul _ _⟩ +#align is_unit.mul IsUnit.mul +#align is_add_unit.add IsAddUnit.add /-- Multiplication by a `u : Mˣ` on the right doesn't affect `IsUnit`. -/ @[to_additive (attr := simp) @@ -659,6 +718,8 @@ theorem unit_of_val_units {a : Mˣ} (h : IsUnit (a : M)) : h.unit = a := @[to_additive (attr := simp)] theorem unit_spec (h : IsUnit a) : ↑h.unit = a := rfl +#align is_unit.unit_spec IsUnit.unit_spec +#align is_add_unit.add_unit_spec IsAddUnit.addUnit_spec @[to_additive (attr := simp)] theorem val_inv_mul (h : IsUnit a) : ↑h.unit⁻¹ * a = 1 := @@ -682,27 +743,39 @@ attribute [instance] IsAddUnit.instDecidableIsAddUnit theorem mul_left_inj (h : IsUnit a) : b * a = c * a ↔ b = c := let ⟨u, hu⟩ := h hu ▸ u.mul_left_inj +#align is_unit.mul_left_inj IsUnit.mul_left_inj +#align is_add_unit.add_left_inj IsAddUnit.add_left_inj @[to_additive] theorem mul_right_inj (h : IsUnit a) : a * b = a * c ↔ b = c := let ⟨u, hu⟩ := h hu ▸ u.mul_right_inj +#align is_unit.mul_right_inj IsUnit.mul_right_inj +#align is_add_unit.add_right_inj IsAddUnit.add_right_inj @[to_additive] protected theorem mul_left_cancel (h : IsUnit a) : a * b = a * c → b = c := h.mul_right_inj.1 +#align is_unit.mul_left_cancel IsUnit.mul_left_cancel +#align is_add_unit.add_left_cancel IsAddUnit.add_left_cancel @[to_additive] protected theorem mul_right_cancel (h : IsUnit b) : a * b = c * b → a = c := h.mul_left_inj.1 +#align is_unit.mul_right_cancel IsUnit.mul_right_cancel +#align is_add_unit.add_right_cancel IsAddUnit.add_right_cancel @[to_additive] protected theorem mul_right_injective (h : IsUnit a) : Injective ((· * ·) a) := fun _ _ => h.mul_left_cancel +#align is_unit.mul_right_injective IsUnit.mul_right_injective +#align is_add_unit.add_right_injective IsAddUnit.add_right_injective @[to_additive] protected theorem mul_left_injective (h : IsUnit b) : Injective (· * b) := fun _ _ => h.mul_right_cancel +#align is_unit.mul_left_injective IsUnit.mul_left_injective +#align is_add_unit.add_left_injective IsAddUnit.add_left_injective end Monoid @@ -712,11 +785,15 @@ variable [DivisionMonoid M] {a : M} protected theorem inv_mul_cancel : IsUnit a → a⁻¹ * a = 1 := by rintro ⟨u, rfl⟩ rw [← Units.val_inv_eq_inv_val, Units.inv_mul] +#align is_unit.inv_mul_cancel IsUnit.inv_mul_cancel +#align is_add_unit.neg_add_cancel IsAddUnit.neg_add_cancel @[to_additive (attr := simp)] protected theorem mul_inv_cancel : IsUnit a → a * a⁻¹ = 1 := by rintro ⟨u, rfl⟩ rw [← Units.val_inv_eq_inv_val, Units.mul_inv] +#align is_unit.mul_inv_cancel IsUnit.mul_inv_cancel +#align is_add_unit.add_neg_cancel IsAddUnit.add_neg_cancel end IsUnit diff --git a/Mathlib/Algebra/Group/WithOne/Basic.lean b/Mathlib/Algebra/Group/WithOne/Basic.lean index ce3eca39f86f0..bdf46f3b47eaa 100644 --- a/Mathlib/Algebra/Group/WithOne/Basic.lean +++ b/Mathlib/Algebra/Group/WithOne/Basic.lean @@ -43,6 +43,7 @@ def coeMulHom [Mul α] : α →ₙ* WithOne α where map_mul' _ _ := rfl #align with_one.coe_mul_hom WithOne.coeMulHom #align with_zero.coe_add_hom WithZero.coeAddHom +#align with_one.coe_mul_hom_apply WithOne.coeMulHom_apply end diff --git a/Mathlib/Algebra/GroupPower/Basic.lean b/Mathlib/Algebra/GroupPower/Basic.lean index f0bbda1a46246..8464fdd05c9fe 100644 --- a/Mathlib/Algebra/GroupPower/Basic.lean +++ b/Mathlib/Algebra/GroupPower/Basic.lean @@ -76,23 +76,28 @@ theorem nsmul_zero (n : ℕ) : n • (0 : A) = 0 := by induction' n with n ih · exact zero_nsmul _ · rw [succ_nsmul, ih, add_zero] +#align nsmul_zero nsmul_zero @[simp] theorem one_nsmul (a : A) : 1 • a = a := by rw [succ_nsmul, zero_nsmul, add_zero] +#align one_nsmul one_nsmul theorem add_nsmul (a : A) (m n : ℕ) : (m + n) • a = m • a + n • a := by induction m with | zero => rw [Nat.zero_add, zero_nsmul, zero_add] | succ m ih => rw [Nat.succ_add, Nat.succ_eq_add_one, succ_nsmul, ih, succ_nsmul, add_assoc] +#align add_nsmul add_nsmul @[to_additive nsmul_zero, simp] theorem one_pow (n : ℕ) : (1 : M) ^ n = 1 := by induction' n with n ih · exact pow_zero _ · rw [pow_succ, ih, one_mul] +#align one_pow one_pow @[to_additive (attr := simp) one_nsmul] theorem pow_one (a : M) : a ^ 1 = a := by rw [pow_succ, pow_zero, mul_one] +#align pow_one pow_one /-- Note that most of the lemmas about powers of two refer to it as `sq`. -/ @[to_additive two_nsmul ""] @@ -101,6 +106,7 @@ theorem pow_two (a : M) : a ^ 2 = a * a := by rw [pow_succ, pow_one] #align two_nsmul two_nsmul alias pow_two ← sq +#align sq sq theorem pow_three' (a : M) : a ^ 3 = a * a * a := by rw [pow_succ', pow_two] #align pow_three' pow_three' @@ -113,6 +119,7 @@ theorem pow_add (a : M) (m n : ℕ) : a ^ (m + n) = a ^ m * a ^ n := by induction' n with n ih · rw [Nat.add_zero, pow_zero, mul_one] · rw [pow_succ', ← mul_assoc, ← ih, ← pow_succ', Nat.add_assoc] +#align pow_add pow_add @[to_additive mul_nsmul] theorem pow_mul (a : M) (m n : ℕ) : a ^ (m * n) = (a ^ m) ^ n := by @@ -264,6 +271,7 @@ def powMonoidHom (n : ℕ) : M →* M where map_mul' a b := mul_pow a b n #align pow_monoid_hom powMonoidHom #align nsmul_add_monoid_hom nsmulAddMonoidHom +#align pow_monoid_hom_apply powMonoidHom_apply end CommMonoid @@ -403,6 +411,7 @@ def zpowGroupHom (n : ℤ) : α →* α where map_mul' a b := mul_zpow a b n #align zpow_group_hom zpowGroupHom #align zsmul_add_group_hom zsmulAddGroupHom +#align zpow_group_hom_apply zpowGroupHom_apply end DivisionCommMonoid diff --git a/Mathlib/Algebra/GroupPower/Identities.lean b/Mathlib/Algebra/GroupPower/Identities.lean index bb00608a8f328..de5060babfafc 100644 --- a/Mathlib/Algebra/GroupPower/Identities.lean +++ b/Mathlib/Algebra/GroupPower/Identities.lean @@ -27,6 +27,7 @@ This sign choice here corresponds to the signs obtained by multiplying two compl theorem sq_add_sq_mul_sq_add_sq : (x₁ ^ 2 + x₂ ^ 2) * (y₁ ^ 2 + y₂ ^ 2) = (x₁ * y₁ - x₂ * y₂) ^ 2 + (x₁ * y₂ + x₂ * y₁) ^ 2 := by ring +#align sq_add_sq_mul_sq_add_sq sq_add_sq_mul_sq_add_sq /-- Brahmagupta's identity, see -/ @@ -34,18 +35,21 @@ theorem sq_add_mul_sq_mul_sq_add_mul_sq : (x₁ ^ 2 + n * x₂ ^ 2) * (y₁ ^ 2 + n * y₂ ^ 2) = (x₁ * y₁ - n * x₂ * y₂) ^ 2 + n * (x₁ * y₂ + x₂ * y₁) ^ 2 := by ring +#align sq_add_mul_sq_mul_sq_add_mul_sq sq_add_mul_sq_mul_sq_add_mul_sq /-- Sophie Germain's identity, see . -/ theorem pow_four_add_four_mul_pow_four : a ^ 4 + 4 * b ^ 4 = ((a - b) ^ 2 + b ^ 2) * ((a + b) ^ 2 + b ^ 2) := by ring +#align pow_four_add_four_mul_pow_four pow_four_add_four_mul_pow_four /-- Sophie Germain's identity, see . -/ theorem pow_four_add_four_mul_pow_four' : a ^ 4 + 4 * b ^ 4 = (a ^ 2 - 2 * a * b + 2 * b ^ 2) * (a ^ 2 + 2 * a * b + 2 * b ^ 2) := by ring +#align pow_four_add_four_mul_pow_four' pow_four_add_four_mul_pow_four' /-- Euler's four-square identity, see . @@ -57,6 +61,7 @@ theorem sum_four_sq_mul_sum_four_sq : (x₁ * y₃ - x₂ * y₄ + x₃ * y₁ + x₄ * y₂) ^ 2 + (x₁ * y₄ + x₂ * y₃ - x₃ * y₂ + x₄ * y₁) ^ 2 := by ring +#align sum_four_sq_mul_sum_four_sq sum_four_sq_mul_sum_four_sq /-- Degen's eight squares identity, see . @@ -74,3 +79,4 @@ theorem sum_eight_sq_mul_sum_eight_sq : (x₁ * y₇ + x₂ * y₈ + x₃ * y₅ - x₄ * y₆ - x₅ * y₃ + x₆ * y₄ + x₇ * y₁ - x₈ * y₂) ^ 2 + (x₁ * y₈ - x₂ * y₇ + x₃ * y₆ + x₄ * y₅ - x₅ * y₄ - x₆ * y₃ + x₇ * y₂ + x₈ * y₁) ^ 2 := by ring +#align sum_eight_sq_mul_sum_eight_sq sum_eight_sq_mul_sum_eight_sq diff --git a/Mathlib/Algebra/GroupPower/Lemmas.lean b/Mathlib/Algebra/GroupPower/Lemmas.lean index a5c1b9b4a8e53..fce1cc389e70a 100644 --- a/Mathlib/Algebra/GroupPower/Lemmas.lean +++ b/Mathlib/Algebra/GroupPower/Lemmas.lean @@ -73,6 +73,7 @@ def Units.ofPow (u : Mˣ) (x : M) {n : ℕ} (hn : n ≠ 0) (hu : x ^ n = u) : M (Commute.self_pow _ _) #align units.of_pow Units.ofPow #align units.of_nsmul AddUnits.ofNSMul +#align add_units.of_nsmul AddUnits.ofNSMul @[to_additive (attr := simp)] theorem isUnit_pow_iff {a : M} {n : ℕ} (hn : n ≠ 0) : IsUnit (a ^ n) ↔ IsUnit a := @@ -737,6 +738,7 @@ lemma natAbs_sq (x : ℤ) : ↑(x.natAbs ^ 2) = x ^ 2 := by rw [sq, Int.natAbs_m #align int.nat_abs_sq Int.natAbs_sq alias natAbs_sq ← natAbs_pow_two +#align int.nat_abs_pow_two Int.natAbs_pow_two theorem natAbs_le_self_sq (a : ℤ) : (Int.natAbs a : ℤ) ≤ a ^ 2 := by rw [← Int.natAbs_sq a, sq] @@ -751,6 +753,7 @@ theorem le_self_sq (b : ℤ) : b ≤ b ^ 2 := #align int.le_self_sq Int.le_self_sq alias le_self_sq ← le_self_pow_two +#align int.le_self_pow_two Int.le_self_pow_two theorem pow_right_injective {x : ℤ} (h : 1 < x.natAbs) : Function.Injective ((· ^ ·) x : ℕ → ℤ) := by diff --git a/Mathlib/Algebra/GroupPower/Order.lean b/Mathlib/Algebra/GroupPower/Order.lean index 22769eb847786..7c14750d2e037 100644 --- a/Mathlib/Algebra/GroupPower/Order.lean +++ b/Mathlib/Algebra/GroupPower/Order.lean @@ -245,6 +245,7 @@ theorem Left.pow_lt_one_iff' [CovariantClass M M (· * ·) (· < ·)] {n : ℕ} x ^ n < 1 ↔ x < 1 := haveI := Mul.to_covariantClass_left M pow_lt_one_iff hn.ne' +#align left.nsmul_neg_iff Left.nsmul_neg_iff theorem Left.pow_lt_one_iff [CovariantClass M M (· * ·) (· < ·)] {n : ℕ} {x : M} (hn : 0 < n) : x ^ n < 1 ↔ x < 1 := Left.pow_lt_one_iff' hn @@ -553,6 +554,7 @@ theorem sq_nonneg (a : R) : 0 ≤ a ^ 2 := #align sq_nonneg sq_nonneg alias sq_nonneg ← pow_two_nonneg +#align pow_two_nonneg pow_two_nonneg theorem pow_bit0_pos {a : R} (h : a ≠ 0) (n : ℕ) : 0 < a ^ bit0 n := (pow_bit0_nonneg a n).lt_of_ne (pow_ne_zero _ h).symm @@ -563,6 +565,7 @@ theorem sq_pos_of_ne_zero (a : R) (h : a ≠ 0) : 0 < a ^ 2 := #align sq_pos_of_ne_zero sq_pos_of_ne_zero alias sq_pos_of_ne_zero ← pow_two_pos_of_ne_zero +#align pow_two_pos_of_ne_zero pow_two_pos_of_ne_zero theorem pow_bit0_pos_iff (a : R) {n : ℕ} (hn : n ≠ 0) : 0 < a ^ bit0 n ↔ a ≠ 0 := by refine' ⟨fun h => _, fun h => pow_bit0_pos h n⟩ @@ -659,6 +662,7 @@ theorem two_mul_le_add_sq (a b : R) : 2 * a * b ≤ a ^ 2 + b ^ 2 := #align two_mul_le_add_sq two_mul_le_add_sq alias two_mul_le_add_sq ← two_mul_le_add_pow_two +#align two_mul_le_add_pow_two two_mul_le_add_pow_two end LinearOrderedCommRing diff --git a/Mathlib/Algebra/GroupPower/Ring.lean b/Mathlib/Algebra/GroupPower/Ring.lean index 6e604163c4e67..04fabe8c630f8 100644 --- a/Mathlib/Algebra/GroupPower/Ring.lean +++ b/Mathlib/Algebra/GroupPower/Ring.lean @@ -34,6 +34,7 @@ variable [MonoidWithZero M] theorem zero_pow : ∀ {n : ℕ}, 0 < n → (0 : M) ^ n = 0 | n + 1, _ => by rw [pow_succ, zero_mul] +#align zero_pow zero_pow @[simp] theorem zero_pow' : ∀ n : ℕ, n ≠ 0 → (0 : M) ^ n = 0 @@ -41,14 +42,17 @@ theorem zero_pow' : ∀ n : ℕ, n ≠ 0 → (0 : M) ^ n = 0 | k + 1, _ => by rw [pow_succ] exact zero_mul _ +#align zero_pow' zero_pow' theorem zero_pow_eq (n : ℕ) : (0 : M) ^ n = if n = 0 then 1 else 0 := by split_ifs with h · rw [h, pow_zero] · rw [zero_pow (Nat.pos_of_ne_zero h)] +#align zero_pow_eq zero_pow_eq theorem pow_eq_zero_of_le {x : M} {n m : ℕ} (hn : n ≤ m) (hx : x ^ n = 0) : x ^ m = 0 := by rw [← tsub_add_cancel_of_le hn, pow_add, hx, mul_zero] +#align pow_eq_zero_of_le pow_eq_zero_of_le theorem pow_eq_zero [NoZeroDivisors M] {x : M} {n : ℕ} (H : x ^ n = 0) : x = 0 := by induction' n with n ih @@ -56,18 +60,22 @@ theorem pow_eq_zero [NoZeroDivisors M] {x : M} {n : ℕ} (H : x ^ n = 0) : x = 0 rw [← mul_one x, H, mul_zero] · rw [pow_succ] at H exact Or.casesOn (mul_eq_zero.1 H) id ih +#align pow_eq_zero pow_eq_zero @[simp] theorem pow_eq_zero_iff [NoZeroDivisors M] {a : M} {n : ℕ} (hn : 0 < n) : a ^ n = 0 ↔ a = 0 := by refine' ⟨pow_eq_zero, _⟩ rintro rfl exact zero_pow hn +#align pow_eq_zero_iff pow_eq_zero_iff theorem pow_eq_zero_iff' [NoZeroDivisors M] [Nontrivial M] {a : M} {n : ℕ} : a ^ n = 0 ↔ a = 0 ∧ n ≠ 0 := by cases (zero_le n).eq_or_gt <;> simp [*, ne_of_gt] +#align pow_eq_zero_iff' pow_eq_zero_iff' theorem pow_ne_zero_iff [NoZeroDivisors M] {a : M} {n : ℕ} (hn : 0 < n) : a ^ n ≠ 0 ↔ a ≠ 0 := (pow_eq_zero_iff hn).not +#align pow_ne_zero_iff pow_ne_zero_iff theorem ne_zero_pow {a : M} {n : ℕ} (hn : n ≠ 0) : a ^ n ≠ 0 → a ≠ 0 := by contrapose! @@ -95,12 +103,14 @@ theorem zero_pow_eq_zero [Nontrivial M] {n : ℕ} : (0 : M) ^ n = 0 ↔ 0 < n := rintro rfl simp at h · exact zero_pow' n h.ne.symm +#align zero_pow_eq_zero zero_pow_eq_zero theorem Ring.inverse_pow (r : M) : ∀ n : ℕ, Ring.inverse r ^ n = Ring.inverse (r ^ n) | 0 => by rw [pow_zero, pow_zero, Ring.inverse_one] | n + 1 => by rw [pow_succ, pow_succ', Ring.mul_inverse_rev' ((Commute.refl r).pow_left n), Ring.inverse_pow r n] +#align ring.inverse_pow Ring.inverse_pow end MonoidWithZero @@ -111,6 +121,7 @@ variable [CommMonoidWithZero M] {n : ℕ} (hn : 0 < n) /-- We define `x ↦ x^n` (for positive `n : ℕ`) as a `MonoidWithZeroHom` -/ def powMonoidWithZeroHom : M →*₀ M := { powMonoidHom n with map_zero' := zero_pow hn } +#align pow_monoid_with_zero_hom powMonoidWithZeroHom @[simp] theorem coe_powMonoidWithZeroHom : (powMonoidWithZeroHom hn : M → M) = fun x ↦ (x^n : M) := rfl @@ -136,6 +147,7 @@ theorem pow_dvd_pow_iff [CancelCommMonoidWithZero R] {x : R} {n m : ℕ} (h0 : x rwa [mul_dvd_mul_iff_left, ← isUnit_iff_dvd_one] at this apply pow_ne_zero m h0 · apply pow_dvd_pow +#align pow_dvd_pow_iff pow_dvd_pow_iff section Semiring @@ -150,6 +162,7 @@ theorem min_pow_dvd_add {n m : ℕ} {a b c : R} (ha : c ^ n ∣ a) (hb : c ^ m replace ha := (pow_dvd_pow c (min_le_left n m)).trans ha replace hb := (pow_dvd_pow c (min_le_right n m)).trans hb exact dvd_add ha hb +#align min_pow_dvd_add min_pow_dvd_add end Semiring @@ -166,6 +179,7 @@ theorem add_sq' (a b : R) : (a + b) ^ 2 = a ^ 2 + b ^ 2 + 2 * a * b := by #align add_sq' add_sq' alias add_sq ← add_pow_two +#align add_pow_two add_pow_two end CommSemiring @@ -180,11 +194,13 @@ theorem neg_one_pow_eq_or : ∀ n : ℕ, (-1 : R) ^ n = 1 ∨ (-1 : R) ^ n = -1 | n + 1 => (neg_one_pow_eq_or n).symm.imp (fun h ↦ by rw [pow_succ, h, neg_one_mul, neg_neg]) (fun h ↦ by rw [pow_succ, h, mul_one]) +#align neg_one_pow_eq_or neg_one_pow_eq_or variable {R} theorem neg_pow (a : R) (n : ℕ) : (-a) ^ n = (-1) ^ n * a ^ n := neg_one_mul a ▸ (Commute.neg_one_left a).mul_pow n +#align neg_pow neg_pow section set_option linter.deprecated false @@ -203,13 +219,17 @@ end @[simp] theorem neg_sq (a : R) : (-a) ^ 2 = a ^ 2 := by simp [sq] +#align neg_sq neg_sq -- Porting note: removed the simp attribute to please the simpNF linter theorem neg_one_sq : (-1 : R) ^ 2 = 1 := by simp [neg_sq, one_pow] +#align neg_one_sq neg_one_sq alias neg_sq ← neg_pow_two +#align neg_pow_two neg_pow_two alias neg_one_sq ← neg_one_pow_two +#align neg_one_pow_two neg_one_pow_two end HasDistribNeg @@ -219,14 +239,17 @@ variable [Ring R] {a b : R} protected theorem Commute.sq_sub_sq (h : Commute a b) : a ^ 2 - b ^ 2 = (a + b) * (a - b) := by rw [sq, sq, h.mul_self_sub_mul_self_eq] +#align commute.sq_sub_sq Commute.sq_sub_sq @[simp] theorem neg_one_pow_mul_eq_zero_iff {n : ℕ} {r : R} : (-1) ^ n * r = 0 ↔ r = 0 := by rcases neg_one_pow_eq_or R n with h | h <;> simp [h] +#align neg_one_pow_mul_eq_zero_iff neg_one_pow_mul_eq_zero_iff @[simp] theorem mul_neg_one_pow_eq_zero_iff {n : ℕ} {r : R} : r * (-1) ^ n = 0 ↔ r = 0 := by rcases neg_one_pow_eq_or R n with h | h <;> simp [h] +#align mul_neg_one_pow_eq_zero_iff mul_neg_one_pow_eq_zero_iff variable [NoZeroDivisors R] @@ -252,33 +275,42 @@ variable [CommRing R] theorem sq_sub_sq (a b : R) : a ^ 2 - b ^ 2 = (a + b) * (a - b) := (Commute.all a b).sq_sub_sq +#align sq_sub_sq sq_sub_sq alias sq_sub_sq ← pow_two_sub_pow_two +#align pow_two_sub_pow_two pow_two_sub_pow_two theorem sub_sq (a b : R) : (a - b) ^ 2 = a ^ 2 - 2 * a * b + b ^ 2 := by rw [sub_eq_add_neg, add_sq, neg_sq, mul_neg, ← sub_eq_add_neg] +#align sub_sq sub_sq alias sub_sq ← sub_pow_two +#align sub_pow_two sub_pow_two theorem sub_sq' (a b : R) : (a - b) ^ 2 = a ^ 2 + b ^ 2 - 2 * a * b := by rw [sub_eq_add_neg, add_sq', neg_sq, mul_neg, ← sub_eq_add_neg] +#align sub_sq' sub_sq' variable [NoZeroDivisors R] {a b : R} theorem sq_eq_sq_iff_eq_or_eq_neg : a ^ 2 = b ^ 2 ↔ a = b ∨ a = -b := (Commute.all a b).sq_eq_sq_iff_eq_or_eq_neg +#align sq_eq_sq_iff_eq_or_eq_neg sq_eq_sq_iff_eq_or_eq_neg theorem eq_or_eq_neg_of_sq_eq_sq (a b : R) : a ^ 2 = b ^ 2 → a = b ∨ a = -b := sq_eq_sq_iff_eq_or_eq_neg.1 +#align eq_or_eq_neg_of_sq_eq_sq eq_or_eq_neg_of_sq_eq_sq -- Copies of the above CommRing lemmas for `Units R`. namespace Units protected theorem sq_eq_sq_iff_eq_or_eq_neg {a b : Rˣ} : a ^ 2 = b ^ 2 ↔ a = b ∨ a = -b := by simp_rw [ext_iff, val_pow_eq_pow_val, sq_eq_sq_iff_eq_or_eq_neg, Units.val_neg] +#align units.sq_eq_sq_iff_eq_or_eq_neg Units.sq_eq_sq_iff_eq_or_eq_neg protected theorem eq_or_eq_neg_of_sq_eq_sq (a b : Rˣ) (h : a ^ 2 = b ^ 2) : a = b ∨ a = -b := Units.sq_eq_sq_iff_eq_or_eq_neg.1 h +#align units.eq_or_eq_neg_of_sq_eq_sq Units.eq_or_eq_neg_of_sq_eq_sq end Units diff --git a/Mathlib/Algebra/GroupRingAction/Basic.lean b/Mathlib/Algebra/GroupRingAction/Basic.lean index 8e33a0607b624..cf4934455cd8b 100644 --- a/Mathlib/Algebra/GroupRingAction/Basic.lean +++ b/Mathlib/Algebra/GroupRingAction/Basic.lean @@ -62,6 +62,7 @@ instance (priority := 100) MulSemiringAction.toMulDistribMulAction [h : MulSemir def MulSemiringAction.toRingHom [MulSemiringAction M R] (x : M) : R →+* R := { MulDistribMulAction.toMonoidHom R x, DistribMulAction.toAddMonoidHom R x with } #align mul_semiring_action.to_ring_hom MulSemiringAction.toRingHom +#align mul_semiring_action.to_ring_hom_apply MulSemiringAction.toRingHom_apply theorem toRingHom_injective [MulSemiringAction M R] [FaithfulSMul M R] : Function.Injective (MulSemiringAction.toRingHom M R) := fun _ _ h => @@ -73,6 +74,8 @@ theorem toRingHom_injective [MulSemiringAction M R] [FaithfulSMul M R] : def MulSemiringAction.toRingEquiv [MulSemiringAction G R] (x : G) : R ≃+* R := { DistribMulAction.toAddEquiv R x, MulSemiringAction.toRingHom G R x with } #align mul_semiring_action.to_ring_equiv MulSemiringAction.toRingEquiv +#align mul_semiring_action.to_ring_equiv_symm_apply MulSemiringAction.toRingEquiv_symmApply +#align mul_semiring_action.to_ring_equiv_apply MulSemiringAction.toRingEquiv_apply section diff --git a/Mathlib/Algebra/GroupWithZero/Basic.lean b/Mathlib/Algebra/GroupWithZero/Basic.lean index 5fe4822cf8967..eab34f684f568 100644 --- a/Mathlib/Algebra/GroupWithZero/Basic.lean +++ b/Mathlib/Algebra/GroupWithZero/Basic.lean @@ -55,23 +55,29 @@ variable [MulZeroClass M₀] {a b : M₀} theorem left_ne_zero_of_mul : a * b ≠ 0 → a ≠ 0 := mt fun h => mul_eq_zero_of_left h b +#align left_ne_zero_of_mul left_ne_zero_of_mul theorem right_ne_zero_of_mul : a * b ≠ 0 → b ≠ 0 := mt (mul_eq_zero_of_right a) +#align right_ne_zero_of_mul right_ne_zero_of_mul theorem ne_zero_and_ne_zero_of_mul (h : a * b ≠ 0) : a ≠ 0 ∧ b ≠ 0 := ⟨left_ne_zero_of_mul h, right_ne_zero_of_mul h⟩ +#align ne_zero_and_ne_zero_of_mul ne_zero_and_ne_zero_of_mul theorem mul_eq_zero_of_ne_zero_imp_eq_zero {a b : M₀} (h : a ≠ 0 → b = 0) : a * b = 0 := if ha : a = 0 then by rw [ha, zero_mul] else by rw [h ha, mul_zero] +#align mul_eq_zero_of_ne_zero_imp_eq_zero mul_eq_zero_of_ne_zero_imp_eq_zero /-- To match `one_mul_eq_id`. -/ theorem zero_mul_eq_const : (· * ·) (0 : M₀) = Function.const _ 0 := funext zero_mul +#align zero_mul_eq_const zero_mul_eq_const /-- To match `mul_one_eq_id`. -/ theorem mul_zero_eq_const : (· * (0 : M₀)) = Function.const _ 0 := funext mul_zero +#align mul_zero_eq_const mul_zero_eq_const end MulZeroClass @@ -81,10 +87,12 @@ variable [Mul M₀] [Zero M₀] [NoZeroDivisors M₀] {a b : M₀} theorem eq_zero_of_mul_self_eq_zero (h : a * a = 0) : a = 0 := (eq_zero_or_eq_zero_of_mul_eq_zero h).elim id id +#align eq_zero_of_mul_self_eq_zero eq_zero_of_mul_self_eq_zero @[field_simps] theorem mul_ne_zero (ha : a ≠ 0) (hb : b ≠ 0) : a * b ≠ 0 := mt eq_zero_or_eq_zero_of_mul_eq_zero <| not_or.mpr ⟨ha, hb⟩ +#align mul_ne_zero mul_ne_zero end Mul @@ -105,6 +113,7 @@ variable [MulZeroOneClass M₀] /-- In a monoid with zero, if zero equals one, then zero is the only element. -/ theorem eq_zero_of_zero_eq_one (h : (0 : M₀) = 1) (a : M₀) : a = 0 := by rw [← mul_one a, ← h, mul_zero] +#align eq_zero_of_zero_eq_one eq_zero_of_zero_eq_one /-- In a monoid with zero, if zero equals one, then zero is the unique element. @@ -119,15 +128,19 @@ def uniqueOfZeroEqOne (h : (0 : M₀) = 1) : Unique M₀ where are equal. -/ theorem subsingleton_iff_zero_eq_one : (0 : M₀) = 1 ↔ Subsingleton M₀ := ⟨fun h => haveI := uniqueOfZeroEqOne h; inferInstance, fun h => @Subsingleton.elim _ h _ _⟩ +#align subsingleton_iff_zero_eq_one subsingleton_iff_zero_eq_one alias subsingleton_iff_zero_eq_one ↔ subsingleton_of_zero_eq_one _ +#align subsingleton_of_zero_eq_one subsingleton_of_zero_eq_one theorem eq_of_zero_eq_one (h : (0 : M₀) = 1) (a b : M₀) : a = b := @Subsingleton.elim _ (subsingleton_of_zero_eq_one h) a b +#align eq_of_zero_eq_one eq_of_zero_eq_one /-- In a monoid with zero, either zero and one are nonequal, or zero is the only element. -/ theorem zero_ne_one_or_forall_eq_0 : (0 : M₀) ≠ 1 ∨ ∀ a : M₀, a = 0 := not_or_of_imp eq_zero_of_zero_eq_one +#align zero_ne_one_or_forall_eq_0 zero_ne_one_or_forall_eq_0 end @@ -137,9 +150,11 @@ variable [MulZeroOneClass M₀] [Nontrivial M₀] {a b : M₀} theorem left_ne_zero_of_mul_eq_one (h : a * b = 1) : a ≠ 0 := left_ne_zero_of_mul <| ne_zero_of_eq_one h +#align left_ne_zero_of_mul_eq_one left_ne_zero_of_mul_eq_one theorem right_ne_zero_of_mul_eq_one (h : a * b = 1) : b ≠ 0 := right_ne_zero_of_mul <| ne_zero_of_eq_one h +#align right_ne_zero_of_mul_eq_one right_ne_zero_of_mul_eq_one end @@ -155,39 +170,47 @@ instance (priority := 10) CancelMonoidWithZero.to_noZeroDivisors : NoZeroDivisor theorem mul_left_inj' (hc : c ≠ 0) : a * c = b * c ↔ a = b := (mul_left_injective₀ hc).eq_iff +#align mul_left_inj' mul_left_inj' theorem mul_right_inj' (ha : a ≠ 0) : a * b = a * c ↔ b = c := (mul_right_injective₀ ha).eq_iff +#align mul_right_inj' mul_right_inj' @[simp] theorem mul_eq_mul_right_iff : a * c = b * c ↔ a = b ∨ c = 0 := by by_cases hc : c = 0 <;> [simp [hc], simp [mul_left_inj', hc]] +#align mul_eq_mul_right_iff mul_eq_mul_right_iff @[simp] theorem mul_eq_mul_left_iff : a * b = a * c ↔ b = c ∨ a = 0 := by by_cases ha : a = 0 <;> [simp [ha], simp [mul_right_inj', ha]] +#align mul_eq_mul_left_iff mul_eq_mul_left_iff theorem mul_right_eq_self₀ : a * b = a ↔ b = 1 ∨ a = 0 := calc a * b = a ↔ a * b = a * 1 := by rw [mul_one] _ ↔ b = 1 ∨ a = 0 := mul_eq_mul_left_iff +#align mul_right_eq_self₀ mul_right_eq_self₀ theorem mul_left_eq_self₀ : a * b = b ↔ a = 1 ∨ b = 0 := calc a * b = b ↔ a * b = 1 * b := by rw [one_mul] _ ↔ a = 1 ∨ b = 0 := mul_eq_mul_right_iff +#align mul_left_eq_self₀ mul_left_eq_self₀ /-- An element of a `CancelMonoidWithZero` fixed by right multiplication by an element other than one must be zero. -/ theorem eq_zero_of_mul_eq_self_right (h₁ : b ≠ 1) (h₂ : a * b = a) : a = 0 := Classical.byContradiction fun ha => h₁ <| mul_left_cancel₀ ha <| h₂.symm ▸ (mul_one a).symm +#align eq_zero_of_mul_eq_self_right eq_zero_of_mul_eq_self_right /-- An element of a `CancelMonoidWithZero` fixed by left multiplication by an element other than one must be zero. -/ theorem eq_zero_of_mul_eq_self_left (h₁ : b ≠ 1) (h₂ : b * a = a) : a = 0 := Classical.byContradiction fun ha => h₁ <| mul_right_cancel₀ ha <| h₂.symm ▸ (one_mul a).symm +#align eq_zero_of_mul_eq_self_left eq_zero_of_mul_eq_self_left end CancelMonoidWithZero @@ -200,6 +223,7 @@ theorem mul_inv_cancel_right₀ (h : b ≠ 0) (a : G₀) : a * b * b⁻¹ = a := calc a * b * b⁻¹ = a * (b * b⁻¹) := mul_assoc _ _ _ _ = a := by simp [h] +#align mul_inv_cancel_right₀ mul_inv_cancel_right₀ @[simp] @@ -207,12 +231,14 @@ theorem mul_inv_cancel_left₀ (h : a ≠ 0) (b : G₀) : a * (a⁻¹ * b) = b : calc a * (a⁻¹ * b) = a * a⁻¹ * b := (mul_assoc _ _ _).symm _ = b := by simp [h] +#align mul_inv_cancel_left₀ mul_inv_cancel_left₀ -- Porting note: used `simpa` to prove `False` in lean3 theorem inv_ne_zero (h : a ≠ 0) : a⁻¹ ≠ 0 := fun a_eq_0 => by have := mul_inv_cancel h simp [a_eq_0] at this +#align inv_ne_zero inv_ne_zero @[simp] theorem inv_mul_cancel (h : a ≠ 0) : a⁻¹ * a = 1 := @@ -220,21 +246,25 @@ theorem inv_mul_cancel (h : a ≠ 0) : a⁻¹ * a = 1 := a⁻¹ * a = a⁻¹ * a * a⁻¹ * a⁻¹⁻¹ := by simp [inv_ne_zero h] _ = a⁻¹ * a⁻¹⁻¹ := by simp [h] _ = 1 := by simp [inv_ne_zero h] +#align inv_mul_cancel inv_mul_cancel theorem GroupWithZero.mul_left_injective (h : x ≠ 0) : Function.Injective fun y => x * y := fun y y' w => by simpa only [← mul_assoc, inv_mul_cancel h, one_mul] using congr_arg (fun y => x⁻¹ * y) w +#align group_with_zero.mul_left_injective GroupWithZero.mul_left_injective theorem GroupWithZero.mul_right_injective (h : x ≠ 0) : Function.Injective fun y => y * x := fun y y' w => by simpa only [mul_assoc, mul_inv_cancel _ h, mul_one] using congr_arg (fun y => y * x⁻¹) w +#align group_with_zero.mul_right_injective GroupWithZero.mul_right_injective @[simp] theorem inv_mul_cancel_right₀ (h : b ≠ 0) (a : G₀) : a * b⁻¹ * b = a := calc a * b⁻¹ * b = a * (b⁻¹ * b) := mul_assoc _ _ _ _ = a := by simp [h] +#align inv_mul_cancel_right₀ inv_mul_cancel_right₀ @[simp] @@ -242,6 +272,7 @@ theorem inv_mul_cancel_left₀ (h : a ≠ 0) (b : G₀) : a⁻¹ * (a * b) = b : calc a⁻¹ * (a * b) = a⁻¹ * a * b := (mul_assoc _ _ _).symm _ = b := by simp [h] +#align inv_mul_cancel_left₀ inv_mul_cancel_left₀ private theorem inv_eq_of_mul (h : a * b = 1) : a⁻¹ = b := by @@ -277,9 +308,11 @@ variable [GroupWithZero G₀] {a b c : G₀} @[simp] theorem zero_div (a : G₀) : 0 / a = 0 := by rw [div_eq_mul_inv, zero_mul] +#align zero_div zero_div @[simp] theorem div_zero (a : G₀) : a / 0 = 0 := by rw [div_eq_mul_inv, inv_zero, mul_zero] +#align div_zero div_zero /-- Multiplying `a` by itself and then by its inverse results in `a` (whether or not `a` is zero). -/ @@ -287,8 +320,8 @@ theorem div_zero (a : G₀) : a / 0 = 0 := by rw [div_eq_mul_inv, inv_zero, mul_ theorem mul_self_mul_inv (a : G₀) : a * a * a⁻¹ = a := by by_cases h : a = 0 · rw [h, inv_zero, mul_zero] - · rw [mul_assoc, mul_inv_cancel h, mul_one] +#align mul_self_mul_inv mul_self_mul_inv /-- Multiplying `a` by its inverse and then by itself results in `a` @@ -297,8 +330,8 @@ theorem mul_self_mul_inv (a : G₀) : a * a * a⁻¹ = a := by theorem mul_inv_mul_self (a : G₀) : a * a⁻¹ * a = a := by by_cases h : a = 0 · rw [h, inv_zero, mul_zero] - · rw [mul_inv_cancel h, one_mul] +#align mul_inv_mul_self mul_inv_mul_self /-- Multiplying `a⁻¹` by `a` twice results in `a` (whether or not `a` @@ -307,19 +340,21 @@ is zero). -/ theorem inv_mul_mul_self (a : G₀) : a⁻¹ * a * a = a := by by_cases h : a = 0 · rw [h, inv_zero, mul_zero] - · rw [inv_mul_cancel h, one_mul] +#align inv_mul_mul_self inv_mul_mul_self /-- Multiplying `a` by itself and then dividing by itself results in `a`, whether or not `a` is zero. -/ @[simp] theorem mul_self_div_self (a : G₀) : a * a / a = a := by rw [div_eq_mul_inv, mul_self_mul_inv a] +#align mul_self_div_self mul_self_div_self /-- Dividing `a` by itself and then multiplying by itself results in `a`, whether or not `a` is zero. -/ @[simp] theorem div_self_mul_self (a : G₀) : a / a * a = a := by rw [div_eq_mul_inv, mul_inv_mul_self a] +#align div_self_mul_self div_self_mul_self attribute [local simp] div_eq_mul_inv mul_comm mul_assoc mul_left_comm @@ -328,17 +363,21 @@ theorem div_self_mul_self' (a : G₀) : a / (a * a) = a⁻¹ := calc a / (a * a) = a⁻¹⁻¹ * a⁻¹ * a⁻¹ := by simp [mul_inv_rev] _ = a⁻¹ := inv_mul_mul_self _ +#align div_self_mul_self' div_self_mul_self' theorem one_div_ne_zero {a : G₀} (h : a ≠ 0) : 1 / a ≠ 0 := by simpa only [one_div] using inv_ne_zero h +#align one_div_ne_zero one_div_ne_zero @[simp] theorem inv_eq_zero {a : G₀} : a⁻¹ = 0 ↔ a = 0 := by rw [inv_eq_iff_inv_eq, inv_zero, eq_comm] +#align inv_eq_zero inv_eq_zero @[simp] theorem zero_eq_inv {a : G₀} : 0 = a⁻¹ ↔ 0 = a := eq_comm.trans <| inv_eq_zero.trans eq_comm +#align zero_eq_inv zero_eq_inv /-- Dividing `a` by the result of dividing `a` by itself results in `a` (whether or not `a` is zero). -/ @@ -346,19 +385,24 @@ theorem zero_eq_inv {a : G₀} : 0 = a⁻¹ ↔ 0 = a := theorem div_div_self (a : G₀) : a / (a / a) = a := by rw [div_div_eq_mul_div] exact mul_self_div_self a +#align div_div_self div_div_self theorem ne_zero_of_one_div_ne_zero {a : G₀} (h : 1 / a ≠ 0) : a ≠ 0 := fun ha : a = 0 => by rw [ha, div_zero] at h contradiction +#align ne_zero_of_one_div_ne_zero ne_zero_of_one_div_ne_zero theorem eq_zero_of_one_div_eq_zero {a : G₀} (h : 1 / a = 0) : a = 0 := Classical.byCases (fun ha => ha) fun ha => ((one_div_ne_zero ha) h).elim +#align eq_zero_of_one_div_eq_zero eq_zero_of_one_div_eq_zero theorem mul_left_surjective₀ {a : G₀} (h : a ≠ 0) : Surjective fun g => a * g := fun g => ⟨a⁻¹ * g, by simp [← mul_assoc, mul_inv_cancel h]⟩ +#align mul_left_surjective₀ mul_left_surjective₀ theorem mul_right_surjective₀ {a : G₀} (h : a ≠ 0) : Surjective fun g => g * a := fun g => ⟨g * a⁻¹, by simp [mul_assoc, inv_mul_cancel h]⟩ +#align mul_right_surjective₀ mul_right_surjective₀ end GroupWithZero @@ -368,6 +412,7 @@ variable [CommGroupWithZero G₀] {a b c d : G₀} theorem div_mul_eq_mul_div₀ (a b c : G₀) : a / c * b = a * b / c := by simp_rw [div_eq_mul_inv, mul_assoc, mul_comm c⁻¹] +#align div_mul_eq_mul_div₀ div_mul_eq_mul_div₀ end CommGroupWithZero diff --git a/Mathlib/Algebra/GroupWithZero/Defs.lean b/Mathlib/Algebra/GroupWithZero/Defs.lean index de0ca7106eb0f..0901545d7d8c3 100644 --- a/Mathlib/Algebra/GroupWithZero/Defs.lean +++ b/Mathlib/Algebra/GroupWithZero/Defs.lean @@ -50,11 +50,13 @@ class MulZeroClass (M₀ : Type u) extends Mul M₀, Zero M₀ where zero_mul : ∀ a : M₀, 0 * a = 0 /-- Zero is a right absorbing element for multiplication -/ mul_zero : ∀ a : M₀, a * 0 = 0 +#align mul_zero_class MulZeroClass /-- A mixin for left cancellative multiplication by nonzero elements. -/ class IsLeftCancelMulZero (M₀ : Type u) [Mul M₀] [Zero M₀] : Prop where /-- Multiplicatin by a nonzero element is left cancellative. -/ protected mul_left_cancel_of_ne_zero : ∀ {a b c : M₀}, a ≠ 0 → a * b = a * c → b = c +#align is_left_cancel_mul_zero IsLeftCancelMulZero section IsLeftCancelMulZero @@ -62,9 +64,11 @@ variable [Mul M₀] [Zero M₀] [IsLeftCancelMulZero M₀] {a b c : M₀} theorem mul_left_cancel₀ (ha : a ≠ 0) (h : a * b = a * c) : b = c := IsLeftCancelMulZero.mul_left_cancel_of_ne_zero ha h +#align mul_left_cancel₀ mul_left_cancel₀ theorem mul_right_injective₀ (ha : a ≠ 0) : Function.Injective ((· * ·) a) := fun _ _ => mul_left_cancel₀ ha +#align mul_right_injective₀ mul_right_injective₀ end IsLeftCancelMulZero @@ -72,6 +76,7 @@ end IsLeftCancelMulZero class IsRightCancelMulZero (M₀ : Type u) [Mul M₀] [Zero M₀] : Prop where /-- Multiplicatin by a nonzero element is right cancellative. -/ protected mul_right_cancel_of_ne_zero : ∀ {a b c : M₀}, b ≠ 0 → a * b = c * b → a = c +#align is_right_cancel_mul_zero IsRightCancelMulZero section IsRightCancelMulZero @@ -79,15 +84,18 @@ variable [Mul M₀] [Zero M₀] [IsRightCancelMulZero M₀] {a b c : M₀} theorem mul_right_cancel₀ (hb : b ≠ 0) (h : a * b = c * b) : a = c := IsRightCancelMulZero.mul_right_cancel_of_ne_zero hb h +#align mul_right_cancel₀ mul_right_cancel₀ theorem mul_left_injective₀ (hb : b ≠ 0) : Function.Injective fun a => a * b := fun _ _ => mul_right_cancel₀ hb +#align mul_left_injective₀ mul_left_injective₀ end IsRightCancelMulZero /-- A mixin for cancellative multiplication by nonzero elements. -/ class IsCancelMulZero (M₀ : Type u) [Mul M₀] [Zero M₀] extends IsLeftCancelMulZero M₀, IsRightCancelMulZero M₀ : Prop +#align is_cancel_mul_zero IsCancelMulZero export MulZeroClass (zero_mul mul_zero) attribute [simp] zero_mul mul_zero @@ -97,26 +105,32 @@ for all `a` and `b` of type `G₀`. -/ class NoZeroDivisors (M₀ : Type _) [Mul M₀] [Zero M₀] : Prop where /-- For all `a` and `b` of `G₀`, `a * b = 0` implies `a = 0` or `b = 0`. -/ eq_zero_or_eq_zero_of_mul_eq_zero : ∀ {a b : M₀}, a * b = 0 → a = 0 ∨ b = 0 +#align no_zero_divisors NoZeroDivisors export NoZeroDivisors (eq_zero_or_eq_zero_of_mul_eq_zero) /-- A type `S₀` is a "semigroup with zero” if it is a semigroup with zero element, and `0` is left and right absorbing. -/ class SemigroupWithZero (S₀ : Type u) extends Semigroup S₀, MulZeroClass S₀ +#align semigroup_with_zero SemigroupWithZero /-- A typeclass for non-associative monoids with zero elements. -/ class MulZeroOneClass (M₀ : Type u) extends MulOneClass M₀, MulZeroClass M₀ +#align mul_zero_one_class MulZeroOneClass /-- A type `M₀` is a “monoid with zero” if it is a monoid with zero element, and `0` is left and right absorbing. -/ class MonoidWithZero (M₀ : Type u) extends Monoid M₀, MulZeroOneClass M₀, SemigroupWithZero M₀ +#align monoid_with_zero MonoidWithZero /-- A type `M` is a `CancelMonoidWithZero` if it is a monoid with zero element, `0` is left and right absorbing, and left/right multiplication by a non-zero element is injective. -/ class CancelMonoidWithZero (M₀ : Type _) extends MonoidWithZero M₀, IsCancelMulZero M₀ +#align cancel_monoid_with_zero CancelMonoidWithZero /-- A type `M` is a commutative “monoid with zero” if it is a commutative monoid with zero element, and `0` is left and right absorbing. -/ class CommMonoidWithZero (M₀ : Type _) extends CommMonoid M₀, MonoidWithZero M₀ +#align comm_monoid_with_zero CommMonoidWithZero section CommSemigroup @@ -126,19 +140,23 @@ lemma IsLeftCancelMulZero.to_isRightCancelMulZero [IsLeftCancelMulZero M₀] : IsRightCancelMulZero M₀ := { mul_right_cancel_of_ne_zero := fun hb h => mul_left_cancel₀ hb <| (mul_comm _ _).trans (h.trans (mul_comm _ _)) } +#align is_left_cancel_mul_zero.to_is_right_cancel_mul_zero IsLeftCancelMulZero.to_isRightCancelMulZero lemma IsRightCancelMulZero.to_isLeftCancelMulZero [IsRightCancelMulZero M₀] : IsLeftCancelMulZero M₀ := { mul_left_cancel_of_ne_zero := fun hb h => mul_right_cancel₀ hb <| (mul_comm _ _).trans (h.trans (mul_comm _ _)) } +#align is_right_cancel_mul_zero.to_is_left_cancel_mul_zero IsRightCancelMulZero.to_isLeftCancelMulZero lemma IsLeftCancelMulZero.to_isCancelMulZero [IsLeftCancelMulZero M₀] : IsCancelMulZero M₀ := { IsLeftCancelMulZero.to_isRightCancelMulZero with } +#align is_left_cancel_mul_zero.to_is_cancel_mul_zero IsLeftCancelMulZero.to_isCancelMulZero lemma IsRightCancelMulZero.to_isCancelMulZero [IsRightCancelMulZero M₀] : IsCancelMulZero M₀ := { IsRightCancelMulZero.to_isLeftCancelMulZero with } +#align is_right_cancel_mul_zero.to_is_cancel_mul_zero IsRightCancelMulZero.to_isCancelMulZero end CommSemigroup @@ -146,6 +164,7 @@ end CommSemigroup `0` is left and right absorbing, and left/right multiplication by a non-zero element is injective. -/ class CancelCommMonoidWithZero (M₀ : Type _) extends CommMonoidWithZero M₀, IsLeftCancelMulZero M₀ +#align cancel_comm_monoid_with_zero CancelCommMonoidWithZero instance (priority := 100) CancelCommMonoidWithZero.toCancelMonoidWithZero [CancelCommMonoidWithZero M₀] : CancelMonoidWithZero M₀ := @@ -162,18 +181,21 @@ class GroupWithZero (G₀ : Type u) extends MonoidWithZero G₀, DivInvMonoid G inv_zero : (0 : G₀)⁻¹ = 0 /-- Every nonzero element of a group with zero is invertible. -/ mul_inv_cancel (a : G₀) : a ≠ 0 → a * a⁻¹ = 1 +#align group_with_zero GroupWithZero export GroupWithZero (inv_zero) attribute [simp] inv_zero @[simp] lemma mul_inv_cancel [GroupWithZero G₀] {a : G₀} (h : a ≠ 0) : a * a⁻¹ = 1 := GroupWithZero.mul_inv_cancel a h +#align mul_inv_cancel mul_inv_cancel /-- A type `G₀` is a commutative “group with zero” if it is a commutative monoid with zero element (distinct from `1`) such that every nonzero element is invertible. The type is required to come with an “inverse” function, and the inverse of `0` must be `0`. -/ class CommGroupWithZero (G₀ : Type _) extends CommMonoidWithZero G₀, GroupWithZero G₀ +#align comm_group_with_zero CommGroupWithZero section NeZero @@ -203,6 +225,7 @@ theorem pullback_nonzero [Zero M₀'] [One M₀'] (f : M₀' → M₀) (zero : f ⟨⟨0, 1, mt (congr_arg f) <| by rw [zero, one] exact zero_ne_one⟩⟩ +#align pullback_nonzero pullback_nonzero end NeZero @@ -211,8 +234,10 @@ section MulZeroClass variable [MulZeroClass M₀] theorem mul_eq_zero_of_left {a : M₀} (h : a = 0) (b : M₀) : a * b = 0 := h.symm ▸ zero_mul b +#align mul_eq_zero_of_left mul_eq_zero_of_left theorem mul_eq_zero_of_right (a : M₀) {b : M₀} (h : b = 0) : a * b = 0 := h.symm ▸ mul_zero a +#align mul_eq_zero_of_right mul_eq_zero_of_right variable [NoZeroDivisors M₀] {a b : M₀} @@ -222,31 +247,40 @@ equals zero. -/ theorem mul_eq_zero : a * b = 0 ↔ a = 0 ∨ b = 0 := ⟨eq_zero_or_eq_zero_of_mul_eq_zero, fun o => o.elim (fun h => mul_eq_zero_of_left h b) (mul_eq_zero_of_right a)⟩ +#align mul_eq_zero mul_eq_zero /-- If `α` has no zero divisors, then the product of two elements equals zero iff one of them equals zero. -/ @[simp] theorem zero_eq_mul : 0 = a * b ↔ a = 0 ∨ b = 0 := by rw [eq_comm, mul_eq_zero] +#align zero_eq_mul zero_eq_mul /-- If `α` has no zero divisors, then the product of two elements is nonzero iff both of them are nonzero. -/ theorem mul_ne_zero_iff : a * b ≠ 0 ↔ a ≠ 0 ∧ b ≠ 0 := mul_eq_zero.not.trans not_or +#align mul_ne_zero_iff mul_ne_zero_iff /-- If `α` has no zero divisors, then for elements `a, b : α`, `a * b` equals zero iff so is `b * a`. -/ theorem mul_eq_zero_comm : a * b = 0 ↔ b * a = 0 := mul_eq_zero.trans <| or_comm.trans mul_eq_zero.symm +#align mul_eq_zero_comm mul_eq_zero_comm /-- If `α` has no zero divisors, then for elements `a, b : α`, `a * b` is nonzero iff so is `b * a`. -/ theorem mul_ne_zero_comm : a * b ≠ 0 ↔ b * a ≠ 0 := mul_eq_zero_comm.not +#align mul_ne_zero_comm mul_ne_zero_comm theorem mul_self_eq_zero : a * a = 0 ↔ a = 0 := by simp +#align mul_self_eq_zero mul_self_eq_zero theorem zero_eq_mul_self : 0 = a * a ↔ a = 0 := by simp +#align zero_eq_mul_self zero_eq_mul_self theorem mul_self_ne_zero : a * a ≠ 0 ↔ a ≠ 0 := mul_self_eq_zero.not +#align mul_self_ne_zero mul_self_ne_zero theorem zero_ne_mul_self : 0 ≠ a * a ↔ a ≠ 0 := zero_eq_mul_self.not +#align zero_ne_mul_self zero_ne_mul_self end MulZeroClass diff --git a/Mathlib/Algebra/Hom/Embedding.lean b/Mathlib/Algebra/Hom/Embedding.lean index c20dc0fcfd848..5b9fca67abd8d 100644 --- a/Mathlib/Algebra/Hom/Embedding.lean +++ b/Mathlib/Algebra/Hom/Embedding.lean @@ -31,6 +31,8 @@ def mulLeftEmbedding {G : Type _} [LeftCancelSemigroup G] (g : G) : G ↪ G wher inj' := mul_right_injective g #align mul_left_embedding mulLeftEmbedding #align add_left_embedding addLeftEmbedding +#align add_left_embedding_apply addLeftEmbedding_apply +#align mul_left_embedding_apply mulLeftEmbedding_apply /-- The embedding of a right cancellative semigroup into itself by right multiplication by a fixed element. @@ -43,11 +45,15 @@ def mulRightEmbedding {G : Type _} [RightCancelSemigroup G] (g : G) : G ↪ G wh inj' := mul_left_injective g #align mul_right_embedding mulRightEmbedding #align add_right_embedding addRightEmbedding +#align mul_right_embedding_apply mulRightEmbedding_apply +#align add_right_embedding_apply addRightEmbedding_apply @[to_additive] theorem mul_left_embedding_eq_mul_right_embedding {G : Type _} [CancelCommMonoid G] (g : G) : mulLeftEmbedding g = mulRightEmbedding g := by ext exact mul_comm _ _ +#align mul_left_embedding_eq_mul_right_embedding mul_left_embedding_eq_mul_right_embedding +#align add_left_embedding_eq_add_right_embedding add_left_embedding_eq_add_right_embedding end LeftOrRightCancelSemigroup diff --git a/Mathlib/Algebra/Hom/Equiv/Basic.lean b/Mathlib/Algebra/Hom/Equiv/Basic.lean index 84d88ad305001..7ad28bbc20f9f 100644 --- a/Mathlib/Algebra/Hom/Equiv/Basic.lean +++ b/Mathlib/Algebra/Hom/Equiv/Basic.lean @@ -57,6 +57,7 @@ def MonoidHom.inverse {A B : Type _} [Monoid A] [Monoid B] (f : A →* B) (g : B { (f : A →ₙ* B).inverse g h₁ h₂ with toFun := g, map_one' := by rw [← f.map_one, h₁] } #align monoid_hom.inverse MonoidHom.inverse #align add_monoid_hom.inverse AddMonoidHom.inverse +#align monoid_hom.inverse_apply MonoidHom.inverse_apply /-- `AddEquiv α β` is the type of an equiv `α ≃ β` which preserves addition. -/ structure AddEquiv (A B : Type _) [Add A] [Add B] extends A ≃ B, AddHom A B @@ -72,9 +73,11 @@ class AddEquivClass (F : Type _) (A B : outParam (Type _)) [Add A] [Add B] /-- The `Equiv` underlying an `AddEquiv`. -/ add_decl_doc AddEquiv.toEquiv +#align add_equiv.to_equiv AddEquiv.toEquiv /-- The `AddHom` underlying a `AddEquiv`. -/ add_decl_doc AddEquiv.toAddHom +#align add_equiv.to_add_hom AddEquiv.toAddHom /-- `MulEquiv α β` is the type of an equiv `α ≃ β` which preserves multiplication. -/ @[to_additive] @@ -86,9 +89,11 @@ attribute [to_additive] MulEquiv.toMulHom /-- The `Equiv` underlying a `MulEquiv`. -/ add_decl_doc MulEquiv.toEquiv +#align mul_equiv.to_equiv MulEquiv.toEquiv /-- The `MulHom` underlying a `MulEquiv`. -/ add_decl_doc MulEquiv.toMulHom +#align mul_equiv.to_mul_hom MulEquiv.toMulHom /-- `MulEquivClass F A B` states that `F` is a type of multiplication-preserving morphisms. You should extend this class when you extend `MulEquiv`. -/ @@ -599,6 +604,7 @@ def arrowCongr {M N P Q : Type _} [Mul P] [Mul Q] (f : M ≃ N) (g : P ≃* Q) : map_mul' h k := by ext; simp #align mul_equiv.arrow_congr MulEquiv.arrowCongr #align add_equiv.arrow_congr AddEquiv.arrowCongr +#align mul_equiv.arrow_congr_apply MulEquiv.arrowCongr_apply /-- A multiplicative analogue of `Equiv.arrowCongr`, for multiplicative maps from a monoid to a commutative monoid. @@ -648,6 +654,7 @@ def piCongrRight {η : Type _} {Ms Ns : η → Type _} [∀ j, Mul (Ms j)] [∀ map_mul' := fun x y => funext fun j => (es j).map_mul (x j) (y j) } #align mul_equiv.Pi_congr_right MulEquiv.piCongrRight #align add_equiv.Pi_congr_right AddEquiv.piCongrRight +#align mul_equiv.Pi_congr_right_apply MulEquiv.piCongrRight_apply @[to_additive (attr := simp)] theorem piCongrRight_refl {η : Type _} {Ms : η → Type _} [∀ j, Mul (Ms j)] : diff --git a/Mathlib/Algebra/Hom/Equiv/Units/Basic.lean b/Mathlib/Algebra/Hom/Equiv/Units/Basic.lean index 78d47f05e1d09..28149465d017a 100644 --- a/Mathlib/Algebra/Hom/Equiv/Units/Basic.lean +++ b/Mathlib/Algebra/Hom/Equiv/Units/Basic.lean @@ -68,6 +68,7 @@ def mulLeft (u : Mˣ) : Equiv.Perm M where right_inv := u.mul_inv_cancel_left #align units.mul_left Units.mulLeft #align add_units.add_left AddUnits.addLeft +#align units.mul_left_apply Units.mulLeft_apply @[to_additive (attr := simp)] theorem mulLeft_symm (u : Mˣ) : u.mulLeft.symm = u⁻¹.mulLeft := @@ -91,6 +92,7 @@ def mulRight (u : Mˣ) : Equiv.Perm M where right_inv x := inv_mul_cancel_right x u #align units.mul_right Units.mulRight #align add_units.add_right AddUnits.addRight +#align units.mul_right_apply Units.mulRight_apply @[to_additive (attr := simp)] theorem mulRight_symm (u : Mˣ) : u.mulRight.symm = u⁻¹.mulRight := @@ -189,6 +191,8 @@ protected def divLeft (a : G) : G ≃ G where right_inv b := by simp [div_eq_mul_inv] #align equiv.div_left Equiv.divLeft #align equiv.sub_left Equiv.subLeft +#align equiv.div_left_apply Equiv.divLeft_apply +#align equiv.div_left_symm_apply Equiv.divLeft_symm_apply @[to_additive] theorem divLeft_eq_inv_trans_mulLeft (a : G) : @@ -207,6 +211,8 @@ protected def divRight (a : G) : G ≃ right_inv b := by simp [div_eq_mul_inv] #align equiv.div_right Equiv.divRight #align equiv.sub_right Equiv.subRight +#align equiv.div_right_symm_apply Equiv.divRight_symm_apply +#align equiv.div_right_apply Equiv.divRight_apply @[to_additive] theorem divRight_eq_mulRight_inv (a : G) : Equiv.divRight a = Equiv.mulRight a⁻¹ := diff --git a/Mathlib/Algebra/Hom/Equiv/Units/GroupWithZero.lean b/Mathlib/Algebra/Hom/Equiv/Units/GroupWithZero.lean index d47d3fa58dc0e..04871674ec77d 100644 --- a/Mathlib/Algebra/Hom/Equiv/Units/GroupWithZero.lean +++ b/Mathlib/Algebra/Hom/Equiv/Units/GroupWithZero.lean @@ -30,6 +30,8 @@ underlying type. -/ protected def mulLeft₀ (a : G) (ha : a ≠ 0) : Perm G := (Units.mk0 a ha).mulLeft #align equiv.mul_left₀ Equiv.mulLeft₀ +#align equiv.mul_left₀_symm_apply Equiv.mulLeft₀_symm_apply +#align equiv.mul_left₀_apply Equiv.mulLeft₀_apply theorem mulLeft_bijective₀ (a : G) (ha : a ≠ 0) : Function.Bijective ((· * ·) a : G → G) := (Equiv.mulLeft₀ a ha).bijective @@ -41,6 +43,8 @@ underlying type. -/ protected def mulRight₀ (a : G) (ha : a ≠ 0) : Perm G := (Units.mk0 a ha).mulRight #align equiv.mul_right₀ Equiv.mulRight₀ +#align equiv.mul_right₀_symm_apply Equiv.mulRight₀_symm_apply +#align equiv.mul_right₀_apply Equiv.mulRight₀_apply theorem mulRight_bijective₀ (a : G) (ha : a ≠ 0) : Function.Bijective ((· * a) : G → G) := (Equiv.mulRight₀ a ha).bijective diff --git a/Mathlib/Algebra/Hom/Freiman.lean b/Mathlib/Algebra/Hom/Freiman.lean index 52ce48ce4f76f..fda573cbce89f 100644 --- a/Mathlib/Algebra/Hom/Freiman.lean +++ b/Mathlib/Algebra/Hom/Freiman.lean @@ -222,6 +222,7 @@ protected def id (A : Set α) (n : ℕ) : A →*[n] α where map_prod_eq_map_prod' _ _ _ _ h := by rw [map_id', map_id', h] #align freiman_hom.id FreimanHom.id #align add_freiman_hom.id AddFreimanHom.id +#align freiman_hom.id_apply FreimanHom.id_apply /-- Composition of Freiman homomorphisms as a Freiman homomorphism. -/ @[to_additive "Composition of additive Freiman homomorphisms as an additive Freiman homomorphism."] @@ -287,6 +288,7 @@ theorem comp_id (f : A →*[n] β) {hf} : f.comp (FreimanHom.id A n) hf = f := theorem id_comp (f : A →*[n] β) {hf} : (FreimanHom.id B n).comp f hf = f := ext fun _ => rfl #align freiman_hom.id_comp FreimanHom.id_comp +#align add_freiman_hom.id_comp AddFreimanHom.id_comp /-- `FreimanHom.const A n b` is the Freiman homomorphism sending everything to `b`. -/ @[to_additive "`AddFreimanHom.const An b` is the Freiman homomorphism sending everything to `b`."] diff --git a/Mathlib/Algebra/Hom/Group.lean b/Mathlib/Algebra/Hom/Group.lean index 728691f26092a..c6ef3eaab7e47 100644 --- a/Mathlib/Algebra/Hom/Group.lean +++ b/Mathlib/Algebra/Hom/Group.lean @@ -955,6 +955,8 @@ def OneHom.id (M : Type _) [One M] : OneHom M M where map_one' := rfl #align one_hom.id OneHom.id #align zero_hom.id ZeroHom.id +#align zero_hom.id_apply ZeroHom.id_apply +#align one_hom.id_apply OneHom.id_apply /-- The identity map from a type with multiplication to itself. -/ @[to_additive (attr := simps)] @@ -963,6 +965,8 @@ def MulHom.id (M : Type _) [Mul M] : M →ₙ* M where map_mul' _ _ := rfl #align mul_hom.id MulHom.id #align add_hom.id AddHom.id +#align add_hom.id_apply AddHom.id_apply +#align mul_hom.id_apply MulHom.id_apply /-- The identity map from a monoid to itself. -/ @[to_additive (attr := simps)] @@ -972,6 +976,8 @@ def MonoidHom.id (M : Type _) [MulOneClass M] : M →* M where map_mul' _ _ := rfl #align monoid_hom.id MonoidHom.id #align add_monoid_hom.id AddMonoidHom.id +#align monoid_hom.id_apply MonoidHom.id_apply +#align add_monoid_hom.id_apply AddMonoidHom.id_apply /-- The identity map from a `MonoidWithZero` to itself. -/ @[simps] @@ -981,6 +987,7 @@ def MonoidWithZeroHom.id (M : Type _) [MulZeroOneClass M] : M →*₀ M where map_one' := rfl map_mul' _ _ := rfl #align monoid_with_zero_hom.id MonoidWithZeroHom.id +#align monoid_with_zero_hom.id_apply MonoidWithZeroHom.id_apply /-- The identity map from an type with zero to itself. -/ add_decl_doc ZeroHom.id @@ -1548,6 +1555,8 @@ def mk' (f : M → G) (map_mul : ∀ a b : M, f (a * b) = f a * f b) : M →* G map_one' := mul_left_eq_self.1 <| by rw [← map_mul, mul_one] #align monoid_hom.mk' MonoidHom.mk' #align add_monoid_hom.mk' AddMonoidHom.mk' +#align add_monoid_hom.mk'_apply AddMonoidHom.mk'_apply +#align monoid_hom.mk'_apply MonoidHom.mk'_apply /-- Makes a group homomorphism from a proof that the map preserves right division `fun x y => x * y⁻¹`. See also `MonoidHom.of_map_div` for a version using `fun x y => x / y`. diff --git a/Mathlib/Algebra/Hom/GroupAction.lean b/Mathlib/Algebra/Hom/GroupAction.lean index 5f9253369655f..1962cd5c27455 100644 --- a/Mathlib/Algebra/Hom/GroupAction.lean +++ b/Mathlib/Algebra/Hom/GroupAction.lean @@ -185,6 +185,7 @@ def inverse (f : A →[M] B) (g : B → A) (h₁ : Function.LeftInverse g f) g (m • x) = g (m • f (g x)) := by rw [h₂] _ = g (f (m • g x)) := by rw [f.map_smul] _ = m • g x := by rw [h₁] +#align mul_action_hom.inverse_to_fun MulActionHom.inverse_toFun #align mul_action_hom.inverse MulActionHom.inverse @@ -196,9 +197,11 @@ structure DistribMulActionHom extends A →[M] B, A →+ B /-- Reinterpret an equivariant additive monoid homomorphism as an additive monoid homomorphism. -/ add_decl_doc DistribMulActionHom.toAddMonoidHom +#align distrib_mul_action_hom.to_add_monoid_hom DistribMulActionHom.toAddMonoidHom /-- Reinterpret an equivariant additive monoid homomorphism as an equivariant function. -/ add_decl_doc DistribMulActionHom.toMulActionHom +#align distrib_mul_action_hom.to_mul_action_hom DistribMulActionHom.toMulActionHom /- Porting note: local notation given a name, conflict with Algebra.Hom.Freiman see https://github.com/leanprover/lean4/issues/2000 -/ @@ -417,9 +420,11 @@ structure MulSemiringActionHom extends R →+[M] S, R →+* S /-- Reinterpret an equivariant ring homomorphism as a ring homomorphism. -/ add_decl_doc MulSemiringActionHom.toRingHom +#align mul_semiring_action_hom.to_ring_hom MulSemiringActionHom.toRingHom /-- Reinterpret an equivariant ring homomorphism as an equivariant additive monoid homomorphism. -/ add_decl_doc MulSemiringActionHom.toDistribMulActionHom +#align mul_semiring_action_hom.to_distrib_mul_action_hom MulSemiringActionHom.toDistribMulActionHom /- Porting note: local notation given a name, conflict with Algebra.Hom.Freiman see https://github.com/leanprover/lean4/issues/2000 -/ diff --git a/Mathlib/Algebra/Hom/GroupInstances.lean b/Mathlib/Algebra/Hom/GroupInstances.lean index e0aaba13f34c4..23351d5e46013 100644 --- a/Mathlib/Algebra/Hom/GroupInstances.lean +++ b/Mathlib/Algebra/Hom/GroupInstances.lean @@ -187,6 +187,7 @@ def eval [MulOneClass M] [CommMonoid N] : M →* (M →* N) →* N := (MonoidHom.id (M →* N)).flip #align monoid_hom.eval MonoidHom.eval #align add_monoid_hom.eval AddMonoidHom.eval +#align monoid_hom.eval_apply_apply MonoidHom.eval_apply_apply /-- The expression `λ g m, g (f m)` as a `MonoidHom`. Equivalently, `(λ g, MonoidHom.comp g f)` as a `MonoidHom`. -/ @@ -200,6 +201,7 @@ def compHom' [MulOneClass M] [MulOneClass N] [CommMonoid P] (f : M →* N) : (N flip <| eval.comp f #align monoid_hom.comp_hom' MonoidHom.compHom' #align add_monoid_hom.comp_hom' AddMonoidHom.compHom' +#align monoid_hom.comp_hom'_apply_apply MonoidHom.compHom'_apply_apply /-- Composition of monoid morphisms (`MonoidHom.comp`) as a monoid morphism. @@ -223,6 +225,7 @@ def compHom [MulOneClass M] [CommMonoid N] [CommMonoid P] : exact mul_comp g₁ g₂ f #align monoid_hom.comp_hom MonoidHom.compHom #align add_monoid_hom.comp_hom AddMonoidHom.compHom +#align monoid_hom.comp_hom_apply_apply MonoidHom.compHom_apply_apply /-- Flipping arguments of monoid morphisms (`MonoidHom.flip`) as a monoid morphism. -/ @[to_additive @@ -236,6 +239,7 @@ def flipHom {_ : MulOneClass M} {_ : MulOneClass N} {_ : CommMonoid P} : map_mul' _ _ := rfl #align monoid_hom.flip_hom MonoidHom.flipHom #align add_monoid_hom.flip_hom AddMonoidHom.flipHom +#align monoid_hom.flip_hom_apply MonoidHom.flipHom_apply /-- The expression `λ m q, f m (g q)` as a `MonoidHom`. @@ -335,11 +339,13 @@ theorem AddMonoidHom.map_mul_iff (f : R →+ S) : def AddMonoid.End.mulLeft : R →+ AddMonoid.End R := AddMonoidHom.mul #align add_monoid.End.mul_left AddMonoid.End.mulLeft +#align add_monoid.End.mul_left_apply_apply AddMonoid.End.mulLeft_apply_apply /-- The right multiplication map: `(a, b) ↦ b * a`. See also `AddMonoidHom.mulRight`. -/ @[simps] def AddMonoid.End.mulRight : R →+ AddMonoid.End R := (AddMonoidHom.mul : R →+ AddMonoid.End R).flip #align add_monoid.End.mul_right AddMonoid.End.mulRight +#align add_monoid.End.mul_right_apply_apply AddMonoid.End.mulRight_apply_apply end Semiring diff --git a/Mathlib/Algebra/Hom/Iterate.lean b/Mathlib/Algebra/Hom/Iterate.lean index a97cae653d5c3..06cebab775c5c 100644 --- a/Mathlib/Algebra/Hom/Iterate.lean +++ b/Mathlib/Algebra/Hom/Iterate.lean @@ -104,6 +104,7 @@ theorem iterate_map_smul (f : M →+ M) (n m : ℕ) (x : M) : (f^[n]) (m • x) #align add_monoid_hom.iterate_map_smul AddMonoidHom.iterate_map_smul attribute [to_additive (reorder := 5)] MonoidHom.iterate_map_pow +#align add_monoid_hom.iterate_map_nsmul AddMonoidHom.iterate_map_nsmul theorem iterate_map_zsmul (f : G →+ G) (n : ℕ) (m : ℤ) (x : G) : (f^[n]) (m • x) = m • (f^[n]) x := f.toMultiplicative.iterate_map_zpow n x m @@ -207,6 +208,7 @@ theorem pow_iterate (n : ℕ) (j : ℕ) : (fun x : G => x ^ n)^[j] = fun x : G = mul_smul := fun m n g => pow_mul' g m n } smul_iterate n j #align pow_iterate pow_iterate +#align nsmul_iterate nsmul_iterate end Monoid diff --git a/Mathlib/Algebra/Hom/Ring.lean b/Mathlib/Algebra/Hom/Ring.lean index cd41aa4bf47df..191d20598b978 100644 --- a/Mathlib/Algebra/Hom/Ring.lean +++ b/Mathlib/Algebra/Hom/Ring.lean @@ -72,10 +72,12 @@ infixr:25 " →ₙ+* " => NonUnitalRingHom /-- Reinterpret a non-unital ring homomorphism `f : α →ₙ+* β` as a semigroup homomorphism `α →ₙ* β`. The `simp`-normal form is `(f : α →ₙ* β)`. -/ add_decl_doc NonUnitalRingHom.toMulHom +#align non_unital_ring_hom.to_mul_hom NonUnitalRingHom.toMulHom /-- Reinterpret a non-unital ring homomorphism `f : α →ₙ+* β` as an additive monoid homomorphism `α →+ β`. The `simp`-normal form is `(f : α →+ β)`. -/ add_decl_doc NonUnitalRingHom.toAddMonoidHom +#align non_unital_ring_hom.to_add_monoid_hom NonUnitalRingHom.toAddMonoidHom section NonUnitalRingHomClass @@ -362,18 +364,22 @@ infixr:25 " →+* " => RingHom /-- Reinterpret a ring homomorphism `f : α →+* β` as a monoid with zero homomorphism `α →*₀ β`. The `simp`-normal form is `(f : α →*₀ β)`. -/ add_decl_doc RingHom.toMonoidWithZeroHom +#align ring_hom.to_monoid_with_zero_hom RingHom.toMonoidWithZeroHom /-- Reinterpret a ring homomorphism `f : α →+* β` as a monoid homomorphism `α →* β`. The `simp`-normal form is `(f : α →* β)`. -/ add_decl_doc RingHom.toMonoidHom +#align ring_hom.to_monoid_hom RingHom.toMonoidHom /-- Reinterpret a ring homomorphism `f : α →+* β` as an additive monoid homomorphism `α →+ β`. The `simp`-normal form is `(f : α →+ β)`. -/ add_decl_doc RingHom.toAddMonoidHom +#align ring_hom.to_add_monoid_hom RingHom.toAddMonoidHom /-- Reinterpret a ring homomorphism `f : α →+* β` as a non-unital ring homomorphism `α →ₙ+* β`. The `simp`-normal form is `(f : α →ₙ+* β)`. -/ add_decl_doc RingHom.toNonUnitalRingHom +#align ring_hom.to_non_unital_ring_hom RingHom.toNonUnitalRingHom section RingHomClass diff --git a/Mathlib/Algebra/Homology/ComplexShape.lean b/Mathlib/Algebra/Homology/ComplexShape.lean index a4fe5637b7402..2885d5aec5834 100644 --- a/Mathlib/Algebra/Homology/ComplexShape.lean +++ b/Mathlib/Algebra/Homology/ComplexShape.lean @@ -70,6 +70,9 @@ structure ComplexShape (ι : Type _) where next_eq : ∀ {i j j'}, Rel i j → Rel i j' → j = j' /-- There is at most one nonzero differential to `X j`. -/ prev_eq : ∀ {i i' j}, Rel i j → Rel i' j → i = i' +#align complex_shape ComplexShape +#align complex_shape.ext ComplexShape.ext +#align complex_shape.ext_iff ComplexShape.ext_iff namespace ComplexShape @@ -84,6 +87,8 @@ def refl (ι : Type _) : ComplexShape ι where Rel i j := i = j next_eq w w' := w.symm.trans w' prev_eq w w' := w.trans w'.symm +#align complex_shape.refl ComplexShape.refl +#align complex_shape.refl_rel ComplexShape.refl_Rel /-- The reverse of a `ComplexShape`. -/ @@ -92,11 +97,14 @@ def symm (c : ComplexShape ι) : ComplexShape ι where Rel i j := c.Rel j i next_eq w w' := c.prev_eq w w' prev_eq w w' := c.next_eq w w' +#align complex_shape.symm ComplexShape.symm +#align complex_shape.symm_rel ComplexShape.symm_Rel @[simp] theorem symm_symm (c : ComplexShape ι) : c.symm.symm = c := by ext simp +#align complex_shape.symm_symm ComplexShape.symm_symm /-- The "composition" of two `ComplexShape`s. @@ -115,6 +123,7 @@ def trans (c₁ c₂ : ComplexShape ι) : ComplexShape ι where obtain ⟨k', w₁', w₂'⟩ := w' rw [c₂.prev_eq w₂ w₂'] at w₁ exact c₁.prev_eq w₁ w₁' +#align complex_shape.trans ComplexShape.trans instance subsingleton_next (c : ComplexShape ι) (i : ι) : Subsingleton { j // c.Rel i j } := by constructor @@ -133,23 +142,27 @@ Returns `i` otherwise. -/ def next (c : ComplexShape ι) (i : ι) : ι := if h : ∃ j, c.Rel i j then h.choose else i +#align complex_shape.next ComplexShape.next /-- An arbitary choice of index `i` such that `Rel i j`, if such exists. Returns `j` otherwise. -/ def prev (c : ComplexShape ι) (j : ι) : ι := if h : ∃ i, c.Rel i j then h.choose else j +#align complex_shape.prev ComplexShape.prev theorem next_eq' (c : ComplexShape ι) {i j : ι} (h : c.Rel i j) : c.next i = j := by apply c.next_eq _ h rw [next] rw [dif_pos] exact Exists.choose_spec ⟨j, h⟩ +#align complex_shape.next_eq' ComplexShape.next_eq' theorem prev_eq' (c : ComplexShape ι) {i j : ι} (h : c.Rel i j) : c.prev j = i := by apply c.prev_eq _ h rw [prev, dif_pos] exact Exists.choose_spec (⟨i, h⟩ : ∃ k, c.Rel k j) +#align complex_shape.prev_eq' ComplexShape.prev_eq' /-- The `ComplexShape` allowing differentials from `X i` to `X (i+a)`. (For example when `a = 1`, a cohomology theory indexed by `ℕ` or `ℤ`) @@ -159,6 +172,8 @@ def up' {α : Type _} [AddRightCancelSemigroup α] (a : α) : ComplexShape α wh Rel i j := i + a = j next_eq hi hj := hi.symm.trans hj prev_eq hi hj := add_right_cancel (hi.trans hj.symm) +#align complex_shape.up' ComplexShape.up' +#align complex_shape.up'_rel ComplexShape.up'_Rel /-- The `ComplexShape` allowing differentials from `X (j+a)` to `X j`. (For example when `a = 1`, a homology theory indexed by `ℕ` or `ℤ`) @@ -168,24 +183,32 @@ def down' {α : Type _} [AddRightCancelSemigroup α] (a : α) : ComplexShape α Rel i j := j + a = i next_eq hi hj := add_right_cancel (hi.trans hj.symm) prev_eq hi hj := hi.symm.trans hj +#align complex_shape.down' ComplexShape.down' +#align complex_shape.down'_rel ComplexShape.down'_Rel theorem down'_mk {α : Type _} [AddRightCancelSemigroup α] (a : α) (i j : α) (h : j + a = i) : (down' a).Rel i j := h +#align complex_shape.down'_mk ComplexShape.down'_mk /-- The `ComplexShape` appropriate for cohomology, so `d : X i ⟶ X j` only when `j = i + 1`. -/ @[simps] def up (α : Type _) [AddRightCancelSemigroup α] [One α] : ComplexShape α := up' 1 +#align complex_shape.up ComplexShape.up +#align complex_shape.up_rel ComplexShape.up_Rel /-- The `ComplexShape` appropriate for homology, so `d : X i ⟶ X j` only when `i = j + 1`. -/ @[simps] def down (α : Type _) [AddRightCancelSemigroup α] [One α] : ComplexShape α := down' 1 +#align complex_shape.down ComplexShape.down +#align complex_shape.down_rel ComplexShape.down_Rel theorem down_mk {α : Type _} [AddRightCancelSemigroup α] [One α] (i j : α) (h : j + 1 = i) : (down α).Rel i j := down'_mk (1 : α) i j h +#align complex_shape.down_mk ComplexShape.down_mk end ComplexShape diff --git a/Mathlib/Algebra/Invertible.lean b/Mathlib/Algebra/Invertible.lean index b6d5806f15191..370881a3e6ec2 100644 --- a/Mathlib/Algebra/Invertible.lean +++ b/Mathlib/Algebra/Invertible.lean @@ -418,6 +418,7 @@ def Invertible.ofLeftInverse {R : Type _} {S : Type _} {G : Type _} [MulOneClass [Invertible (f r)] : Invertible r := (Invertible.map g (f r)).copy _ (h r).symm #align invertible.of_left_inverse Invertible.ofLeftInverse +#align invertible.of_left_inverse_inv_of Invertible.ofLeftInverse_invOf /-- Invertibility on either side of a monoid hom with a left-inverse is equivalent. -/ @[simps] @@ -431,3 +432,5 @@ def invertibleEquivOfLeftInverse {R : Type _} {S : Type _} {F G : Type _} [Monoi left_inv _ := Subsingleton.elim _ _ right_inv _ := Subsingleton.elim _ _ #align invertible_equiv_of_left_inverse invertibleEquivOfLeftInverse +#align invertible_equiv_of_left_inverse_symm_apply invertibleEquivOfLeftInverse_symm_apply +#align invertible_equiv_of_left_inverse_apply invertibleEquivOfLeftInverse_apply diff --git a/Mathlib/Algebra/Module/Basic.lean b/Mathlib/Algebra/Module/Basic.lean index f7e9d92138b4b..ce4e7c39d9229 100644 --- a/Mathlib/Algebra/Module/Basic.lean +++ b/Mathlib/Algebra/Module/Basic.lean @@ -63,6 +63,8 @@ class Module (R : Type u) (M : Type v) [Semiring R] [AddCommMonoid M] extends /-- Scalar multiplication by zero gives zero. -/ protected zero_smul : ∀ x : M, (0 : R) • x = 0 #align module Module +#align module.ext Module.ext +#align module.ext_iff Module.ext_iff section AddCommMonoid @@ -185,6 +187,7 @@ def Module.toAddMonoidEnd : R →+* AddMonoid.End M := map_add' := fun x y => AddMonoidHom.ext fun r => show (x + y) • r = x • r + y • r by simp [add_smul] } #align module.to_add_monoid_End Module.toAddMonoidEnd +#align module.to_add_monoid_End_apply_apply Module.toAddMonoidEnd_apply_apply /-- A convenience alias for `Module.toAddMonoidEnd` as an `AddMonoidHom`, usually to allow the use of `AddMonoidHom.flip`. -/ diff --git a/Mathlib/Algebra/Module/Equiv.lean b/Mathlib/Algebra/Module/Equiv.lean index f982a9a55df9c..6d7349b51f94c 100644 --- a/Mathlib/Algebra/Module/Equiv.lean +++ b/Mathlib/Algebra/Module/Equiv.lean @@ -57,9 +57,11 @@ structure LinearEquiv {R : Type _} {S : Type _} [Semiring R] [Semiring S] (σ : /-- The linear map underlying a linear equivalence. -/ add_decl_doc LinearEquiv.toLinearMap +#align linear_equiv.to_linear_map LinearEquiv.toLinearMap /-- The additive equivalence of types underlying a linear equivalence. -/ add_decl_doc LinearEquiv.toAddEquiv +#align linear_equiv.to_add_equiv LinearEquiv.toAddEquiv /-- The backwards directed function underlying a linear equivalence. -/ add_decl_doc LinearEquiv.invFun @@ -565,6 +567,7 @@ def _root_.RingEquiv.toSemilinearEquiv (f : R ≃+* S) : toFun := f map_smul' := f.map_mul } #align ring_equiv.to_semilinear_equiv RingEquiv.toSemilinearEquiv +#align ring_equiv.to_semilinear_equiv_symm_apply RingEquiv.toSemilinearEquiv_symmApply variable [Semiring R₁] [Semiring R₂] [Semiring R₃] @@ -601,6 +604,7 @@ def restrictScalars (f : M ≃ₗ[S] M₂) : M ≃ₗ[R] M₂ := left_inv := f.left_inv right_inv := f.right_inv } #align linear_equiv.restrict_scalars LinearEquiv.restrictScalars +#align linear_equiv.restrict_scalars_symm_apply LinearEquiv.restrictScalars_symmApply theorem restrictScalars_injective : Function.Injective (restrictScalars R : (M ≃ₗ[S] M₂) → M ≃ₗ[R] M₂) := fun _ _ h => @@ -639,6 +643,7 @@ def automorphismGroup.toLinearMapMonoidHom : (M ≃ₗ[R] M) →* M →ₗ[R] M map_one' := rfl map_mul' _ _ := rfl #align linear_equiv.automorphism_group.to_linear_map_monoid_hom LinearEquiv.automorphismGroup.toLinearMapMonoidHom +#align linear_equiv.automorphism_group.to_linear_map_monoid_hom_apply LinearEquiv.automorphismGroup.toLinearMapMonoidHom_apply /-- The tautological action by `M ≃ₗ[R] M` on `M`. @@ -686,6 +691,7 @@ def ofSubsingleton : M ≃ₗ[R] M₂ := left_inv := fun _ => Subsingleton.elim _ _ right_inv := fun _ => Subsingleton.elim _ _ } #align linear_equiv.of_subsingleton LinearEquiv.ofSubsingleton +#align linear_equiv.of_subsingleton_symm_apply LinearEquiv.ofSubsingleton_symmApply @[simp] theorem ofSubsingleton_self : ofSubsingleton M M = refl R M := by @@ -712,6 +718,7 @@ def compHom.toLinearEquiv {R S : Type _} [Semiring R] [Semiring S] (g : R ≃+* invFun := (g.symm : S → R) map_smul' := g.map_mul } #align module.comp_hom.to_linear_equiv Module.compHom.toLinearEquiv +#align module.comp_hom.to_linear_equiv_symm_apply Module.compHom.toLinearEquiv_symmApply end Module @@ -728,6 +735,8 @@ This is a stronger version of `DistribMulAction.toAddEquiv`. -/ def toLinearEquiv (s : S) : M ≃ₗ[R] M := { toAddEquiv M s, toLinearMap R M s with } #align distrib_mul_action.to_linear_equiv DistribMulAction.toLinearEquiv +#align distrib_mul_action.to_linear_equiv_apply DistribMulAction.toLinearEquiv_apply +#align distrib_mul_action.to_linear_equiv_symm_apply DistribMulAction.toLinearEquiv_symmApply /-- Each element of the group defines a module automorphism. @@ -738,6 +747,7 @@ def toModuleAut : S →* M ≃ₗ[R] M where map_one' := LinearEquiv.ext <| one_smul _ map_mul' _ _ := LinearEquiv.ext <| mul_smul _ _ #align distrib_mul_action.to_module_aut DistribMulAction.toModuleAut +#align distrib_mul_action.to_module_aut_apply DistribMulAction.toModuleAut_apply end DistribMulAction diff --git a/Mathlib/Algebra/Module/LinearMap.lean b/Mathlib/Algebra/Module/LinearMap.lean index 3c17ca4c2cf8b..91a894cd372f4 100644 --- a/Mathlib/Algebra/Module/LinearMap.lean +++ b/Mathlib/Algebra/Module/LinearMap.lean @@ -104,6 +104,7 @@ structure LinearMap {R : Type _} {S : Type _} [Semiring R] [Semiring S] (σ : R /-- The `add_hom` underlying a `LinearMap`. -/ add_decl_doc LinearMap.toAddHom +#align linear_map.to_add_hom LinearMap.toAddHom -- mathport name: «expr →ₛₗ[ ] » /-- `M →ₛₗ[σ] N` is the type of `σ`-semilinear maps from `M` to `N`. -/ @@ -518,6 +519,7 @@ def _root_.RingHom.toSemilinearMap (f : R →+* S) : R →ₛₗ[f] S := toFun := f map_smul' := f.map_mul } #align ring_hom.to_semilinear_map RingHom.toSemilinearMap +#align ring_hom.to_semilinear_map_apply RingHom.toSemilinearMap_apply section @@ -639,6 +641,7 @@ letI := compHom S g map_add' := g.map_add map_smul' := g.map_mul } #align module.comp_hom.to_linear_map Module.compHom.toLinearMap +#align module.comp_hom.to_linear_map_apply Module.compHom.toLinearMap_apply end Module @@ -1184,6 +1187,7 @@ def toLinearMap (s : S) : M →ₗ[R] M where map_add' := smul_add s map_smul' _ _ := smul_comm _ _ _ #align distrib_mul_action.to_linear_map DistribMulAction.toLinearMap +#align distrib_mul_action.to_linear_map_apply DistribMulAction.toLinearMap_apply /-- Each element of the monoid defines a module endomorphism. @@ -1195,6 +1199,7 @@ def toModuleEnd : S →* Module.End R M map_one' := LinearMap.ext <| one_smul _ map_mul' _ _ := LinearMap.ext <| mul_smul _ _ #align distrib_mul_action.to_module_End DistribMulAction.toModuleEnd +#align distrib_mul_action.to_module_End_apply DistribMulAction.toModuleEnd_apply end DistribMulAction @@ -1216,6 +1221,7 @@ def toModuleEnd : S →+* Module.End R M := map_zero' := LinearMap.ext <| zero_smul _ map_add' := fun _ _ ↦ LinearMap.ext <| add_smul _ _ } #align module.to_module_End Module.toModuleEnd +#align module.to_module_End_apply Module.toModuleEnd_apply /-- The canonical (semi)ring isomorphism from `Rᵐᵒᵖ` to `Module.End R R` induced by the right multiplication. -/ @@ -1227,6 +1233,7 @@ def moduleEndSelf : Rᵐᵒᵖ ≃+* Module.End R R := left_inv := mul_one right_inv := fun _ ↦ LinearMap.ext_ring <| one_mul _ } #align module.module_End_self Module.moduleEndSelf +#align module.module_End_self_apply Module.moduleEndSelf_apply /-- The canonical (semi)ring isomorphism from `R` to `Module.End Rᵐᵒᵖ R` induced by the left multiplication. -/ @@ -1238,6 +1245,8 @@ def moduleEndSelfOp : R ≃+* Module.End Rᵐᵒᵖ R := left_inv := mul_one right_inv := fun _ ↦ LinearMap.ext_ring_op <| mul_one _ } #align module.module_End_self_op Module.moduleEndSelfOp +#align module.module_End_self_op_symm_apply Module.moduleEndSelfOp_symmApply +#align module.module_End_self_op_apply Module.moduleEndSelfOp_apply theorem End.natCast_def (n : ℕ) [AddCommMonoid N₁] [Module R N₁] : (↑n : Module.End R N₁) = Module.toModuleEnd R N₁ n := diff --git a/Mathlib/Algebra/Module/Submodule/Basic.lean b/Mathlib/Algebra/Module/Submodule/Basic.lean index 859d845780935..5d88cc8c1ff4f 100644 --- a/Mathlib/Algebra/Module/Submodule/Basic.lean +++ b/Mathlib/Algebra/Module/Submodule/Basic.lean @@ -52,9 +52,11 @@ structure Submodule (R : Type u) (M : Type v) [Semiring R] [AddCommMonoid M] [Mo /-- Reinterpret a `Submodule` as an `AddSubmonoid`. -/ add_decl_doc Submodule.toAddSubmonoid +#align submodule.to_add_submonoid Submodule.toAddSubmonoid /-- Reinterpret a `Submodule` as an `SubMulAction`. -/ add_decl_doc Submodule.toSubMulAction +#align submodule.to_sub_mul_action Submodule.toSubMulAction namespace Submodule @@ -443,6 +445,7 @@ def restrictScalarsEmbedding : Submodule R M ↪o Submodule S M inj' := restrictScalars_injective S R M map_rel_iff' := by simp [SetLike.le_def] #align submodule.restrict_scalars_embedding Submodule.restrictScalarsEmbedding +#align submodule.restrict_scalars_embedding_apply Submodule.restrictScalarsEmbedding_apply /-- Turning `p : Submodule R M` into an `S`-submodule gives the same module structure as turning it into a type and adding a module structure. -/ @@ -453,6 +456,7 @@ def restrictScalarsEquiv (p : Submodule R M) : p.restrictScalars S ≃ₗ[R] p : invFun := id map_smul' := fun _ _ => rfl } #align submodule.restrict_scalars_equiv Submodule.restrictScalarsEquiv +#align submodule.restrict_scalars_equiv_symm_apply Submodule.restrictScalarsEquiv_symmApply end RestrictScalars diff --git a/Mathlib/Algebra/NeZero.lean b/Mathlib/Algebra/NeZero.lean index 9abc3ffbe7769..ede5d3d73585b 100644 --- a/Mathlib/Algebra/NeZero.lean +++ b/Mathlib/Algebra/NeZero.lean @@ -27,12 +27,15 @@ We create a typeclass `NeZero n` which carries around the fact that `(n : R) ≠ class NeZero {R} [Zero R] (n : R) : Prop where /-- The proposition that `n` is not zero. -/ out : n ≠ 0 +#align ne_zero NeZero theorem NeZero.ne {R} [Zero R] (n : R) [h : NeZero n] : n ≠ 0 := h.out +#align ne_zero.ne NeZero.ne theorem NeZero.ne' {R} [Zero R] (n : R) [h : NeZero n] : 0 ≠ n := h.out.symm +#align ne_zero.ne' NeZero.ne' theorem neZero_iff {R : Type _} [Zero R] {n : R} : NeZero n ↔ n ≠ 0 := ⟨fun h ↦ h.out, NeZero.mk⟩ @@ -50,12 +53,18 @@ variable {α : Type _} [Zero α] @[simp] lemma zero_ne_one [One α] [NeZero (1 : α)] : (0 : α) ≠ 1 := NeZero.ne' (1 : α) @[simp] lemma one_ne_zero [One α] [NeZero (1 : α)] : (1 : α) ≠ 0 := NeZero.ne (1 : α) +#align one_ne_zero one_ne_zero +#align zero_ne_one zero_ne_one lemma ne_zero_of_eq_one [One α] [NeZero (1 : α)] {a : α} (h : a = 1) : a ≠ 0 := h ▸ one_ne_zero +#align ne_zero_of_eq_one ne_zero_of_eq_one lemma two_ne_zero [OfNat α 2] [NeZero (2 : α)] : (2 : α) ≠ 0 := NeZero.ne (2 : α) lemma three_ne_zero [OfNat α 3] [NeZero (3 : α)] : (3 : α) ≠ 0 := NeZero.ne (3 : α) lemma four_ne_zero [OfNat α 4] [NeZero (4 : α)] : (4 : α) ≠ 0 := NeZero.ne (4 : α) +#align four_ne_zero four_ne_zero +#align three_ne_zero three_ne_zero +#align two_ne_zero two_ne_zero variable (α) @@ -64,6 +73,11 @@ lemma one_ne_zero' [One α] [NeZero (1 : α)] : (1 : α) ≠ 0 := one_ne_zero lemma two_ne_zero' [OfNat α 2] [NeZero (2 : α)] : (2 : α) ≠ 0 := two_ne_zero lemma three_ne_zero' [OfNat α 3] [NeZero (3 : α)] : (3 : α) ≠ 0 := three_ne_zero lemma four_ne_zero' [OfNat α 4] [NeZero (4 : α)] : (4 : α) ≠ 0 := four_ne_zero +#align four_ne_zero' four_ne_zero' +#align three_ne_zero' three_ne_zero' +#align two_ne_zero' two_ne_zero' +#align one_ne_zero' one_ne_zero' +#align zero_ne_one' zero_ne_one' end @@ -74,5 +88,6 @@ variable {M : Type _} {x : M} instance succ : NeZero (n + 1) := ⟨n.succ_ne_zero⟩ theorem of_pos [Preorder M] [Zero M] (h : 0 < x) : NeZero x := ⟨ne_of_gt h⟩ +#align ne_zero.of_pos NeZero.of_pos end NeZero diff --git a/Mathlib/Algebra/Opposites.lean b/Mathlib/Algebra/Opposites.lean index c50b2cb5ce23f..46cc1d07e8b99 100644 --- a/Mathlib/Algebra/Opposites.lean +++ b/Mathlib/Algebra/Opposites.lean @@ -55,6 +55,8 @@ structure MulOpposite (α : Type u) : Type u where /-- The element of `MulOpposite α` that represents `x : α`. -/ op :: /-- The element of `α` represented by `x : αᵐᵒᵖ`. -/ unop : α #align mul_opposite.op MulOpposite.op +#align mul_opposite.unop MulOpposite.unop +#align mul_opposite MulOpposite -- porting note: the attribute `pp_nodot` does not exist yet; `op` and `unop` were -- both tagged with it in mathlib3 @@ -64,6 +66,9 @@ structure MulOpposite (α : Type u) : Type u where structure AddOpposite (α : Type u) : Type u where /-- The element of `αᵃᵒᵖ` that represents `x : α`. -/ op :: /-- The element of `α` represented by `x : αᵃᵒᵖ`. -/ unop : α +#align add_opposite.unop AddOpposite.unop +#align add_opposite.op AddOpposite.op +#align add_opposite AddOpposite -- porting note: the attribute `pp_nodot` does not exist yet; `op` and `unop` were -- both tagged with it in mathlib3 @@ -81,18 +86,26 @@ namespace MulOpposite -- porting note: `simp` can prove this in Lean 4 @[to_additive] theorem unop_op (x : α) : unop (op x) = x := rfl +#align mul_opposite.unop_op MulOpposite.unop_op +#align add_opposite.unop_op AddOpposite.unop_op @[to_additive (attr := simp)] theorem op_unop (x : αᵐᵒᵖ) : op (unop x) = x := rfl +#align mul_opposite.op_unop MulOpposite.op_unop +#align add_opposite.op_unop AddOpposite.op_unop @[to_additive (attr := simp)] theorem op_comp_unop : (op : α → αᵐᵒᵖ) ∘ unop = id := rfl +#align mul_opposite.op_comp_unop MulOpposite.op_comp_unop +#align add_opposite.op_comp_unop AddOpposite.op_comp_unop @[to_additive (attr := simp)] theorem unop_comp_op : (unop : αᵐᵒᵖ → α) ∘ op = id := rfl +#align mul_opposite.unop_comp_op MulOpposite.unop_comp_op +#align add_opposite.unop_comp_op AddOpposite.unop_comp_op /-- A recursor for `MulOpposite`. Use as `induction x using MulOpposite.rec'`. -/ @[to_additive (attr := simp) @@ -106,22 +119,34 @@ protected def rec' {F : ∀ _ : αᵐᵒᵖ, Sort v} (h : ∀ X, F (op X)) : ∀ simps (config := { fullyApplied := false }) apply symm_apply] def opEquiv : α ≃ αᵐᵒᵖ := ⟨op, unop, unop_op, op_unop⟩ +#align mul_opposite.op_equiv MulOpposite.opEquiv +#align mul_opposite.op_equiv_apply MulOpposite.opEquiv_apply +#align mul_opposite.op_equiv_symm_apply MulOpposite.opEquiv_symm_apply +#align add_opposite.op_equiv AddOpposite.opEquiv @[to_additive] theorem op_bijective : Bijective (op : α → αᵐᵒᵖ) := opEquiv.bijective +#align mul_opposite.op_bijective MulOpposite.op_bijective +#align add_opposite.op_bijective AddOpposite.op_bijective @[to_additive] theorem unop_bijective : Bijective (unop : αᵐᵒᵖ → α) := opEquiv.symm.bijective +#align mul_opposite.unop_bijective MulOpposite.unop_bijective +#align add_opposite.unop_bijective AddOpposite.unop_bijective @[to_additive] theorem op_injective : Injective (op : α → αᵐᵒᵖ) := op_bijective.injective +#align mul_opposite.op_injective MulOpposite.op_injective +#align add_opposite.op_injective AddOpposite.op_injective @[to_additive] theorem op_surjective : Surjective (op : α → αᵐᵒᵖ) := op_bijective.surjective +#align mul_opposite.op_surjective MulOpposite.op_surjective +#align add_opposite.op_surjective AddOpposite.op_surjective @[to_additive] theorem unop_injective : Injective (unop : αᵐᵒᵖ → α) := @@ -132,14 +157,20 @@ theorem unop_injective : Injective (unop : αᵐᵒᵖ → α) := @[to_additive] theorem unop_surjective : Surjective (unop : αᵐᵒᵖ → α) := unop_bijective.surjective +#align mul_opposite.unop_surjective MulOpposite.unop_surjective +#align add_opposite.unop_surjective AddOpposite.unop_surjective -- porting note: `simp` can prove this @[to_additive] theorem op_inj {x y : α} : op x = op y ↔ x = y := by simp +#align mul_opposite.op_inj MulOpposite.op_inj +#align add_opposite.op_inj AddOpposite.op_inj @[to_additive (attr := simp, nolint simpComm)] theorem unop_inj {x y : αᵐᵒᵖ} : unop x = unop y ↔ x = y := unop_injective.eq_iff +#align mul_opposite.unop_inj MulOpposite.unop_inj +#align add_opposite.unop_inj AddOpposite.unop_inj attribute [nolint simpComm] AddOpposite.unop_inj @@ -207,10 +238,14 @@ theorem unop_zero [Zero α] : unop (0 : αᵐᵒᵖ) = 0 := @[to_additive (attr := simp)] theorem op_one [One α] : op (1 : α) = 1 := rfl +#align mul_opposite.op_one MulOpposite.op_one +#align add_opposite.op_zero AddOpposite.op_zero @[to_additive (attr := simp)] theorem unop_one [One α] : unop (1 : αᵐᵒᵖ) = 1 := rfl +#align mul_opposite.unop_one MulOpposite.unop_one +#align add_opposite.unop_zero AddOpposite.unop_zero variable {α} @@ -237,19 +272,26 @@ theorem unop_neg [Neg α] (x : αᵐᵒᵖ) : unop (-x) = -unop x := @[to_additive (attr := simp)] theorem op_mul [Mul α] (x y : α) : op (x * y) = op y * op x := rfl +#align mul_opposite.op_mul MulOpposite.op_mul +#align add_opposite.op_add AddOpposite.op_add @[to_additive (attr := simp)] theorem unop_mul [Mul α] (x y : αᵐᵒᵖ) : unop (x * y) = unop y * unop x := rfl +#align mul_opposite.unop_mul MulOpposite.unop_mul +#align add_opposite.unop_add AddOpposite.unop_add @[to_additive (attr := simp)] theorem op_inv [Inv α] (x : α) : op x⁻¹ = (op x)⁻¹ := rfl +#align mul_opposite.op_inv MulOpposite.op_inv +#align add_opposite.op_neg AddOpposite.op_neg @[to_additive (attr := simp)] theorem unop_inv [Inv α] (x : αᵐᵒᵖ) : unop x⁻¹ = (unop x)⁻¹ := rfl #align mul_opposite.unop_inv MulOpposite.unop_inv +#align add_opposite.unop_neg AddOpposite.unop_neg @[simp] theorem op_sub [Sub α] (x y : α) : op (x - y) = op x - op y := @@ -264,10 +306,14 @@ theorem unop_sub [Sub α] (x y : αᵐᵒᵖ) : unop (x - y) = unop x - unop y : @[to_additive (attr := simp)] theorem op_smul {R : Type _} [SMul R α] (c : R) (a : α) : op (c • a) = c • op a := rfl +#align mul_opposite.op_smul MulOpposite.op_smul +#align add_opposite.op_vadd AddOpposite.op_vadd @[to_additive (attr := simp)] theorem unop_smul {R : Type _} [SMul R α] (c : R) (a : αᵐᵒᵖ) : unop (c • a) = c • unop a := rfl +#align mul_opposite.unop_smul MulOpposite.unop_smul +#align add_opposite.unop_vadd AddOpposite.unop_vadd end @@ -294,12 +340,16 @@ theorem op_ne_zero_iff [Zero α] (a : α) : op a ≠ (0 : αᵐᵒᵖ) ↔ a ≠ @[to_additive (attr := simp, nolint simpComm)] theorem unop_eq_one_iff [One α] (a : αᵐᵒᵖ) : a.unop = 1 ↔ a = 1 := unop_injective.eq_iff' rfl +#align mul_opposite.unop_eq_one_iff MulOpposite.unop_eq_one_iff +#align add_opposite.unop_eq_zero_iff AddOpposite.unop_eq_zero_iff attribute [nolint simpComm] AddOpposite.unop_eq_zero_iff @[to_additive (attr := simp)] theorem op_eq_one_iff [One α] (a : α) : op a = 1 ↔ a = 1 := op_injective.eq_iff' rfl +#align mul_opposite.op_eq_one_iff MulOpposite.op_eq_one_iff +#align add_opposite.op_eq_zero_iff AddOpposite.op_eq_zero_iff end MulOpposite diff --git a/Mathlib/Algebra/Order/Field/Basic.lean b/Mathlib/Algebra/Order/Field/Basic.lean index e4588c094aaf4..95753aaa374bb 100644 --- a/Mathlib/Algebra/Order/Field/Basic.lean +++ b/Mathlib/Algebra/Order/Field/Basic.lean @@ -31,12 +31,16 @@ variable [LinearOrderedSemifield α] {a b c d e : α} {m n : ℤ} def OrderIso.mulLeft₀ (a : α) (ha : 0 < a) : α ≃o α := { Equiv.mulLeft₀ a ha.ne' with map_rel_iff' := @fun _ _ => mul_le_mul_left ha } #align order_iso.mul_left₀ OrderIso.mulLeft₀ +#align order_iso.mul_left₀_symm_apply OrderIso.mulLeft₀_symmApply +#align order_iso.mul_left₀_apply OrderIso.mulLeft₀_apply /-- `Equiv.mulRight₀` as an order_iso. -/ @[simps (config := { simpRhs := true })] def OrderIso.mulRight₀ (a : α) (ha : 0 < a) : α ≃o α := { Equiv.mulRight₀ a ha.ne' with map_rel_iff' := @fun _ _ => mul_le_mul_right ha } #align order_iso.mul_right₀ OrderIso.mulRight₀ +#align order_iso.mul_right₀_symm_apply OrderIso.mulRight₀_symmApply +#align order_iso.mul_right₀_apply OrderIso.mulRight₀_apply /-! ### Lemmas about pos, nonneg, nonpos, neg @@ -50,6 +54,7 @@ theorem inv_pos : 0 < a⁻¹ ↔ 0 < a := #align inv_pos inv_pos alias inv_pos ↔ _ inv_pos_of_pos +#align inv_pos_of_pos inv_pos_of_pos @[simp] theorem inv_nonneg : 0 ≤ a⁻¹ ↔ 0 ≤ a := by @@ -57,6 +62,7 @@ theorem inv_nonneg : 0 ≤ a⁻¹ ↔ 0 ≤ a := by #align inv_nonneg inv_nonneg alias inv_nonneg ↔ _ inv_nonneg_of_nonneg +#align inv_nonneg_of_nonneg inv_nonneg_of_nonneg @[simp] theorem inv_lt_zero : a⁻¹ < 0 ↔ a < 0 := by simp only [← not_le, inv_nonneg] @@ -939,9 +945,13 @@ theorem mul_sub_mul_div_mul_nonpos_iff (hc : c ≠ 0) (hd : d ≠ 0) : #align mul_sub_mul_div_mul_nonpos_iff mul_sub_mul_div_mul_nonpos_iff alias mul_sub_mul_div_mul_neg_iff ↔ div_lt_div_of_mul_sub_mul_div_neg mul_sub_mul_div_mul_neg +#align mul_sub_mul_div_mul_neg mul_sub_mul_div_mul_neg +#align div_lt_div_of_mul_sub_mul_div_neg div_lt_div_of_mul_sub_mul_div_neg alias mul_sub_mul_div_mul_nonpos_iff ↔ div_le_div_of_mul_sub_mul_div_nonpos mul_sub_mul_div_mul_nonpos +#align div_le_div_of_mul_sub_mul_div_nonpos div_le_div_of_mul_sub_mul_div_nonpos +#align mul_sub_mul_div_mul_nonpos mul_sub_mul_div_mul_nonpos theorem exists_add_lt_and_pos_of_lt (h : b < a) : ∃ c, b + c < a ∧ 0 < c := ⟨(a - b) / 2, add_sub_div_two_lt h, div_pos (sub_pos_of_lt h) zero_lt_two⟩ diff --git a/Mathlib/Algebra/Order/Group/Defs.lean b/Mathlib/Algebra/Order/Group/Defs.lean index f8a5e432f3717..9bea0327e478e 100644 --- a/Mathlib/Algebra/Order/Group/Defs.lean +++ b/Mathlib/Algebra/Order/Group/Defs.lean @@ -343,6 +343,7 @@ theorem inv_le_inv_iff : a⁻¹ ≤ b⁻¹ ↔ b ≤ a := by #align neg_le_neg_iff neg_le_neg_iff alias neg_le_neg_iff ↔ le_of_neg_le_neg _ +#align le_of_neg_le_neg le_of_neg_le_neg @[to_additive] theorem mul_inv_le_inv_mul_iff : a * b⁻¹ ≤ d⁻¹ * c ↔ d * a ≤ c * b := by @@ -366,6 +367,7 @@ theorem le_div_self_iff (a : α) {b : α} : a ≤ a / b ↔ b ≤ 1 := by #align le_sub_self_iff le_sub_self_iff alias sub_le_self_iff ↔ _ sub_le_self +#align sub_le_self sub_le_self end TypeclassesLeftRightLE @@ -392,12 +394,16 @@ theorem lt_inv' : a < b⁻¹ ↔ b < a⁻¹ := by rw [← inv_lt_inv_iff, inv_in #align lt_neg lt_neg alias lt_inv' ↔ lt_inv_of_lt_inv _ +#align lt_inv_of_lt_inv lt_inv_of_lt_inv attribute [to_additive] lt_inv_of_lt_inv +#align lt_neg_of_lt_neg lt_neg_of_lt_neg alias inv_lt' ↔ inv_lt_of_inv_lt' _ +#align inv_lt_of_inv_lt' inv_lt_of_inv_lt' attribute [to_additive neg_lt_of_neg_lt] inv_lt_of_inv_lt' +#align neg_lt_of_neg_lt neg_lt_of_neg_lt @[to_additive] theorem mul_inv_lt_inv_mul_iff : a * b⁻¹ < d⁻¹ * c ↔ d * a < c * b := by @@ -414,6 +420,7 @@ theorem div_lt_self_iff (a : α) {b : α} : a / b < a ↔ 1 < b := by #align sub_lt_self_iff sub_lt_self_iff alias sub_lt_self_iff ↔ _ sub_lt_self +#align sub_lt_self sub_lt_self end TypeclassesLeftRightLT @@ -432,6 +439,7 @@ theorem Left.inv_le_self (h : 1 ≤ a) : a⁻¹ ≤ a := #align left.neg_le_self Left.neg_le_self alias Left.neg_le_self ← neg_le_self +#align neg_le_self neg_le_self @[to_additive] theorem Left.self_le_inv (h : a ≤ 1) : a ≤ a⁻¹ := @@ -452,6 +460,7 @@ theorem Left.inv_lt_self (h : 1 < a) : a⁻¹ < a := #align left.neg_lt_self Left.neg_lt_self alias Left.neg_lt_self ← neg_lt_self +#align neg_lt_self neg_lt_self @[to_additive] theorem Left.self_lt_inv (h : a < 1) : a < a⁻¹ := @@ -556,91 +565,134 @@ end LT end CommGroup alias Left.inv_le_one_iff ↔ one_le_of_inv_le_one _ +#align one_le_of_inv_le_one one_le_of_inv_le_one attribute [to_additive] one_le_of_inv_le_one +#align nonneg_of_neg_nonpos nonneg_of_neg_nonpos alias Left.one_le_inv_iff ↔ le_one_of_one_le_inv _ +#align le_one_of_one_le_inv le_one_of_one_le_inv attribute [to_additive nonpos_of_neg_nonneg] le_one_of_one_le_inv +#align nonpos_of_neg_nonneg nonpos_of_neg_nonneg alias inv_lt_inv_iff ↔ lt_of_inv_lt_inv _ +#align lt_of_inv_lt_inv lt_of_inv_lt_inv attribute [to_additive] lt_of_inv_lt_inv +#align lt_of_neg_lt_neg lt_of_neg_lt_neg alias Left.inv_lt_one_iff ↔ one_lt_of_inv_lt_one _ +#align one_lt_of_inv_lt_one one_lt_of_inv_lt_one attribute [to_additive] one_lt_of_inv_lt_one +#align pos_of_neg_neg pos_of_neg_neg alias Left.inv_lt_one_iff ← inv_lt_one_iff_one_lt +#align inv_lt_one_iff_one_lt inv_lt_one_iff_one_lt attribute [to_additive] inv_lt_one_iff_one_lt +#align neg_neg_iff_pos neg_neg_iff_pos alias Left.inv_lt_one_iff ← inv_lt_one' +#align inv_lt_one' inv_lt_one' attribute [to_additive neg_lt_zero] inv_lt_one' +#align neg_lt_zero neg_lt_zero alias Left.one_lt_inv_iff ↔ inv_of_one_lt_inv _ +#align inv_of_one_lt_inv inv_of_one_lt_inv attribute [to_additive neg_of_neg_pos] inv_of_one_lt_inv +#align neg_of_neg_pos neg_of_neg_pos alias Left.one_lt_inv_iff ↔ _ one_lt_inv_of_inv +#align one_lt_inv_of_inv one_lt_inv_of_inv attribute [to_additive neg_pos_of_neg] one_lt_inv_of_inv +#align neg_pos_of_neg neg_pos_of_neg alias le_inv_mul_iff_mul_le ↔ mul_le_of_le_inv_mul _ +#align mul_le_of_le_inv_mul mul_le_of_le_inv_mul attribute [to_additive] mul_le_of_le_inv_mul +#align add_le_of_le_neg_add add_le_of_le_neg_add alias le_inv_mul_iff_mul_le ↔ _ le_inv_mul_of_mul_le +#align le_inv_mul_of_mul_le le_inv_mul_of_mul_le attribute [to_additive] le_inv_mul_of_mul_le +#align le_neg_add_of_add_le le_neg_add_of_add_le alias inv_mul_le_iff_le_mul ↔ _ inv_mul_le_of_le_mul +#align inv_mul_le_of_le_mul inv_mul_le_of_le_mul -- Porting note: was `inv_mul_le_iff_le_mul` attribute [to_additive] inv_mul_le_of_le_mul alias lt_inv_mul_iff_mul_lt ↔ mul_lt_of_lt_inv_mul _ +#align mul_lt_of_lt_inv_mul mul_lt_of_lt_inv_mul attribute [to_additive] mul_lt_of_lt_inv_mul +#align add_lt_of_lt_neg_add add_lt_of_lt_neg_add alias lt_inv_mul_iff_mul_lt ↔ _ lt_inv_mul_of_mul_lt +#align lt_inv_mul_of_mul_lt lt_inv_mul_of_mul_lt attribute [to_additive] lt_inv_mul_of_mul_lt +#align lt_neg_add_of_add_lt lt_neg_add_of_add_lt alias inv_mul_lt_iff_lt_mul ↔ lt_mul_of_inv_mul_lt inv_mul_lt_of_lt_mul +#align lt_mul_of_inv_mul_lt lt_mul_of_inv_mul_lt +#align inv_mul_lt_of_lt_mul inv_mul_lt_of_lt_mul attribute [to_additive] lt_mul_of_inv_mul_lt +#align lt_add_of_neg_add_lt lt_add_of_neg_add_lt attribute [to_additive] inv_mul_lt_of_lt_mul +#align neg_add_lt_of_lt_add neg_add_lt_of_lt_add alias lt_mul_of_inv_mul_lt ← lt_mul_of_inv_mul_lt_left +#align lt_mul_of_inv_mul_lt_left lt_mul_of_inv_mul_lt_left attribute [to_additive] lt_mul_of_inv_mul_lt_left +#align lt_add_of_neg_add_lt_left lt_add_of_neg_add_lt_left alias Left.inv_le_one_iff ← inv_le_one' +#align inv_le_one' inv_le_one' attribute [to_additive neg_nonpos] inv_le_one' +#align neg_nonpos neg_nonpos alias Left.one_le_inv_iff ← one_le_inv' +#align one_le_inv' one_le_inv' attribute [to_additive neg_nonneg] one_le_inv' +#align neg_nonneg neg_nonneg alias Left.one_lt_inv_iff ← one_lt_inv' +#align one_lt_inv' one_lt_inv' attribute [to_additive neg_pos] one_lt_inv' +#align neg_pos neg_pos alias mul_lt_mul_left' ← OrderedCommGroup.mul_lt_mul_left' +#align ordered_comm_group.mul_lt_mul_left' OrderedCommGroup.mul_lt_mul_left' attribute [to_additive OrderedAddCommGroup.add_lt_add_left] OrderedCommGroup.mul_lt_mul_left' +#align ordered_add_comm_group.add_lt_add_left OrderedAddCommGroup.add_lt_add_left alias le_of_mul_le_mul_left' ← OrderedCommGroup.le_of_mul_le_mul_left +#align ordered_comm_group.le_of_mul_le_mul_left OrderedCommGroup.le_of_mul_le_mul_left attribute [to_additive] OrderedCommGroup.le_of_mul_le_mul_left +#align ordered_add_comm_group.le_of_add_le_add_left OrderedAddCommGroup.le_of_add_le_add_left alias lt_of_mul_lt_mul_left' ← OrderedCommGroup.lt_of_mul_lt_mul_left +#align ordered_comm_group.lt_of_mul_lt_mul_left OrderedCommGroup.lt_of_mul_lt_mul_left attribute [to_additive] OrderedCommGroup.lt_of_mul_lt_mul_left +#align ordered_add_comm_group.lt_of_add_lt_add_left OrderedAddCommGroup.lt_of_add_lt_add_left -- Most of the lemmas that are primed in this section appear in ordered_field. -- I (DT) did not try to minimise the assumptions. @@ -671,6 +723,8 @@ theorem one_le_div' : 1 ≤ a / b ↔ b ≤ a := by #align sub_nonneg sub_nonneg alias sub_nonneg ↔ le_of_sub_nonneg sub_nonneg_of_le +#align sub_nonneg_of_le sub_nonneg_of_le +#align le_of_sub_nonneg le_of_sub_nonneg @[to_additive (attr := simp) sub_nonpos] theorem div_le_one' : a / b ≤ 1 ↔ a ≤ b := by @@ -679,6 +733,8 @@ theorem div_le_one' : a / b ≤ 1 ↔ a ≤ b := by #align sub_nonpos sub_nonpos alias sub_nonpos ↔ le_of_sub_nonpos sub_nonpos_of_le +#align sub_nonpos_of_le sub_nonpos_of_le +#align le_of_sub_nonpos le_of_sub_nonpos @[to_additive] theorem le_div_iff_mul_le : a ≤ c / b ↔ a * b ≤ c := by @@ -687,6 +743,8 @@ theorem le_div_iff_mul_le : a ≤ c / b ↔ a * b ≤ c := by #align le_sub_iff_add_le le_sub_iff_add_le alias le_sub_iff_add_le ↔ add_le_of_le_sub_right le_sub_right_of_add_le +#align add_le_of_le_sub_right add_le_of_le_sub_right +#align le_sub_right_of_add_le le_sub_right_of_add_le @[to_additive] theorem div_le_iff_le_mul : a / c ≤ b ↔ a ≤ b * c := by @@ -747,6 +805,8 @@ theorem le_div_iff_mul_le' : b ≤ c / a ↔ a * b ≤ c := by rw [le_div_iff_mu #align le_sub_iff_add_le' le_sub_iff_add_le' alias le_sub_iff_add_le' ↔ add_le_of_le_sub_left le_sub_left_of_add_le +#align le_sub_left_of_add_le le_sub_left_of_add_le +#align add_le_of_le_sub_left add_le_of_le_sub_left @[to_additive] theorem div_le_iff_le_mul' : a / b ≤ c ↔ a ≤ b * c := by rw [div_le_iff_le_mul, mul_comm] @@ -754,6 +814,8 @@ theorem div_le_iff_le_mul' : a / b ≤ c ↔ a ≤ b * c := by rw [div_le_iff_le #align sub_le_iff_le_add' sub_le_iff_le_add' alias sub_le_iff_le_add' ↔ le_add_of_sub_left_le sub_left_le_of_le_add +#align sub_left_le_of_le_add sub_left_le_of_le_add +#align le_add_of_sub_left_le le_add_of_sub_left_le @[to_additive (attr := simp)] theorem inv_le_div_iff_le_mul : b⁻¹ ≤ a / c ↔ c ≤ a * b := @@ -824,6 +886,8 @@ theorem one_lt_div' : 1 < a / b ↔ b < a := by #align sub_pos sub_pos alias sub_pos ↔ lt_of_sub_pos sub_pos_of_lt +#align lt_of_sub_pos lt_of_sub_pos +#align sub_pos_of_lt sub_pos_of_lt @[to_additive (attr := simp) sub_neg] theorem div_lt_one' : a / b < 1 ↔ a < b := by @@ -832,8 +896,11 @@ theorem div_lt_one' : a / b < 1 ↔ a < b := by #align sub_neg sub_neg alias sub_neg ↔ lt_of_sub_neg sub_neg_of_lt +#align lt_of_sub_neg lt_of_sub_neg +#align sub_neg_of_lt sub_neg_of_lt alias sub_neg ← sub_lt_zero +#align sub_lt_zero sub_lt_zero @[to_additive] theorem lt_div_iff_mul_lt : a < c / b ↔ a * b < c := by @@ -842,6 +909,8 @@ theorem lt_div_iff_mul_lt : a < c / b ↔ a * b < c := by #align lt_sub_iff_add_lt lt_sub_iff_add_lt alias lt_sub_iff_add_lt ↔ add_lt_of_lt_sub_right lt_sub_right_of_add_lt +#align add_lt_of_lt_sub_right add_lt_of_lt_sub_right +#align lt_sub_right_of_add_lt lt_sub_right_of_add_lt @[to_additive] theorem div_lt_iff_lt_mul : a / c < b ↔ a < b * c := by @@ -850,6 +919,8 @@ theorem div_lt_iff_lt_mul : a / c < b ↔ a < b * c := by #align sub_lt_iff_lt_add sub_lt_iff_lt_add alias sub_lt_iff_lt_add ↔ lt_add_of_sub_right_lt sub_right_lt_of_lt_add +#align lt_add_of_sub_right_lt lt_add_of_sub_right_lt +#align sub_right_lt_of_lt_add sub_right_lt_of_lt_add end Right @@ -901,6 +972,8 @@ theorem lt_div_iff_mul_lt' : b < c / a ↔ a * b < c := by rw [lt_div_iff_mul_lt #align lt_sub_iff_add_lt' lt_sub_iff_add_lt' alias lt_sub_iff_add_lt' ↔ add_lt_of_lt_sub_left lt_sub_left_of_add_lt +#align lt_sub_left_of_add_lt lt_sub_left_of_add_lt +#align add_lt_of_lt_sub_left add_lt_of_lt_sub_left @[to_additive] theorem div_lt_iff_lt_mul' : a / b < c ↔ a < b * c := by rw [div_lt_iff_lt_mul, mul_comm] @@ -908,6 +981,8 @@ theorem div_lt_iff_lt_mul' : a / b < c ↔ a < b * c := by rw [div_lt_iff_lt_mul #align sub_lt_iff_lt_add' sub_lt_iff_lt_add' alias sub_lt_iff_lt_add' ↔ lt_add_of_sub_left_lt sub_left_lt_of_lt_add +#align lt_add_of_sub_left_lt lt_add_of_sub_left_lt +#align sub_left_lt_of_lt_add sub_left_lt_of_lt_add @[to_additive] theorem inv_lt_div_iff_lt_mul' : b⁻¹ < a / c ↔ c < a * b := @@ -1112,6 +1187,7 @@ structure TotalPositiveCone (α : Type _) [AddCommGroup α] extends PositiveCone /-- Forget that a `TotalPositiveCone` is total. -/ add_decl_doc TotalPositiveCone.toPositiveCone +#align add_comm_group.total_positive_cone.to_positive_cone AddCommGroup.TotalPositiveCone.toPositiveCone end AddCommGroup diff --git a/Mathlib/Algebra/Order/Group/MinMax.lean b/Mathlib/Algebra/Order/Group/MinMax.lean index 57e621f119b60..77d7e612aaf14 100644 --- a/Mathlib/Algebra/Order/Group/MinMax.lean +++ b/Mathlib/Algebra/Order/Group/MinMax.lean @@ -27,6 +27,7 @@ theorem max_one_div_max_inv_one_eq_self (a : α) : max a 1 / max a⁻¹ 1 = a := #align max_zero_sub_max_neg_zero_eq_self max_zero_sub_max_neg_zero_eq_self alias max_zero_sub_max_neg_zero_eq_self ← max_zero_sub_eq_self +#align max_zero_sub_eq_self max_zero_sub_eq_self end diff --git a/Mathlib/Algebra/Order/Group/OrderIso.lean b/Mathlib/Algebra/Order/Group/OrderIso.lean index 85152c34a0dba..517898d866bb4 100644 --- a/Mathlib/Algebra/Order/Group/OrderIso.lean +++ b/Mathlib/Algebra/Order/Group/OrderIso.lean @@ -43,6 +43,8 @@ def OrderIso.inv : α ≃o αᵒᵈ where map_rel_iff' {_ _} := @inv_le_inv_iff α _ _ _ _ _ _ #align order_iso.inv OrderIso.inv #align order_iso.neg OrderIso.neg +#align order_iso.inv_apply OrderIso.inv_apply +#align order_iso.inv_symm_apply OrderIso.inv_symmApply end @@ -53,8 +55,10 @@ theorem inv_le' : a⁻¹ ≤ b ↔ b⁻¹ ≤ a := #align neg_le neg_le alias inv_le' ↔ inv_le_of_inv_le' _ +#align inv_le_of_inv_le' inv_le_of_inv_le' attribute [to_additive neg_le_of_neg_le] inv_le_of_inv_le' +#align neg_le_of_neg_le neg_le_of_neg_le @[to_additive le_neg] theorem le_inv' : a ≤ b⁻¹ ↔ b ≤ a⁻¹ := @@ -67,8 +71,10 @@ end TypeclassesLeftRightLE end Group alias le_inv' ↔ le_inv_of_le_inv _ +#align le_inv_of_le_inv le_inv_of_le_inv attribute [to_additive] le_inv_of_le_inv +#align le_neg_of_le_neg le_neg_of_le_neg section Group @@ -86,6 +92,8 @@ def OrderIso.mulRight (a : α) : α ≃o α where toEquiv := Equiv.mulRight a #align order_iso.mul_right OrderIso.mulRight #align order_iso.add_right OrderIso.addRight +#align order_iso.mul_right_apply OrderIso.mulRight_apply +#align order_iso.mul_right_to_equiv OrderIso.mulRight_toEquiv @[to_additive (attr := simp)] theorem OrderIso.mulRight_symm (a : α) : (OrderIso.mulRight a).symm = OrderIso.mulRight a⁻¹ := by @@ -108,6 +116,8 @@ def OrderIso.mulLeft (a : α) : α ≃o α where toEquiv := Equiv.mulLeft a #align order_iso.mul_left OrderIso.mulLeft #align order_iso.add_left OrderIso.addLeft +#align order_iso.mul_left_apply OrderIso.mulLeft_apply +#align order_iso.mul_left_to_equiv OrderIso.mulLeft_toEquiv @[to_additive (attr := simp)] theorem OrderIso.mulLeft_symm (a : α) : (OrderIso.mulLeft a).symm = OrderIso.mulLeft a⁻¹ := by diff --git a/Mathlib/Algebra/Order/Hom/Monoid.lean b/Mathlib/Algebra/Order/Hom/Monoid.lean index 913638b223ec1..30afaad7a080f 100644 --- a/Mathlib/Algebra/Order/Hom/Monoid.lean +++ b/Mathlib/Algebra/Order/Hom/Monoid.lean @@ -422,6 +422,7 @@ theorem coe_comp (f : β →*o γ) (g : α →*o β) : (f.comp g : α → γ) = theorem comp_apply (f : β →*o γ) (g : α →*o β) (a : α) : (f.comp g) a = f (g a) := rfl #align order_add_monoid_hom.comp_apply OrderAddMonoidHom.comp_apply +#align order_monoid_hom.comp_apply OrderMonoidHom.comp_apply @[to_additive] theorem coe_comp_monoidHom (f : β →*o γ) (g : α →*o β) : diff --git a/Mathlib/Algebra/Order/Monoid/Basic.lean b/Mathlib/Algebra/Order/Monoid/Basic.lean index 421f34ebb69f4..b1588639c3f30 100644 --- a/Mathlib/Algebra/Order/Monoid/Basic.lean +++ b/Mathlib/Algebra/Order/Monoid/Basic.lean @@ -68,6 +68,7 @@ def OrderEmbedding.mulLeft {α : Type _} [Mul α] [LinearOrder α] OrderEmbedding.ofStrictMono (fun n => m * n) fun _ _ w => mul_lt_mul_left' w m #align order_embedding.mul_left OrderEmbedding.mulLeft #align order_embedding.add_left OrderEmbedding.addLeft +#align order_embedding.mul_left_apply OrderEmbedding.mulLeft_apply /-- The order embedding sending `b` to `b * a`, for some fixed `a`. See also `OrderIso.mulRight` when working in an ordered group. -/ @@ -80,3 +81,4 @@ def OrderEmbedding.mulRight {α : Type _} [Mul α] [LinearOrder α] OrderEmbedding.ofStrictMono (fun n => n * m) fun _ _ w => mul_lt_mul_right' w m #align order_embedding.mul_right OrderEmbedding.mulRight #align order_embedding.add_right OrderEmbedding.addRight +#align order_embedding.mul_right_apply OrderEmbedding.mulRight_apply diff --git a/Mathlib/Algebra/Order/Monoid/Lemmas.lean b/Mathlib/Algebra/Order/Monoid/Lemmas.lean index 4c6d4ee4df13c..2986d5672644d 100644 --- a/Mathlib/Algebra/Order/Monoid/Lemmas.lean +++ b/Mathlib/Algebra/Order/Monoid/Lemmas.lean @@ -55,12 +55,16 @@ is taken by the analogous lemma for semiring, with an extra non-negativity assum theorem mul_le_mul_left' [CovariantClass α α (· * ·) (· ≤ ·)] {b c : α} (bc : b ≤ c) (a : α) : a * b ≤ a * c := CovariantClass.elim _ bc +#align mul_le_mul_left' mul_le_mul_left' +#align add_le_add_left add_le_add_left @[to_additive le_of_add_le_add_left] theorem le_of_mul_le_mul_left' [ContravariantClass α α (· * ·) (· ≤ ·)] {a b c : α} (bc : a * b ≤ a * c) : b ≤ c := ContravariantClass.elim _ bc +#align le_of_mul_le_mul_left' le_of_mul_le_mul_left' +#align le_of_add_le_add_left le_of_add_le_add_left /- The prime on this lemma is present only on the multiplicative version. The unprimed version is taken by the analogous lemma for semiring, with an extra non-negativity assumption. -/ @@ -69,24 +73,32 @@ theorem mul_le_mul_right' [i : CovariantClass α α (swap (· * ·)) (· ≤ ·) (a : α) : b * a ≤ c * a := i.elim a bc +#align mul_le_mul_right' mul_le_mul_right' +#align add_le_add_right add_le_add_right @[to_additive le_of_add_le_add_right] theorem le_of_mul_le_mul_right' [i : ContravariantClass α α (swap (· * ·)) (· ≤ ·)] {a b c : α} (bc : b * a ≤ c * a) : b ≤ c := i.elim a bc +#align le_of_mul_le_mul_right' le_of_mul_le_mul_right' +#align le_of_add_le_add_right le_of_add_le_add_right @[to_additive (attr := simp)] theorem mul_le_mul_iff_left [CovariantClass α α (· * ·) (· ≤ ·)] [ContravariantClass α α (· * ·) (· ≤ ·)] (a : α) {b c : α} : a * b ≤ a * c ↔ b ≤ c := rel_iff_cov α α (· * ·) (· ≤ ·) a +#align mul_le_mul_iff_left mul_le_mul_iff_left +#align add_le_add_iff_left add_le_add_iff_left @[to_additive (attr := simp)] theorem mul_le_mul_iff_right [CovariantClass α α (swap (· * ·)) (· ≤ ·)] [ContravariantClass α α (swap (· * ·)) (· ≤ ·)] (a : α) {b c : α} : b * a ≤ c * a ↔ b ≤ c := rel_iff_cov α α (swap (· * ·)) (· ≤ ·) a +#align mul_le_mul_iff_right mul_le_mul_iff_right +#align add_le_add_iff_right add_le_add_iff_right end LE @@ -99,35 +111,47 @@ theorem mul_lt_mul_iff_left [CovariantClass α α (· * ·) (· < ·)] [ContravariantClass α α (· * ·) (· < ·)] (a : α) {b c : α} : a * b < a * c ↔ b < c := rel_iff_cov α α (· * ·) (· < ·) a +#align mul_lt_mul_iff_left mul_lt_mul_iff_left +#align add_lt_add_iff_left add_lt_add_iff_left @[to_additive (attr := simp)] theorem mul_lt_mul_iff_right [CovariantClass α α (swap (· * ·)) (· < ·)] [ContravariantClass α α (swap (· * ·)) (· < ·)] (a : α) {b c : α} : b * a < c * a ↔ b < c := rel_iff_cov α α (swap (· * ·)) (· < ·) a +#align mul_lt_mul_iff_right mul_lt_mul_iff_right +#align add_lt_add_iff_right add_lt_add_iff_right @[to_additive add_lt_add_left] theorem mul_lt_mul_left' [CovariantClass α α (· * ·) (· < ·)] {b c : α} (bc : b < c) (a : α) : a * b < a * c := CovariantClass.elim _ bc +#align mul_lt_mul_left' mul_lt_mul_left' +#align add_lt_add_left add_lt_add_left @[to_additive lt_of_add_lt_add_left] theorem lt_of_mul_lt_mul_left' [ContravariantClass α α (· * ·) (· < ·)] {a b c : α} (bc : a * b < a * c) : b < c := ContravariantClass.elim _ bc +#align lt_of_mul_lt_mul_left' lt_of_mul_lt_mul_left' +#align lt_of_add_lt_add_left lt_of_add_lt_add_left @[to_additive add_lt_add_right] theorem mul_lt_mul_right' [i : CovariantClass α α (swap (· * ·)) (· < ·)] {b c : α} (bc : b < c) (a : α) : b * a < c * a := i.elim a bc +#align mul_lt_mul_right' mul_lt_mul_right' +#align add_lt_add_right add_lt_add_right @[to_additive lt_of_add_lt_add_right] theorem lt_of_mul_lt_mul_right' [i : ContravariantClass α α (swap (· * ·)) (· < ·)] {a b c : α} (bc : b * a < c * a) : b < c := i.elim a bc +#align lt_of_mul_lt_mul_right' lt_of_mul_lt_mul_right' +#align lt_of_add_lt_add_right lt_of_add_lt_add_right end LT @@ -142,20 +166,27 @@ theorem mul_lt_mul_of_lt_of_lt [CovariantClass α α (· * ·) (· < ·)] calc a * c < a * d := mul_lt_mul_left' h₂ a _ < b * d := mul_lt_mul_right' h₁ d +#align mul_lt_mul_of_lt_of_lt mul_lt_mul_of_lt_of_lt +#align add_lt_add_of_lt_of_lt add_lt_add_of_lt_of_lt alias add_lt_add_of_lt_of_lt ← add_lt_add +#align add_lt_add add_lt_add @[to_additive] theorem mul_lt_mul_of_le_of_lt [CovariantClass α α (· * ·) (· < ·)] [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b c d : α} (h₁ : a ≤ b) (h₂ : c < d) : a * c < b * d := (mul_le_mul_right' h₁ _).trans_lt (mul_lt_mul_left' h₂ b) +#align mul_lt_mul_of_le_of_lt mul_lt_mul_of_le_of_lt +#align add_lt_add_of_le_of_lt add_lt_add_of_le_of_lt @[to_additive] theorem mul_lt_mul_of_lt_of_le [CovariantClass α α (· * ·) (· ≤ ·)] [CovariantClass α α (swap (· * ·)) (· < ·)] {a b c d : α} (h₁ : a < b) (h₂ : c ≤ d) : a * c < b * d := (mul_le_mul_left' h₂ _).trans_lt (mul_lt_mul_right' h₁ d) +#align mul_lt_mul_of_lt_of_le mul_lt_mul_of_lt_of_le +#align add_lt_add_of_lt_of_le add_lt_add_of_lt_of_le /-- Only assumes left strict covariance. -/ @[to_additive "Only assumes left strict covariance"] @@ -181,6 +212,8 @@ theorem mul_le_mul' [CovariantClass α α (· * ·) (· ≤ ·)] [CovariantClass {a b c d : α} (h₁ : a ≤ b) (h₂ : c ≤ d) : a * c ≤ b * d := (mul_le_mul_left' h₂ _).trans (mul_le_mul_right' h₁ d) +#align mul_le_mul' mul_le_mul' +#align add_le_add add_le_add @[to_additive] theorem mul_le_mul_three [CovariantClass α α (· * ·) (· ≤ ·)] @@ -188,54 +221,72 @@ theorem mul_le_mul_three [CovariantClass α α (· * ·) (· ≤ ·)] (h₃ : c ≤ f) : a * b * c ≤ d * e * f := mul_le_mul' (mul_le_mul' h₁ h₂) h₃ +#align mul_le_mul_three mul_le_mul_three +#align add_le_add_three add_le_add_three @[to_additive] theorem mul_lt_of_mul_lt_left [CovariantClass α α (· * ·) (· ≤ ·)] {a b c d : α} (h : a * b < c) (hle : d ≤ b) : a * d < c := (mul_le_mul_left' hle a).trans_lt h +#align mul_lt_of_mul_lt_left mul_lt_of_mul_lt_left +#align add_lt_of_add_lt_left add_lt_of_add_lt_left @[to_additive] theorem mul_le_of_mul_le_left [CovariantClass α α (· * ·) (· ≤ ·)] {a b c d : α} (h : a * b ≤ c) (hle : d ≤ b) : a * d ≤ c := @act_rel_of_rel_of_act_rel _ _ _ (· ≤ ·) _ _ a _ _ _ hle h +#align mul_le_of_mul_le_left mul_le_of_mul_le_left +#align add_le_of_add_le_left add_le_of_add_le_left @[to_additive] theorem mul_lt_of_mul_lt_right [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b c d : α} (h : a * b < c) (hle : d ≤ a) : d * b < c := (mul_le_mul_right' hle b).trans_lt h +#align mul_lt_of_mul_lt_right mul_lt_of_mul_lt_right +#align add_lt_of_add_lt_right add_lt_of_add_lt_right @[to_additive] theorem mul_le_of_mul_le_right [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b c d : α} (h : a * b ≤ c) (hle : d ≤ a) : d * b ≤ c := (mul_le_mul_right' hle b).trans h +#align mul_le_of_mul_le_right mul_le_of_mul_le_right +#align add_le_of_add_le_right add_le_of_add_le_right @[to_additive] theorem lt_mul_of_lt_mul_left [CovariantClass α α (· * ·) (· ≤ ·)] {a b c d : α} (h : a < b * c) (hle : c ≤ d) : a < b * d := h.trans_le (mul_le_mul_left' hle b) +#align lt_mul_of_lt_mul_left lt_mul_of_lt_mul_left +#align lt_add_of_lt_add_left lt_add_of_lt_add_left @[to_additive] theorem le_mul_of_le_mul_left [CovariantClass α α (· * ·) (· ≤ ·)] {a b c d : α} (h : a ≤ b * c) (hle : c ≤ d) : a ≤ b * d := @rel_act_of_rel_of_rel_act _ _ _ (· ≤ ·) _ _ b _ _ _ hle h +#align le_mul_of_le_mul_left le_mul_of_le_mul_left +#align le_add_of_le_add_left le_add_of_le_add_left @[to_additive] theorem lt_mul_of_lt_mul_right [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b c d : α} (h : a < b * c) (hle : b ≤ d) : a < d * c := h.trans_le (mul_le_mul_right' hle c) +#align lt_mul_of_lt_mul_right lt_mul_of_lt_mul_right +#align lt_add_of_lt_add_right lt_add_of_lt_add_right @[to_additive] theorem le_mul_of_le_mul_right [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b c d : α} (h : a ≤ b * c) (hle : b ≤ d) : a ≤ d * c := h.trans (mul_le_mul_right' hle c) +#align le_mul_of_le_mul_right le_mul_of_le_mul_right +#align le_add_of_le_add_right le_add_of_le_add_right end Preorder @@ -247,12 +298,16 @@ variable [PartialOrder α] theorem mul_left_cancel'' [ContravariantClass α α (· * ·) (· ≤ ·)] {a b c : α} (h : a * b = a * c) : b = c := (le_of_mul_le_mul_left' h.le).antisymm (le_of_mul_le_mul_left' h.ge) +#align mul_left_cancel'' mul_left_cancel'' +#align add_left_cancel'' add_left_cancel'' @[to_additive] theorem mul_right_cancel'' [ContravariantClass α α (swap (· * ·)) (· ≤ ·)] {a b c : α} (h : a * b = c * b) : a = c := le_antisymm (le_of_mul_le_mul_right' h.le) (le_of_mul_le_mul_right' h.ge) +#align mul_right_cancel'' mul_right_cancel'' +#align add_right_cancel'' add_right_cancel'' end PartialOrder @@ -262,6 +317,8 @@ variable [LinearOrder α] {a b c d : α} [CovariantClass α α (· * ·) (· < @[to_additive] lemma min_le_max_of_mul_le_mul (h : a * b ≤ c * d) : min a b ≤ max c d := by simp_rw [min_le_iff, le_max_iff]; contrapose! h; exact mul_lt_mul_of_lt_of_lt h.1.1 h.2.2 +#align min_le_max_of_add_le_add min_le_max_of_add_le_add +#align min_le_max_of_mul_le_mul min_le_max_of_mul_le_mul end LinearOrder end Mul @@ -281,6 +338,8 @@ theorem le_mul_of_one_le_right' [CovariantClass α α (· * ·) (· ≤ ·)] {a calc a = a * 1 := (mul_one a).symm _ ≤ a * b := mul_le_mul_left' h a +#align le_mul_of_one_le_right' le_mul_of_one_le_right' +#align le_add_of_nonneg_right le_add_of_nonneg_right @[to_additive add_le_of_nonpos_right] theorem mul_le_of_le_one_right' [CovariantClass α α (· * ·) (· ≤ ·)] {a b : α} (h : b ≤ 1) : @@ -288,6 +347,8 @@ theorem mul_le_of_le_one_right' [CovariantClass α α (· * ·) (· ≤ ·)] {a calc a * b ≤ a * 1 := mul_le_mul_left' h a _ = a := mul_one a +#align mul_le_of_le_one_right' mul_le_of_le_one_right' +#align add_le_of_nonpos_right add_le_of_nonpos_right @[to_additive le_add_of_nonneg_left] theorem le_mul_of_one_le_left' [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b : α} (h : 1 ≤ b) : @@ -295,6 +356,8 @@ theorem le_mul_of_one_le_left' [CovariantClass α α (swap (· * ·)) (· ≤ · calc a = 1 * a := (one_mul a).symm _ ≤ b * a := mul_le_mul_right' h a +#align le_mul_of_one_le_left' le_mul_of_one_le_left' +#align le_add_of_nonneg_left le_add_of_nonneg_left @[to_additive add_le_of_nonpos_left] theorem mul_le_of_le_one_left' [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b : α} (h : b ≤ 1) : @@ -302,52 +365,70 @@ theorem mul_le_of_le_one_left' [CovariantClass α α (swap (· * ·)) (· ≤ · calc b * a ≤ 1 * a := mul_le_mul_right' h a _ = a := one_mul a +#align mul_le_of_le_one_left' mul_le_of_le_one_left' +#align add_le_of_nonpos_left add_le_of_nonpos_left @[to_additive] theorem one_le_of_le_mul_right [ContravariantClass α α (· * ·) (· ≤ ·)] {a b : α} (h : a ≤ a * b) : 1 ≤ b := le_of_mul_le_mul_left' <| by simpa only [mul_one] +#align one_le_of_le_mul_right one_le_of_le_mul_right +#align nonneg_of_le_add_right nonneg_of_le_add_right @[to_additive] theorem le_one_of_mul_le_right [ContravariantClass α α (· * ·) (· ≤ ·)] {a b : α} (h : a * b ≤ a) : b ≤ 1 := le_of_mul_le_mul_left' <| by simpa only [mul_one] +#align le_one_of_mul_le_right le_one_of_mul_le_right +#align nonpos_of_add_le_right nonpos_of_add_le_right @[to_additive] theorem one_le_of_le_mul_left [ContravariantClass α α (swap (· * ·)) (· ≤ ·)] {a b : α} (h : b ≤ a * b) : 1 ≤ a := le_of_mul_le_mul_right' <| by simpa only [one_mul] +#align one_le_of_le_mul_left one_le_of_le_mul_left +#align nonneg_of_le_add_left nonneg_of_le_add_left @[to_additive] theorem le_one_of_mul_le_left [ContravariantClass α α (swap (· * ·)) (· ≤ ·)] {a b : α} (h : a * b ≤ b) : a ≤ 1 := le_of_mul_le_mul_right' <| by simpa only [one_mul] +#align le_one_of_mul_le_left le_one_of_mul_le_left +#align nonpos_of_add_le_left nonpos_of_add_le_left @[to_additive (attr := simp) le_add_iff_nonneg_right] theorem le_mul_iff_one_le_right' [CovariantClass α α (· * ·) (· ≤ ·)] [ContravariantClass α α (· * ·) (· ≤ ·)] (a : α) {b : α} : a ≤ a * b ↔ 1 ≤ b := Iff.trans (by rw [mul_one]) (mul_le_mul_iff_left a) +#align le_mul_iff_one_le_right' le_mul_iff_one_le_right' +#align le_add_iff_nonneg_right le_add_iff_nonneg_right @[to_additive (attr := simp) le_add_iff_nonneg_left] theorem le_mul_iff_one_le_left' [CovariantClass α α (swap (· * ·)) (· ≤ ·)] [ContravariantClass α α (swap (· * ·)) (· ≤ ·)] (a : α) {b : α} : a ≤ b * a ↔ 1 ≤ b := Iff.trans (by rw [one_mul]) (mul_le_mul_iff_right a) +#align le_mul_iff_one_le_left' le_mul_iff_one_le_left' +#align le_add_iff_nonneg_left le_add_iff_nonneg_left @[to_additive (attr := simp) add_le_iff_nonpos_right] theorem mul_le_iff_le_one_right' [CovariantClass α α (· * ·) (· ≤ ·)] [ContravariantClass α α (· * ·) (· ≤ ·)] (a : α) {b : α} : a * b ≤ a ↔ b ≤ 1 := Iff.trans (by rw [mul_one]) (mul_le_mul_iff_left a) +#align mul_le_iff_le_one_right' mul_le_iff_le_one_right' +#align add_le_iff_nonpos_right add_le_iff_nonpos_right @[to_additive (attr := simp) add_le_iff_nonpos_left] theorem mul_le_iff_le_one_left' [CovariantClass α α (swap (· * ·)) (· ≤ ·)] [ContravariantClass α α (swap (· * ·)) (· ≤ ·)] {a b : α} : a * b ≤ b ↔ a ≤ 1 := Iff.trans (by rw [one_mul]) (mul_le_mul_iff_right b) +#align mul_le_iff_le_one_left' mul_le_iff_le_one_left' +#align add_le_iff_nonpos_left add_le_iff_nonpos_left end LE @@ -361,6 +442,8 @@ theorem lt_mul_of_one_lt_right' [CovariantClass α α (· * ·) (· < ·)] (a : calc a = a * 1 := (mul_one a).symm _ < a * b := mul_lt_mul_left' h a +#align lt_mul_of_one_lt_right' lt_mul_of_one_lt_right' +#align lt_add_of_pos_right lt_add_of_pos_right @[to_additive add_lt_of_neg_right] theorem mul_lt_of_lt_one_right' [CovariantClass α α (· * ·) (· < ·)] (a : α) {b : α} (h : b < 1) : @@ -368,6 +451,8 @@ theorem mul_lt_of_lt_one_right' [CovariantClass α α (· * ·) (· < ·)] (a : calc a * b < a * 1 := mul_lt_mul_left' h a _ = a := mul_one a +#align mul_lt_of_lt_one_right' mul_lt_of_lt_one_right' +#align add_lt_of_neg_right add_lt_of_neg_right @[to_additive lt_add_of_pos_left] theorem lt_mul_of_one_lt_left' [CovariantClass α α (swap (· * ·)) (· < ·)] (a : α) {b : α} @@ -376,6 +461,8 @@ theorem lt_mul_of_one_lt_left' [CovariantClass α α (swap (· * ·)) (· < ·)] calc a = 1 * a := (one_mul a).symm _ < b * a := mul_lt_mul_right' h a +#align lt_mul_of_one_lt_left' lt_mul_of_one_lt_left' +#align lt_add_of_pos_left lt_add_of_pos_left @[to_additive add_lt_of_neg_left] theorem mul_lt_of_lt_one_left' [CovariantClass α α (swap (· * ·)) (· < ·)] (a : α) {b : α} @@ -384,50 +471,68 @@ theorem mul_lt_of_lt_one_left' [CovariantClass α α (swap (· * ·)) (· < ·)] calc b * a < 1 * a := mul_lt_mul_right' h a _ = a := one_mul a +#align mul_lt_of_lt_one_left' mul_lt_of_lt_one_left' +#align add_lt_of_neg_left add_lt_of_neg_left @[to_additive] theorem one_lt_of_lt_mul_right [ContravariantClass α α (· * ·) (· < ·)] {a b : α} (h : a < a * b) : 1 < b := lt_of_mul_lt_mul_left' <| by simpa only [mul_one] +#align one_lt_of_lt_mul_right one_lt_of_lt_mul_right +#align pos_of_lt_add_right pos_of_lt_add_right @[to_additive] theorem lt_one_of_mul_lt_right [ContravariantClass α α (· * ·) (· < ·)] {a b : α} (h : a * b < a) : b < 1 := lt_of_mul_lt_mul_left' <| by simpa only [mul_one] +#align lt_one_of_mul_lt_right lt_one_of_mul_lt_right +#align neg_of_add_lt_right neg_of_add_lt_right @[to_additive] theorem one_lt_of_lt_mul_left [ContravariantClass α α (swap (· * ·)) (· < ·)] {a b : α} (h : b < a * b) : 1 < a := lt_of_mul_lt_mul_right' <| by simpa only [one_mul] +#align one_lt_of_lt_mul_left one_lt_of_lt_mul_left +#align pos_of_lt_add_left pos_of_lt_add_left @[to_additive] theorem lt_one_of_mul_lt_left [ContravariantClass α α (swap (· * ·)) (· < ·)] {a b : α} (h : a * b < b) : a < 1 := lt_of_mul_lt_mul_right' <| by simpa only [one_mul] +#align lt_one_of_mul_lt_left lt_one_of_mul_lt_left +#align neg_of_add_lt_left neg_of_add_lt_left @[to_additive (attr := simp) lt_add_iff_pos_right] theorem lt_mul_iff_one_lt_right' [CovariantClass α α (· * ·) (· < ·)] [ContravariantClass α α (· * ·) (· < ·)] (a : α) {b : α} : a < a * b ↔ 1 < b := Iff.trans (by rw [mul_one]) (mul_lt_mul_iff_left a) +#align lt_mul_iff_one_lt_right' lt_mul_iff_one_lt_right' +#align lt_add_iff_pos_right lt_add_iff_pos_right @[to_additive (attr := simp) lt_add_iff_pos_left] theorem lt_mul_iff_one_lt_left' [CovariantClass α α (swap (· * ·)) (· < ·)] [ContravariantClass α α (swap (· * ·)) (· < ·)] (a : α) {b : α} : a < b * a ↔ 1 < b := Iff.trans (by rw [one_mul]) (mul_lt_mul_iff_right a) +#align lt_mul_iff_one_lt_left' lt_mul_iff_one_lt_left' +#align lt_add_iff_pos_left lt_add_iff_pos_left @[to_additive (attr := simp) add_lt_iff_neg_left] theorem mul_lt_iff_lt_one_left' [CovariantClass α α (· * ·) (· < ·)] [ContravariantClass α α (· * ·) (· < ·)] {a b : α} : a * b < a ↔ b < 1 := Iff.trans (by rw [mul_one]) (mul_lt_mul_iff_left a) +#align mul_lt_iff_lt_one_left' mul_lt_iff_lt_one_left' +#align add_lt_iff_neg_left add_lt_iff_neg_left @[to_additive (attr := simp) add_lt_iff_neg_right] theorem mul_lt_iff_lt_one_right' [CovariantClass α α (swap (· * ·)) (· < ·)] [ContravariantClass α α (swap (· * ·)) (· < ·)] {a : α} (b : α) : a * b < b ↔ a < 1 := Iff.trans (by rw [one_mul]) (mul_lt_mul_iff_right b) +#align mul_lt_iff_lt_one_right' mul_lt_iff_lt_one_right' +#align add_lt_iff_neg_right add_lt_iff_neg_right end LT @@ -447,6 +552,8 @@ theorem mul_le_of_le_of_le_one [CovariantClass α α (· * ·) (· ≤ ·)] {a b b * a ≤ b * 1 := mul_le_mul_left' ha b _ = b := mul_one b _ ≤ c := hbc +#align mul_le_of_le_of_le_one mul_le_of_le_of_le_one +#align add_le_of_le_of_nonpos add_le_of_le_of_nonpos @[to_additive] theorem mul_lt_of_le_of_lt_one [CovariantClass α α (· * ·) (· < ·)] {a b c : α} (hbc : b ≤ c) @@ -456,6 +563,8 @@ theorem mul_lt_of_le_of_lt_one [CovariantClass α α (· * ·) (· < ·)] {a b c b * a < b * 1 := mul_lt_mul_left' ha b _ = b := mul_one b _ ≤ c := hbc +#align mul_lt_of_le_of_lt_one mul_lt_of_le_of_lt_one +#align add_lt_of_le_of_neg add_lt_of_le_of_neg @[to_additive] theorem mul_lt_of_lt_of_le_one [CovariantClass α α (· * ·) (· ≤ ·)] {a b c : α} (hbc : b < c) @@ -465,6 +574,8 @@ theorem mul_lt_of_lt_of_le_one [CovariantClass α α (· * ·) (· ≤ ·)] {a b b * a ≤ b * 1 := mul_le_mul_left' ha b _ = b := mul_one b _ < c := hbc +#align mul_lt_of_lt_of_le_one mul_lt_of_lt_of_le_one +#align add_lt_of_lt_of_nonpos add_lt_of_lt_of_nonpos @[to_additive] theorem mul_lt_of_lt_of_lt_one [CovariantClass α α (· * ·) (· < ·)] {a b c : α} (hbc : b < c) @@ -474,12 +585,16 @@ theorem mul_lt_of_lt_of_lt_one [CovariantClass α α (· * ·) (· < ·)] {a b c b * a < b * 1 := mul_lt_mul_left' ha b _ = b := mul_one b _ < c := hbc +#align mul_lt_of_lt_of_lt_one mul_lt_of_lt_of_lt_one +#align add_lt_of_lt_of_neg add_lt_of_lt_of_neg @[to_additive] theorem mul_lt_of_lt_of_lt_one' [CovariantClass α α (· * ·) (· ≤ ·)] {a b c : α} (hbc : b < c) (ha : a < 1) : b * a < c := mul_lt_of_lt_of_le_one hbc ha.le +#align mul_lt_of_lt_of_lt_one' mul_lt_of_lt_of_lt_one' +#align add_lt_of_lt_of_neg' add_lt_of_lt_of_neg' /-- Assumes left covariance. The lemma assuming right covariance is `Right.mul_le_one`. -/ @@ -547,6 +662,8 @@ theorem le_mul_of_le_of_one_le [CovariantClass α α (· * ·) (· ≤ ·)] {a b b ≤ c := hbc _ = c * 1 := (mul_one c).symm _ ≤ c * a := mul_le_mul_left' ha c +#align le_mul_of_le_of_one_le le_mul_of_le_of_one_le +#align le_add_of_le_of_nonneg le_add_of_le_of_nonneg @[to_additive] theorem lt_mul_of_le_of_one_lt [CovariantClass α α (· * ·) (· < ·)] {a b c : α} (hbc : b ≤ c) @@ -556,6 +673,8 @@ theorem lt_mul_of_le_of_one_lt [CovariantClass α α (· * ·) (· < ·)] {a b c b ≤ c := hbc _ = c * 1 := (mul_one c).symm _ < c * a := mul_lt_mul_left' ha c +#align lt_mul_of_le_of_one_lt lt_mul_of_le_of_one_lt +#align lt_add_of_le_of_pos lt_add_of_le_of_pos @[to_additive] theorem lt_mul_of_lt_of_one_le [CovariantClass α α (· * ·) (· ≤ ·)] {a b c : α} (hbc : b < c) @@ -565,6 +684,8 @@ theorem lt_mul_of_lt_of_one_le [CovariantClass α α (· * ·) (· ≤ ·)] {a b b < c := hbc _ = c * 1 := (mul_one c).symm _ ≤ c * a := mul_le_mul_left' ha c +#align lt_mul_of_lt_of_one_le lt_mul_of_lt_of_one_le +#align lt_add_of_lt_of_nonneg lt_add_of_lt_of_nonneg @[to_additive] theorem lt_mul_of_lt_of_one_lt [CovariantClass α α (· * ·) (· < ·)] {a b c : α} (hbc : b < c) @@ -574,12 +695,16 @@ theorem lt_mul_of_lt_of_one_lt [CovariantClass α α (· * ·) (· < ·)] {a b c b < c := hbc _ = c * 1 := (mul_one c).symm _ < c * a := mul_lt_mul_left' ha c +#align lt_mul_of_lt_of_one_lt lt_mul_of_lt_of_one_lt +#align lt_add_of_lt_of_pos lt_add_of_lt_of_pos @[to_additive] theorem lt_mul_of_lt_of_one_lt' [CovariantClass α α (· * ·) (· ≤ ·)] {a b c : α} (hbc : b < c) (ha : 1 < a) : b < c * a := lt_mul_of_lt_of_one_le hbc ha.le +#align lt_mul_of_lt_of_one_lt' lt_mul_of_lt_of_one_lt' +#align lt_add_of_lt_of_pos' lt_add_of_lt_of_pos' /-- Assumes left covariance. The lemma assuming right covariance is `Right.one_le_mul`. -/ @@ -647,6 +772,8 @@ theorem mul_le_of_le_one_of_le [CovariantClass α α (swap (· * ·)) (· ≤ · a * b ≤ 1 * b := mul_le_mul_right' ha b _ = b := one_mul b _ ≤ c := hbc +#align mul_le_of_le_one_of_le mul_le_of_le_one_of_le +#align add_le_of_nonpos_of_le add_le_of_nonpos_of_le @[to_additive] theorem mul_lt_of_lt_one_of_le [CovariantClass α α (swap (· * ·)) (· < ·)] {a b c : α} (ha : a < 1) @@ -656,6 +783,8 @@ theorem mul_lt_of_lt_one_of_le [CovariantClass α α (swap (· * ·)) (· < ·)] a * b < 1 * b := mul_lt_mul_right' ha b _ = b := one_mul b _ ≤ c := hbc +#align mul_lt_of_lt_one_of_le mul_lt_of_lt_one_of_le +#align add_lt_of_neg_of_le add_lt_of_neg_of_le @[to_additive] theorem mul_lt_of_le_one_of_lt [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b c : α} (ha : a ≤ 1) @@ -665,6 +794,8 @@ theorem mul_lt_of_le_one_of_lt [CovariantClass α α (swap (· * ·)) (· ≤ · a * b ≤ 1 * b := mul_le_mul_right' ha b _ = b := one_mul b _ < c := hb +#align mul_lt_of_le_one_of_lt mul_lt_of_le_one_of_lt +#align add_lt_of_nonpos_of_lt add_lt_of_nonpos_of_lt @[to_additive] theorem mul_lt_of_lt_one_of_lt [CovariantClass α α (swap (· * ·)) (· < ·)] {a b c : α} (ha : a < 1) @@ -674,12 +805,16 @@ theorem mul_lt_of_lt_one_of_lt [CovariantClass α α (swap (· * ·)) (· < ·)] a * b < 1 * b := mul_lt_mul_right' ha b _ = b := one_mul b _ < c := hb +#align mul_lt_of_lt_one_of_lt mul_lt_of_lt_one_of_lt +#align add_lt_of_neg_of_lt add_lt_of_neg_of_lt @[to_additive] theorem mul_lt_of_lt_one_of_lt' [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b c : α} (ha : a < 1) (hbc : b < c) : a * b < c := mul_lt_of_le_one_of_lt ha.le hbc +#align mul_lt_of_lt_one_of_lt' mul_lt_of_lt_one_of_lt' +#align add_lt_of_neg_of_lt' add_lt_of_neg_of_lt' /-- Assumes right covariance. The lemma assuming left covariance is `Left.mul_le_one`. -/ @@ -750,6 +885,8 @@ theorem le_mul_of_one_le_of_le [CovariantClass α α (swap (· * ·)) (· ≤ · b ≤ c := hbc _ = 1 * c := (one_mul c).symm _ ≤ a * c := mul_le_mul_right' ha c +#align le_mul_of_one_le_of_le le_mul_of_one_le_of_le +#align le_add_of_nonneg_of_le le_add_of_nonneg_of_le @[to_additive] theorem lt_mul_of_one_lt_of_le [CovariantClass α α (swap (· * ·)) (· < ·)] {a b c : α} (ha : 1 < a) @@ -759,6 +896,8 @@ theorem lt_mul_of_one_lt_of_le [CovariantClass α α (swap (· * ·)) (· < ·)] b ≤ c := hbc _ = 1 * c := (one_mul c).symm _ < a * c := mul_lt_mul_right' ha c +#align lt_mul_of_one_lt_of_le lt_mul_of_one_lt_of_le +#align lt_add_of_pos_of_le lt_add_of_pos_of_le @[to_additive] theorem lt_mul_of_one_le_of_lt [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b c : α} (ha : 1 ≤ a) @@ -768,6 +907,8 @@ theorem lt_mul_of_one_le_of_lt [CovariantClass α α (swap (· * ·)) (· ≤ · b < c := hbc _ = 1 * c := (one_mul c).symm _ ≤ a * c := mul_le_mul_right' ha c +#align lt_mul_of_one_le_of_lt lt_mul_of_one_le_of_lt +#align lt_add_of_nonneg_of_lt lt_add_of_nonneg_of_lt @[to_additive] theorem lt_mul_of_one_lt_of_lt [CovariantClass α α (swap (· * ·)) (· < ·)] {a b c : α} (ha : 1 < a) @@ -777,12 +918,16 @@ theorem lt_mul_of_one_lt_of_lt [CovariantClass α α (swap (· * ·)) (· < ·)] b < c := hbc _ = 1 * c := (one_mul c).symm _ < a * c := mul_lt_mul_right' ha c +#align lt_mul_of_one_lt_of_lt lt_mul_of_one_lt_of_lt +#align lt_add_of_pos_of_lt lt_add_of_pos_of_lt @[to_additive] theorem lt_mul_of_one_lt_of_lt' [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b c : α} (ha : 1 < a) (hbc : b < c) : b < a * c := lt_mul_of_one_le_of_lt ha.le hbc +#align lt_mul_of_one_lt_of_lt' lt_mul_of_one_lt_of_lt' +#align lt_add_of_pos_of_lt' lt_add_of_pos_of_lt' /-- Assumes right covariance. The lemma assuming left covariance is `Left.one_le_mul`. -/ @@ -842,96 +987,132 @@ theorem Right.one_lt_mul' [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a #align right.add_pos' Right.add_pos' alias Left.mul_le_one ← mul_le_one' +#align mul_le_one' mul_le_one' alias Left.mul_lt_one_of_le_of_lt ← mul_lt_one_of_le_of_lt +#align mul_lt_one_of_le_of_lt mul_lt_one_of_le_of_lt alias Left.mul_lt_one_of_lt_of_le ← mul_lt_one_of_lt_of_le +#align mul_lt_one_of_lt_of_le mul_lt_one_of_lt_of_le alias Left.mul_lt_one ← mul_lt_one +#align mul_lt_one mul_lt_one alias Left.mul_lt_one' ← mul_lt_one' +#align mul_lt_one' mul_lt_one' attribute [to_additive add_nonpos "**Alias** of `Left.add_nonpos`."] mul_le_one' +#align add_nonpos add_nonpos attribute [to_additive add_neg_of_nonpos_of_neg "**Alias** of `Left.add_neg_of_nonpos_of_neg`."] mul_lt_one_of_le_of_lt +#align add_neg_of_nonpos_of_neg add_neg_of_nonpos_of_neg attribute [to_additive add_neg_of_neg_of_nonpos "**Alias** of `Left.add_neg_of_neg_of_nonpos`."] mul_lt_one_of_lt_of_le +#align add_neg_of_neg_of_nonpos add_neg_of_neg_of_nonpos attribute [to_additive "**Alias** of `Left.add_neg`."] mul_lt_one +#align add_neg add_neg attribute [to_additive "**Alias** of `Left.add_neg'`."] mul_lt_one' +#align add_neg' add_neg' alias Left.one_le_mul ← one_le_mul +#align one_le_mul one_le_mul alias Left.one_lt_mul_of_le_of_lt ← one_lt_mul_of_le_of_lt' +#align one_lt_mul_of_le_of_lt' one_lt_mul_of_le_of_lt' alias Left.one_lt_mul_of_lt_of_le ← one_lt_mul_of_lt_of_le' +#align one_lt_mul_of_lt_of_le' one_lt_mul_of_lt_of_le' alias Left.one_lt_mul ← one_lt_mul' +#align one_lt_mul' one_lt_mul' alias Left.one_lt_mul' ← one_lt_mul'' +#align one_lt_mul'' one_lt_mul'' attribute [to_additive add_nonneg "**Alias** of `Left.add_nonneg`."] one_le_mul +#align add_nonneg add_nonneg attribute [to_additive add_pos_of_nonneg_of_pos "**Alias** of `Left.add_pos_of_nonneg_of_pos`."] one_lt_mul_of_le_of_lt' +#align add_pos_of_nonneg_of_pos add_pos_of_nonneg_of_pos attribute [to_additive add_pos_of_pos_of_nonneg "**Alias** of `Left.add_pos_of_pos_of_nonneg`."] one_lt_mul_of_lt_of_le' +#align add_pos_of_pos_of_nonneg add_pos_of_pos_of_nonneg attribute [to_additive add_pos "**Alias** of `Left.add_pos`."] one_lt_mul' +#align add_pos add_pos attribute [to_additive add_pos' "**Alias** of `Left.add_pos'`."] one_lt_mul'' +#align add_pos' add_pos' @[to_additive] theorem lt_of_mul_lt_of_one_le_left [CovariantClass α α (· * ·) (· ≤ ·)] {a b c : α} (h : a * b < c) (hle : 1 ≤ b) : a < c := (le_mul_of_one_le_right' hle).trans_lt h +#align lt_of_mul_lt_of_one_le_left lt_of_mul_lt_of_one_le_left +#align lt_of_add_lt_of_nonneg_left lt_of_add_lt_of_nonneg_left @[to_additive] theorem le_of_mul_le_of_one_le_left [CovariantClass α α (· * ·) (· ≤ ·)] {a b c : α} (h : a * b ≤ c) (hle : 1 ≤ b) : a ≤ c := (le_mul_of_one_le_right' hle).trans h +#align le_of_mul_le_of_one_le_left le_of_mul_le_of_one_le_left +#align le_of_add_le_of_nonneg_left le_of_add_le_of_nonneg_left @[to_additive] theorem lt_of_lt_mul_of_le_one_left [CovariantClass α α (· * ·) (· ≤ ·)] {a b c : α} (h : a < b * c) (hle : c ≤ 1) : a < b := h.trans_le (mul_le_of_le_one_right' hle) +#align lt_of_lt_mul_of_le_one_left lt_of_lt_mul_of_le_one_left +#align lt_of_lt_add_of_nonpos_left lt_of_lt_add_of_nonpos_left @[to_additive] theorem le_of_le_mul_of_le_one_left [CovariantClass α α (· * ·) (· ≤ ·)] {a b c : α} (h : a ≤ b * c) (hle : c ≤ 1) : a ≤ b := h.trans (mul_le_of_le_one_right' hle) +#align le_of_le_mul_of_le_one_left le_of_le_mul_of_le_one_left +#align le_of_le_add_of_nonpos_left le_of_le_add_of_nonpos_left @[to_additive] theorem lt_of_mul_lt_of_one_le_right [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b c : α} (h : a * b < c) (hle : 1 ≤ a) : b < c := (le_mul_of_one_le_left' hle).trans_lt h +#align lt_of_mul_lt_of_one_le_right lt_of_mul_lt_of_one_le_right +#align lt_of_add_lt_of_nonneg_right lt_of_add_lt_of_nonneg_right @[to_additive] theorem le_of_mul_le_of_one_le_right [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b c : α} (h : a * b ≤ c) (hle : 1 ≤ a) : b ≤ c := (le_mul_of_one_le_left' hle).trans h +#align le_of_mul_le_of_one_le_right le_of_mul_le_of_one_le_right +#align le_of_add_le_of_nonneg_right le_of_add_le_of_nonneg_right @[to_additive] theorem lt_of_lt_mul_of_le_one_right [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b c : α} (h : a < b * c) (hle : b ≤ 1) : a < c := h.trans_le (mul_le_of_le_one_left' hle) +#align lt_of_lt_mul_of_le_one_right lt_of_lt_mul_of_le_one_right +#align lt_of_lt_add_of_nonpos_right lt_of_lt_add_of_nonpos_right @[to_additive] theorem le_of_le_mul_of_le_one_right [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b c : α} (h : a ≤ b * c) (hle : b ≤ 1) : a ≤ c := h.trans (mul_le_of_le_one_left' hle) +#align le_of_le_mul_of_le_one_right le_of_le_mul_of_le_one_right +#align le_of_le_add_of_nonpos_right le_of_le_add_of_nonpos_right end Preorder @@ -954,6 +1135,8 @@ theorem mul_eq_one_iff' [CovariantClass α α (· * ·) (· ≤ ·)] -- porting note: original proof of the second implication, -- `fun ⟨ha', hb'⟩ => by rw [ha', hb', mul_one]`, -- had its `to_additive`-ization fail due to some bug +#align mul_eq_one_iff' mul_eq_one_iff' +#align add_eq_zero_iff' add_eq_zero_iff' @[to_additive] lemma mul_le_mul_iff_of_ge [CovariantClass α α (· * ·) (· ≤ ·)] [CovariantClass α α (swap (· * ·)) (· ≤ ·)] [CovariantClass α α (· * ·) (· < ·)] @@ -964,6 +1147,8 @@ theorem mul_eq_one_iff' [CovariantClass α α (· * ·) (· ≤ ·)] refine' ⟨fun ha ↦ h.not_lt _, fun hb ↦ h.not_lt _⟩ { exact mul_lt_mul_of_lt_of_le ha hb } { exact mul_lt_mul_of_le_of_lt ha hb } +#align add_le_add_iff_of_ge add_le_add_iff_of_ge +#align mul_le_mul_iff_of_ge mul_le_mul_iff_of_ge section Left @@ -972,10 +1157,14 @@ variable [CovariantClass α α (· * ·) (· ≤ ·)] {a b : α} @[to_additive eq_zero_of_add_nonneg_left] theorem eq_one_of_one_le_mul_left (ha : a ≤ 1) (hb : b ≤ 1) (hab : 1 ≤ a * b) : a = 1 := ha.eq_of_not_lt fun h => hab.not_lt <| mul_lt_one_of_lt_of_le h hb +#align eq_one_of_one_le_mul_left eq_one_of_one_le_mul_left +#align eq_zero_of_add_nonneg_left eq_zero_of_add_nonneg_left @[to_additive] theorem eq_one_of_mul_le_one_left (ha : 1 ≤ a) (hb : 1 ≤ b) (hab : a * b ≤ 1) : a = 1 := ha.eq_of_not_gt fun h => hab.not_lt <| one_lt_mul_of_lt_of_le' h hb +#align eq_one_of_mul_le_one_left eq_one_of_mul_le_one_left +#align eq_zero_of_add_nonpos_left eq_zero_of_add_nonpos_left end Left @@ -986,10 +1175,14 @@ variable [CovariantClass α α (swap (· * ·)) (· ≤ ·)] {a b : α} @[to_additive eq_zero_of_add_nonneg_right] theorem eq_one_of_one_le_mul_right (ha : a ≤ 1) (hb : b ≤ 1) (hab : 1 ≤ a * b) : b = 1 := hb.eq_of_not_lt fun h => hab.not_lt <| Right.mul_lt_one_of_le_of_lt ha h +#align eq_one_of_one_le_mul_right eq_one_of_one_le_mul_right +#align eq_zero_of_add_nonneg_right eq_zero_of_add_nonneg_right @[to_additive] theorem eq_one_of_mul_le_one_right (ha : 1 ≤ a) (hb : 1 ≤ b) (hab : a * b ≤ 1) : b = 1 := hb.eq_of_not_gt fun h => hab.not_lt <| Right.one_lt_mul_of_le_of_lt ha h +#align eq_one_of_mul_le_one_right eq_one_of_mul_le_one_right +#align eq_zero_of_add_nonpos_right eq_zero_of_add_nonpos_right end Right @@ -1005,10 +1198,10 @@ theorem exists_square_le [CovariantClass α α (· * ·) (· < ·)] (a : α) : have : a * a < a * 1 := mul_lt_mul_left' h a rw [mul_one] at this exact le_of_lt this - · use 1 push_neg at h rwa [mul_one] +#align exists_square_le exists_square_le end LinearOrder @@ -1081,8 +1274,10 @@ theorem Right.mul_eq_mul_iff_eq_and_eq [CovariantClass α α (· * ·) (· ≤ #align right.add_eq_add_iff_eq_and_eq Right.add_eq_add_iff_eq_and_eq alias Left.mul_eq_mul_iff_eq_and_eq ← mul_eq_mul_iff_eq_and_eq +#align mul_eq_mul_iff_eq_and_eq mul_eq_mul_iff_eq_and_eq attribute [to_additive] mul_eq_mul_iff_eq_and_eq +#align add_eq_add_iff_eq_and_eq add_eq_add_iff_eq_and_eq end PartialOrder @@ -1362,12 +1557,16 @@ theorem cmp_mul_left' {α : Type _} [Mul α] [LinearOrder α] [CovariantClass α (a b c : α) : cmp (a * b) (a * c) = cmp b c := (strictMono_id.const_mul' a).cmp_map_eq b c +#align cmp_mul_left' cmp_mul_left' +#align cmp_add_left cmp_add_left @[to_additive (attr := simp) cmp_add_right] theorem cmp_mul_right' {α : Type _} [Mul α] [LinearOrder α] [CovariantClass α α (swap (· * ·)) (· < ·)] (a b c : α) : cmp (a * c) (b * c) = cmp a b := (strictMono_id.mul_const' c).cmp_map_eq a b +#align cmp_mul_right' cmp_mul_right' +#align cmp_add_right cmp_add_right end Mono diff --git a/Mathlib/Algebra/Order/Monoid/NatCast.lean b/Mathlib/Algebra/Order/Monoid/NatCast.lean index 1d9c1e6817e68..cc02fe27e5284 100644 --- a/Mathlib/Algebra/Order/Monoid/NatCast.lean +++ b/Mathlib/Algebra/Order/Monoid/NatCast.lean @@ -23,10 +23,12 @@ open Function lemma lt_add_one [One α] [AddZeroClass α] [PartialOrder α] [ZeroLEOneClass α] [NeZero (1 : α)] [CovariantClass α α (·+·) (·<·)] (a : α) : a < a + 1 := lt_add_of_pos_right _ zero_lt_one +#align lt_add_one lt_add_one lemma lt_one_add [One α] [AddZeroClass α] [PartialOrder α] [ZeroLEOneClass α] [NeZero (1 : α)] [CovariantClass α α (swap (·+·)) (·<·)] (a : α) : a < 1 + a := lt_add_of_pos_left _ zero_lt_one +#align lt_one_add lt_one_add variable [AddMonoidWithOne α] @@ -34,28 +36,33 @@ lemma zero_le_two [Preorder α] [ZeroLEOneClass α] [CovariantClass α α (·+· (0 : α) ≤ 2 := by rw [← one_add_one_eq_two] exact add_nonneg zero_le_one zero_le_one +#align zero_le_two zero_le_two lemma zero_le_three [Preorder α] [ZeroLEOneClass α] [CovariantClass α α (·+·) (·≤·)] : (0 : α) ≤ 3 := by rw [← two_add_one_eq_three] exact add_nonneg zero_le_two zero_le_one +#align zero_le_three zero_le_three lemma zero_le_four [Preorder α] [ZeroLEOneClass α] [CovariantClass α α (·+·) (·≤·)] : (0 : α) ≤ 4 := by rw [← three_add_one_eq_four] exact add_nonneg zero_le_three zero_le_one +#align zero_le_four zero_le_four lemma one_le_two [LE α] [ZeroLEOneClass α] [CovariantClass α α (·+·) (·≤·)] : (1 : α) ≤ 2 := calc 1 = 1 + 0 := (add_zero 1).symm _ ≤ 1 + 1 := add_le_add_left zero_le_one _ _ = 2 := one_add_one_eq_two +#align one_le_two one_le_two lemma one_le_two' [LE α] [ZeroLEOneClass α] [CovariantClass α α (swap (·+·)) (·≤·)] : (1 : α) ≤ 2 := calc 1 = 0 + 1 := (zero_add 1).symm _ ≤ 1 + 1 := add_le_add_right zero_le_one _ _ = 2 := one_add_one_eq_two +#align one_le_two' one_le_two' section variable [PartialOrder α] [ZeroLEOneClass α] [NeZero (1 : α)] @@ -73,6 +80,9 @@ variable [CovariantClass α α (·+·) (·≤·)] @[simp] lemma zero_lt_four : (0 : α) < 4 := by rw [← three_add_one_eq_four] exact lt_add_of_lt_of_nonneg zero_lt_three zero_le_one +#align zero_lt_four zero_lt_four +#align zero_lt_three zero_lt_three +#align zero_lt_two zero_lt_two variable (α) @@ -82,6 +92,9 @@ lemma zero_lt_two' : (0 : α) < 2 := zero_lt_two lemma zero_lt_three' : (0 : α) < 3 := zero_lt_three /-- See `zero_lt_four` for a version with the type implicit. -/ lemma zero_lt_four' : (0 : α) < 4 := zero_lt_four +#align zero_lt_four' zero_lt_four' +#align zero_lt_three' zero_lt_three' +#align zero_lt_two' zero_lt_two' instance ZeroLEOneClass.neZero.two : NeZero (2 : α) := ⟨zero_lt_two.ne'⟩ instance ZeroLEOneClass.neZero.three : NeZero (3 : α) := ⟨zero_lt_three.ne'⟩ @@ -92,9 +105,13 @@ end lemma one_lt_two [CovariantClass α α (·+·) (·<·)] : (1 : α) < 2 := by rw [← one_add_one_eq_two] exact lt_add_one _ +#align one_lt_two one_lt_two end alias zero_lt_two ← two_pos alias zero_lt_three ← three_pos alias zero_lt_four ← four_pos +#align four_pos four_pos +#align three_pos three_pos +#align two_pos two_pos diff --git a/Mathlib/Algebra/Order/Monoid/WithTop.lean b/Mathlib/Algebra/Order/Monoid/WithTop.lean index 3b3e6b1f58ee5..34e9cdaee6447 100644 --- a/Mathlib/Algebra/Order/Monoid/WithTop.lean +++ b/Mathlib/Algebra/Order/Monoid/WithTop.lean @@ -51,11 +51,13 @@ theorem coe_eq_one {a : α} : (a : WithTop α) = 1 ↔ a = 1 := theorem one_le_coe [LE α] {a : α} : 1 ≤ (a : WithTop α) ↔ 1 ≤ a := coe_le_coe #align with_top.one_le_coe WithTop.one_le_coe +#align with_top.coe_nonneg WithTop.coe_nonneg @[to_additive (attr := simp, norm_cast) coe_le_zero] theorem coe_le_one [LE α] {a : α} : (a : WithTop α) ≤ 1 ↔ a ≤ 1 := coe_le_coe #align with_top.coe_le_one WithTop.coe_le_one +#align with_top.coe_le_zero WithTop.coe_le_zero @[to_additive (attr := simp, norm_cast) coe_pos] theorem one_lt_coe [LT α] {a : α} : 1 < (a : WithTop α) ↔ 1 < a := @@ -410,6 +412,7 @@ protected def _root_.OneHom.withTopMap {M N : Type _} [One M] [One N] (f : OneHo map_one' := by rw [WithTop.map_one, map_one, coe_one] #align one_hom.with_top_map OneHom.withTopMap #align zero_hom.with_top_map ZeroHom.withTopMap +#align one_hom.with_top_map_apply OneHom.withTopMap_apply /-- A version of `WithTop.map` for `AddHom`s. -/ @[simps (config := { fullyApplied := false })] @@ -418,6 +421,7 @@ protected def _root_.AddHom.withTopMap {M N : Type _} [Add M] [Add N] (f : AddHo toFun := WithTop.map f map_add' := WithTop.map_add f #align add_hom.with_top_map AddHom.withTopMap +#align add_hom.with_top_map_apply AddHom.withTopMap_apply /-- A version of `WithTop.map` for `AddMonoidHom`s. -/ @[simps (config := { fullyApplied := false })] @@ -425,6 +429,7 @@ protected def _root_.AddMonoidHom.withTopMap {M N : Type _} [AddZeroClass M] [Ad (f : M →+ N) : WithTop M →+ WithTop N := { ZeroHom.withTopMap f.toZeroHom, AddHom.withTopMap f.toAddHom with toFun := WithTop.map f } #align add_monoid_hom.with_top_map AddMonoidHom.withTopMap +#align add_monoid_hom.with_top_map_apply AddMonoidHom.withTopMap_apply end WithTop @@ -479,11 +484,13 @@ theorem coe_eq_one [One α] {a : α} : (a : WithBot α) = 1 ↔ a = 1 := theorem one_le_coe [One α] [LE α] {a : α} : 1 ≤ (a : WithBot α) ↔ 1 ≤ a := coe_le_coe #align with_bot.one_le_coe WithBot.one_le_coe +#align with_bot.coe_nonneg WithBot.coe_nonneg @[to_additive (attr := simp, norm_cast) coe_le_zero] theorem coe_le_one [One α] [LE α] {a : α} : (a : WithBot α) ≤ 1 ↔ a ≤ 1 := coe_le_coe #align with_bot.coe_le_one WithBot.coe_le_one +#align with_bot.coe_le_zero WithBot.coe_le_zero @[to_additive (attr := simp, norm_cast) coe_pos] theorem one_lt_coe [One α] [LT α] {a : α} : 1 < (a : WithBot α) ↔ 1 < a := @@ -598,6 +605,7 @@ protected def _root_.OneHom.withBotMap {M N : Type _} [One M] [One N] (f : OneHo map_one' := by rw [WithBot.map_one, map_one, coe_one] #align one_hom.with_bot_map OneHom.withBotMap #align zero_hom.with_bot_map ZeroHom.withBotMap +#align one_hom.with_bot_map_apply OneHom.withBotMap_apply /-- A version of `WithBot.map` for `AddHom`s. -/ @[simps (config := { fullyApplied := false })] @@ -606,6 +614,7 @@ protected def _root_.AddHom.withBotMap {M N : Type _} [Add M] [Add N] (f : AddHo toFun := WithBot.map f map_add' := WithBot.map_add f #align add_hom.with_bot_map AddHom.withBotMap +#align add_hom.with_bot_map_apply AddHom.withBotMap_apply /-- A version of `WithBot.map` for `AddMonoidHom`s. -/ @[simps (config := { fullyApplied := false })] @@ -613,6 +622,7 @@ protected def _root_.AddMonoidHom.withBotMap {M N : Type _} [AddZeroClass M] [Ad (f : M →+ N) : WithBot M →+ WithBot N := { ZeroHom.withBotMap f.toZeroHom, AddHom.withBotMap f.toAddHom with toFun := WithBot.map f } #align add_monoid_hom.with_bot_map AddMonoidHom.withBotMap +#align add_monoid_hom.with_bot_map_apply AddMonoidHom.withBotMap_apply variable [Preorder α] diff --git a/Mathlib/Algebra/Order/Ring/Cone.lean b/Mathlib/Algebra/Order/Ring/Cone.lean index 5ff36a18f8cf4..36f3cbcd40ebb 100644 --- a/Mathlib/Algebra/Order/Ring/Cone.lean +++ b/Mathlib/Algebra/Order/Ring/Cone.lean @@ -34,6 +34,7 @@ structure PositiveCone (α : Type _) [Ring α] extends AddCommGroup.PositiveCone /-- Forget that a positive cone in a ring respects the multiplicative structure. -/ add_decl_doc PositiveCone.toPositiveCone +#align ring.positive_cone.to_positive_cone Ring.PositiveCone.toPositiveCone /-- A total positive cone in a nontrivial ring induces a linear order. -/ structure TotalPositiveCone (α : Type _) [Ring α] extends PositiveCone α, @@ -46,6 +47,7 @@ add_decl_doc TotalPositiveCone.toPositiveCone_1 /-- Forget that a `TotalPositiveCone` in a ring respects the multiplicative structure. -/ add_decl_doc TotalPositiveCone.toTotalPositiveCone +#align ring.total_positive_cone.to_total_positive_cone Ring.TotalPositiveCone.toTotalPositiveCone theorem PositiveCone.one_pos (C : PositiveCone α) : C.pos 1 := (C.pos_iff _).2 ⟨C.one_nonneg, fun h => one_ne_zero <| C.nonneg_antisymm C.one_nonneg h⟩ diff --git a/Mathlib/Algebra/Order/Ring/Defs.lean b/Mathlib/Algebra/Order/Ring/Defs.lean index e93804af79ec3..f9abe5b924954 100644 --- a/Mathlib/Algebra/Order/Ring/Defs.lean +++ b/Mathlib/Algebra/Order/Ring/Defs.lean @@ -325,6 +325,7 @@ theorem one_lt_mul_of_lt_of_le (ha : 1 < a) (hb : 1 ≤ b) : 1 < a * b := #align one_lt_mul_of_lt_of_le one_lt_mul_of_lt_of_le alias one_lt_mul_of_le_of_lt ← one_lt_mul +#align one_lt_mul one_lt_mul theorem mul_lt_one_of_nonneg_of_lt_one_left (ha₀ : 0 ≤ a) (ha : a < 1) (hb : b ≤ 1) : a * b < 1 := (mul_le_of_le_one_right ha₀ hb).trans_lt ha diff --git a/Mathlib/Algebra/Order/Ring/Lemmas.lean b/Mathlib/Algebra/Order/Ring/Lemmas.lean index 91bcb142fc249..da401f67e34f6 100644 --- a/Mathlib/Algebra/Order/Ring/Lemmas.lean +++ b/Mathlib/Algebra/Order/Ring/Lemmas.lean @@ -76,21 +76,25 @@ notation "α>0" => { x : α // 0 < x } expressing that multiplication by nonnegative elements on the left is monotone. -/ abbrev PosMulMono : Prop := CovariantClass α≥0 α (fun x y => x * y) (· ≤ ·) +#align pos_mul_mono PosMulMono /-- `MulPosMono α` is an abbreviation for `CovariantClass α≥0 α (λ x y, y * x) (≤)`, expressing that multiplication by nonnegative elements on the right is monotone. -/ abbrev MulPosMono : Prop := CovariantClass α≥0 α (fun x y => y * x) (· ≤ ·) +#align mul_pos_mono MulPosMono /-- `PosMulStrictMono α` is an abbreviation for `CovariantClass α>0 α (λ x y, x * y) (<)`, expressing that multiplication by positive elements on the left is strictly monotone. -/ abbrev PosMulStrictMono : Prop := CovariantClass α>0 α (fun x y => x * y) (· < ·) +#align pos_mul_strict_mono PosMulStrictMono /-- `MulPosStrictMono α` is an abbreviation for `CovariantClass α>0 α (λ x y, y * x) (<)`, expressing that multiplication by positive elements on the right is strictly monotone. -/ abbrev MulPosStrictMono : Prop := CovariantClass α>0 α (fun x y => y * x) (· < ·) +#align mul_pos_strict_mono MulPosStrictMono /-- `PosMulReflectLT α` is an abbreviation for `ContravariantClas α≥0 α (λ x y, x * y) (<)`, expressing that multiplication by nonnegative elements on the left is strictly reverse monotone. -/ @@ -108,11 +112,13 @@ abbrev MulPosReflectLT : Prop := expressing that multiplication by positive elements on the left is reverse monotone. -/ abbrev PosMulMonoRev : Prop := ContravariantClass α>0 α (fun x y => x * y) (· ≤ ·) +#align pos_mul_mono_rev PosMulMonoRev /-- `MulPosMonoRev α` is an abbreviation for `ContravariantClas α>0 α (λ x y, y * x) (≤)`, expressing that multiplication by positive elements on the right is reverse monotone. -/ abbrev MulPosMonoRev : Prop := ContravariantClass α>0 α (fun x y => y * x) (· ≤ ·) +#align mul_pos_mono_rev MulPosMonoRev end Abbreviations @@ -148,93 +154,119 @@ instance MulPosReflectLT.to_contravariantClass_pos_mul_lt [MulPosReflectLT α] : theorem mul_le_mul_of_nonneg_left [PosMulMono α] (h : b ≤ c) (a0 : 0 ≤ a) : a * b ≤ a * c := @CovariantClass.elim α≥0 α (fun x y => x * y) (· ≤ ·) _ ⟨a, a0⟩ _ _ h +#align mul_le_mul_of_nonneg_left mul_le_mul_of_nonneg_left theorem mul_le_mul_of_nonneg_right [MulPosMono α] (h : b ≤ c) (a0 : 0 ≤ a) : b * a ≤ c * a := @CovariantClass.elim α≥0 α (fun x y => y * x) (· ≤ ·) _ ⟨a, a0⟩ _ _ h +#align mul_le_mul_of_nonneg_right mul_le_mul_of_nonneg_right theorem mul_lt_mul_of_pos_left [PosMulStrictMono α] (bc : b < c) (a0 : 0 < a) : a * b < a * c := @CovariantClass.elim α>0 α (fun x y => x * y) (· < ·) _ ⟨a, a0⟩ _ _ bc +#align mul_lt_mul_of_pos_left mul_lt_mul_of_pos_left theorem mul_lt_mul_of_pos_right [MulPosStrictMono α] (bc : b < c) (a0 : 0 < a) : b * a < c * a := @CovariantClass.elim α>0 α (fun x y => y * x) (· < ·) _ ⟨a, a0⟩ _ _ bc +#align mul_lt_mul_of_pos_right mul_lt_mul_of_pos_right theorem lt_of_mul_lt_mul_left [PosMulReflectLT α] (h : a * b < a * c) (a0 : 0 ≤ a) : b < c := @ContravariantClass.elim α≥0 α (fun x y => x * y) (· < ·) _ ⟨a, a0⟩ _ _ h +#align lt_of_mul_lt_mul_left lt_of_mul_lt_mul_left theorem lt_of_mul_lt_mul_right [MulPosReflectLT α] (h : b * a < c * a) (a0 : 0 ≤ a) : b < c := @ContravariantClass.elim α≥0 α (fun x y => y * x) (· < ·) _ ⟨a, a0⟩ _ _ h +#align lt_of_mul_lt_mul_right lt_of_mul_lt_mul_right theorem le_of_mul_le_mul_left [PosMulMonoRev α] (bc : a * b ≤ a * c) (a0 : 0 < a) : b ≤ c := @ContravariantClass.elim α>0 α (fun x y => x * y) (· ≤ ·) _ ⟨a, a0⟩ _ _ bc +#align le_of_mul_le_mul_left le_of_mul_le_mul_left theorem le_of_mul_le_mul_right [MulPosMonoRev α] (bc : b * a ≤ c * a) (a0 : 0 < a) : b ≤ c := @ContravariantClass.elim α>0 α (fun x y => y * x) (· ≤ ·) _ ⟨a, a0⟩ _ _ bc +#align le_of_mul_le_mul_right le_of_mul_le_mul_right alias lt_of_mul_lt_mul_left ← lt_of_mul_lt_mul_of_nonneg_left +#align lt_of_mul_lt_mul_of_nonneg_left lt_of_mul_lt_mul_of_nonneg_left alias lt_of_mul_lt_mul_right ← lt_of_mul_lt_mul_of_nonneg_right +#align lt_of_mul_lt_mul_of_nonneg_right lt_of_mul_lt_mul_of_nonneg_right alias le_of_mul_le_mul_left ← le_of_mul_le_mul_of_pos_left +#align le_of_mul_le_mul_of_pos_left le_of_mul_le_mul_of_pos_left alias le_of_mul_le_mul_right ← le_of_mul_le_mul_of_pos_right +#align le_of_mul_le_mul_of_pos_right le_of_mul_le_mul_of_pos_right @[simp] theorem mul_lt_mul_left [PosMulStrictMono α] [PosMulReflectLT α] (a0 : 0 < a) : a * b < a * c ↔ b < c := @rel_iff_cov α>0 α (fun x y => x * y) (· < ·) _ _ ⟨a, a0⟩ _ _ +#align mul_lt_mul_left mul_lt_mul_left @[simp] theorem mul_lt_mul_right [MulPosStrictMono α] [MulPosReflectLT α] (a0 : 0 < a) : b * a < c * a ↔ b < c := @rel_iff_cov α>0 α (fun x y => y * x) (· < ·) _ _ ⟨a, a0⟩ _ _ +#align mul_lt_mul_right mul_lt_mul_right @[simp] theorem mul_le_mul_left [PosMulMono α] [PosMulMonoRev α] (a0 : 0 < a) : a * b ≤ a * c ↔ b ≤ c := @rel_iff_cov α>0 α (fun x y => x * y) (· ≤ ·) _ _ ⟨a, a0⟩ _ _ +#align mul_le_mul_left mul_le_mul_left @[simp] theorem mul_le_mul_right [MulPosMono α] [MulPosMonoRev α] (a0 : 0 < a) : b * a ≤ c * a ↔ b ≤ c := @rel_iff_cov α>0 α (fun x y => y * x) (· ≤ ·) _ _ ⟨a, a0⟩ _ _ +#align mul_le_mul_right mul_le_mul_right theorem mul_lt_mul_of_pos_of_nonneg [PosMulStrictMono α] [MulPosMono α] (h₁ : a ≤ b) (h₂ : c < d) (a0 : 0 < a) (d0 : 0 ≤ d) : a * c < b * d := (mul_lt_mul_of_pos_left h₂ a0).trans_le (mul_le_mul_of_nonneg_right h₁ d0) +#align mul_lt_mul_of_pos_of_nonneg mul_lt_mul_of_pos_of_nonneg theorem mul_lt_mul_of_le_of_le' [PosMulStrictMono α] [MulPosMono α] (h₁ : a ≤ b) (h₂ : c < d) (b0 : 0 < b) (c0 : 0 ≤ c) : a * c < b * d := (mul_le_mul_of_nonneg_right h₁ c0).trans_lt (mul_lt_mul_of_pos_left h₂ b0) +#align mul_lt_mul_of_le_of_le' mul_lt_mul_of_le_of_le' theorem mul_lt_mul_of_nonneg_of_pos [PosMulMono α] [MulPosStrictMono α] (h₁ : a < b) (h₂ : c ≤ d) (a0 : 0 ≤ a) (d0 : 0 < d) : a * c < b * d := (mul_le_mul_of_nonneg_left h₂ a0).trans_lt (mul_lt_mul_of_pos_right h₁ d0) +#align mul_lt_mul_of_nonneg_of_pos mul_lt_mul_of_nonneg_of_pos theorem mul_lt_mul_of_le_of_lt' [PosMulMono α] [MulPosStrictMono α] (h₁ : a < b) (h₂ : c ≤ d) (b0 : 0 ≤ b) (c0 : 0 < c) : a * c < b * d := (mul_lt_mul_of_pos_right h₁ c0).trans_le (mul_le_mul_of_nonneg_left h₂ b0) +#align mul_lt_mul_of_le_of_lt' mul_lt_mul_of_le_of_lt' theorem mul_lt_mul_of_pos_of_pos [PosMulStrictMono α] [MulPosStrictMono α] (h₁ : a < b) (h₂ : c < d) (a0 : 0 < a) (d0 : 0 < d) : a * c < b * d := (mul_lt_mul_of_pos_left h₂ a0).trans (mul_lt_mul_of_pos_right h₁ d0) +#align mul_lt_mul_of_pos_of_pos mul_lt_mul_of_pos_of_pos theorem mul_lt_mul_of_lt_of_lt' [PosMulStrictMono α] [MulPosStrictMono α] (h₁ : a < b) (h₂ : c < d) (b0 : 0 < b) (c0 : 0 < c) : a * c < b * d := (mul_lt_mul_of_pos_right h₁ c0).trans (mul_lt_mul_of_pos_left h₂ b0) +#align mul_lt_mul_of_lt_of_lt' mul_lt_mul_of_lt_of_lt' theorem mul_lt_of_mul_lt_of_nonneg_left [PosMulMono α] (h : a * b < c) (hdb : d ≤ b) (ha : 0 ≤ a) : a * d < c := (mul_le_mul_of_nonneg_left hdb ha).trans_lt h +#align mul_lt_of_mul_lt_of_nonneg_left mul_lt_of_mul_lt_of_nonneg_left theorem lt_mul_of_lt_mul_of_nonneg_left [PosMulMono α] (h : a < b * c) (hcd : c ≤ d) (hb : 0 ≤ b) : a < b * d := h.trans_le <| mul_le_mul_of_nonneg_left hcd hb +#align lt_mul_of_lt_mul_of_nonneg_left lt_mul_of_lt_mul_of_nonneg_left theorem mul_lt_of_mul_lt_of_nonneg_right [MulPosMono α] (h : a * b < c) (hda : d ≤ a) (hb : 0 ≤ b) : d * b < c := (mul_le_mul_of_nonneg_right hda hb).trans_lt h +#align mul_lt_of_mul_lt_of_nonneg_right mul_lt_of_mul_lt_of_nonneg_right theorem lt_mul_of_lt_mul_of_nonneg_right [MulPosMono α] (h : a < b * c) (hbd : b ≤ d) (hc : 0 ≤ c) : a < d * c := h.trans_le <| mul_le_mul_of_nonneg_right hbd hc +#align lt_mul_of_lt_mul_of_nonneg_right lt_mul_of_lt_mul_of_nonneg_right end Preorder @@ -254,15 +286,19 @@ instance (priority := 100) MulPosStrictMono.toMulPosMonoRev [MulPosStrictMono α theorem PosMulMonoRev.toPosMulStrictMono [PosMulMonoRev α] : PosMulStrictMono α := ⟨fun x _ _ h => lt_of_not_ge fun h' => h.not_le <| le_of_mul_le_mul_of_pos_left h' x.prop⟩ +#align pos_mul_mono_rev.to_pos_mul_strict_mono PosMulMonoRev.toPosMulStrictMono theorem MulPosMonoRev.toMulPosStrictMono [MulPosMonoRev α] : MulPosStrictMono α := ⟨fun x _ _ h => lt_of_not_ge fun h' => h.not_le <| le_of_mul_le_mul_of_pos_right h' x.prop⟩ +#align mul_pos_mono_rev.to_mul_pos_strict_mono MulPosMonoRev.toMulPosStrictMono theorem posMulStrictMono_iff_posMulMonoRev : PosMulStrictMono α ↔ PosMulMonoRev α := ⟨@PosMulStrictMono.toPosMulMonoRev _ _ _ _, @PosMulMonoRev.toPosMulStrictMono _ _ _ _⟩ +#align pos_mul_strict_mono_iff_pos_mul_mono_rev posMulStrictMono_iff_posMulMonoRev theorem mulPosStrictMono_iff_mulPosMonoRev : MulPosStrictMono α ↔ MulPosMonoRev α := ⟨@MulPosStrictMono.toMulPosMonoRev _ _ _ _, @MulPosMonoRev.toMulPosStrictMono _ _ _ _⟩ +#align mul_pos_strict_mono_iff_mul_pos_mono_rev mulPosStrictMono_iff_mulPosMonoRev theorem PosMulReflectLT.toPosMulMono [PosMulReflectLT α] : PosMulMono α := ⟨fun x _ _ h => le_of_not_lt fun h' => h.not_lt <| lt_of_mul_lt_mul_left h' x.prop⟩ @@ -274,15 +310,19 @@ theorem MulPosReflectLT.toMulPosMono [MulPosReflectLT α] : MulPosMono α := theorem PosMulMono.toPosMulReflectLT [PosMulMono α] : PosMulReflectLT α := ⟨fun x _ _ h => lt_of_not_ge fun h' => h.not_le <| mul_le_mul_of_nonneg_left h' x.prop⟩ +#align pos_mul_mono.to_pos_mul_reflect_lt PosMulMono.toPosMulReflectLT theorem MulPosMono.toMulPosReflectLT [MulPosMono α] : MulPosReflectLT α := ⟨fun x _ _ h => lt_of_not_ge fun h' => h.not_le <| mul_le_mul_of_nonneg_right h' x.prop⟩ +#align mul_pos_mono.to_mul_pos_reflect_lt MulPosMono.toMulPosReflectLT theorem posMulMono_iff_posMulReflectLT : PosMulMono α ↔ PosMulReflectLT α := ⟨@PosMulMono.toPosMulReflectLT _ _ _ _, @PosMulReflectLT.toPosMulMono _ _ _ _⟩ +#align pos_mul_mono_iff_pos_mul_reflect_lt posMulMono_iff_posMulReflectLT theorem mulPosMono_iff_mulPosReflectLT : MulPosMono α ↔ MulPosReflectLT α := ⟨@MulPosMono.toMulPosReflectLT _ _ _ _, @MulPosReflectLT.toMulPosMono _ _ _ _⟩ +#align mul_pos_mono_iff_mul_pos_reflect_lt mulPosMono_iff_mulPosReflectLT end LinearOrder @@ -299,84 +339,106 @@ variable [Preorder α] /-- Assumes left covariance. -/ theorem Left.mul_pos [PosMulStrictMono α] (ha : 0 < a) (hb : 0 < b) : 0 < a * b := by simpa only [mul_zero] using mul_lt_mul_of_pos_left hb ha +#align left.mul_pos Left.mul_pos alias Left.mul_pos ← mul_pos +#align mul_pos mul_pos theorem mul_neg_of_pos_of_neg [PosMulStrictMono α] (ha : 0 < a) (hb : b < 0) : a * b < 0 := by simpa only [mul_zero] using mul_lt_mul_of_pos_left hb ha +#align mul_neg_of_pos_of_neg mul_neg_of_pos_of_neg @[simp] theorem zero_lt_mul_left [PosMulStrictMono α] [PosMulReflectLT α] (h : 0 < c) : 0 < c * b ↔ 0 < b := by rw [←mul_zero c, mul_lt_mul_left h] simp +#align zero_lt_mul_left zero_lt_mul_left /-- Assumes right covariance. -/ theorem Right.mul_pos [MulPosStrictMono α] (ha : 0 < a) (hb : 0 < b) : 0 < a * b := by simpa only [zero_mul] using mul_lt_mul_of_pos_right ha hb +#align right.mul_pos Right.mul_pos theorem mul_neg_of_neg_of_pos [MulPosStrictMono α] (ha : a < 0) (hb : 0 < b) : a * b < 0 := by simpa only [zero_mul] using mul_lt_mul_of_pos_right ha hb +#align mul_neg_of_neg_of_pos mul_neg_of_neg_of_pos @[simp] theorem zero_lt_mul_right [MulPosStrictMono α] [MulPosReflectLT α] (h : 0 < c) : 0 < b * c ↔ 0 < b := by rw [←zero_mul c, mul_lt_mul_right h] simp +#align zero_lt_mul_right zero_lt_mul_right /-- Assumes left covariance. -/ theorem Left.mul_nonneg [PosMulMono α] (ha : 0 ≤ a) (hb : 0 ≤ b) : 0 ≤ a * b := by simpa only [mul_zero] using mul_le_mul_of_nonneg_left hb ha +#align left.mul_nonneg Left.mul_nonneg alias Left.mul_nonneg ← mul_nonneg +#align mul_nonneg mul_nonneg theorem mul_nonpos_of_nonneg_of_nonpos [PosMulMono α] (ha : 0 ≤ a) (hb : b ≤ 0) : a * b ≤ 0 := by simpa only [mul_zero] using mul_le_mul_of_nonneg_left hb ha +#align mul_nonpos_of_nonneg_of_nonpos mul_nonpos_of_nonneg_of_nonpos /-- Assumes right covariance. -/ theorem Right.mul_nonneg [MulPosMono α] (ha : 0 ≤ a) (hb : 0 ≤ b) : 0 ≤ a * b := by simpa only [zero_mul] using mul_le_mul_of_nonneg_right ha hb +#align right.mul_nonneg Right.mul_nonneg theorem mul_nonpos_of_nonpos_of_nonneg [MulPosMono α] (ha : a ≤ 0) (hb : 0 ≤ b) : a * b ≤ 0 := by simpa only [zero_mul] using mul_le_mul_of_nonneg_right ha hb +#align mul_nonpos_of_nonpos_of_nonneg mul_nonpos_of_nonpos_of_nonneg theorem pos_of_mul_pos_right [PosMulReflectLT α] (h : 0 < a * b) (ha : 0 ≤ a) : 0 < b := lt_of_mul_lt_mul_left ((mul_zero a).symm ▸ h : a * 0 < a * b) ha +#align pos_of_mul_pos_right pos_of_mul_pos_right theorem pos_of_mul_pos_left [MulPosReflectLT α] (h : 0 < a * b) (hb : 0 ≤ b) : 0 < a := lt_of_mul_lt_mul_right ((zero_mul b).symm ▸ h : 0 * b < a * b) hb +#align pos_of_mul_pos_left pos_of_mul_pos_left theorem pos_iff_pos_of_mul_pos [PosMulReflectLT α] [MulPosReflectLT α] (hab : 0 < a * b) : 0 < a ↔ 0 < b := ⟨pos_of_mul_pos_right hab ∘ le_of_lt, pos_of_mul_pos_left hab ∘ le_of_lt⟩ +#align pos_iff_pos_of_mul_pos pos_iff_pos_of_mul_pos theorem mul_le_mul_of_le_of_le [PosMulMono α] [MulPosMono α] (h₁ : a ≤ b) (h₂ : c ≤ d) (a0 : 0 ≤ a) (d0 : 0 ≤ d) : a * c ≤ b * d := (mul_le_mul_of_nonneg_left h₂ a0).trans <| mul_le_mul_of_nonneg_right h₁ d0 +#align mul_le_mul_of_le_of_le mul_le_mul_of_le_of_le theorem mul_le_mul [PosMulMono α] [MulPosMono α] (h₁ : a ≤ b) (h₂ : c ≤ d) (c0 : 0 ≤ c) (b0 : 0 ≤ b) : a * c ≤ b * d := (mul_le_mul_of_nonneg_right h₁ c0).trans <| mul_le_mul_of_nonneg_left h₂ b0 +#align mul_le_mul mul_le_mul theorem mul_self_le_mul_self [PosMulMono α] [MulPosMono α] (ha : 0 ≤ a) (hab : a ≤ b) : a * a ≤ b * b := mul_le_mul hab hab ha <| ha.trans hab +#align mul_self_le_mul_self mul_self_le_mul_self theorem mul_le_of_mul_le_of_nonneg_left [PosMulMono α] (h : a * b ≤ c) (hle : d ≤ b) (a0 : 0 ≤ a) : a * d ≤ c := (mul_le_mul_of_nonneg_left hle a0).trans h +#align mul_le_of_mul_le_of_nonneg_left mul_le_of_mul_le_of_nonneg_left theorem le_mul_of_le_mul_of_nonneg_left [PosMulMono α] (h : a ≤ b * c) (hle : c ≤ d) (b0 : 0 ≤ b) : a ≤ b * d := h.trans (mul_le_mul_of_nonneg_left hle b0) +#align le_mul_of_le_mul_of_nonneg_left le_mul_of_le_mul_of_nonneg_left theorem mul_le_of_mul_le_of_nonneg_right [MulPosMono α] (h : a * b ≤ c) (hle : d ≤ a) (b0 : 0 ≤ b) : d * b ≤ c := (mul_le_mul_of_nonneg_right hle b0).trans h +#align mul_le_of_mul_le_of_nonneg_right mul_le_of_mul_le_of_nonneg_right theorem le_mul_of_le_mul_of_nonneg_right [MulPosMono α] (h : a ≤ b * c) (hle : b ≤ d) (c0 : 0 ≤ c) : a ≤ d * c := h.trans (mul_le_mul_of_nonneg_right hle c0) +#align le_mul_of_le_mul_of_nonneg_right le_mul_of_le_mul_of_nonneg_right end Preorder @@ -459,10 +521,12 @@ instance (priority := 100) MulPosMonoRev.toMulPosReflectLT [MulPosMonoRev α] : theorem mul_left_cancel_iff_of_pos [PosMulMonoRev α] (a0 : 0 < a) : a * b = a * c ↔ b = c := ⟨fun h => (le_of_mul_le_mul_of_pos_left h.le a0).antisymm <| le_of_mul_le_mul_of_pos_left h.ge a0, congr_arg _⟩ +#align mul_left_cancel_iff_of_pos mul_left_cancel_iff_of_pos theorem mul_right_cancel_iff_of_pos [MulPosMonoRev α] (b0 : 0 < b) : a * b = c * b ↔ a = c := ⟨fun h => (le_of_mul_le_mul_of_pos_right h.le b0).antisymm <| le_of_mul_le_mul_of_pos_right h.ge b0, congr_arg (· * b)⟩ +#align mul_right_cancel_iff_of_pos mul_right_cancel_iff_of_pos theorem mul_eq_mul_iff_eq_and_eq_of_pos [PosMulStrictMono α] [MulPosStrictMono α] [PosMulMonoRev α] [MulPosMonoRev α] (hac : a ≤ b) (hbd : c ≤ d) (a0 : 0 < a) (d0 : 0 < d) : @@ -470,11 +534,10 @@ theorem mul_eq_mul_iff_eq_and_eq_of_pos [PosMulStrictMono α] [MulPosStrictMono refine' ⟨fun h => _, fun h => congr_arg₂ (· * ·) h.1 h.2⟩ rcases hac.eq_or_lt with (rfl | hac) · exact ⟨rfl, (mul_left_cancel_iff_of_pos a0).mp h⟩ - rcases eq_or_lt_of_le hbd with (rfl | hbd) · exact ⟨(mul_right_cancel_iff_of_pos d0).mp h, rfl⟩ - exact ((mul_lt_mul_of_pos_of_pos hac hbd a0 d0).ne h).elim +#align mul_eq_mul_iff_eq_and_eq_of_pos mul_eq_mul_iff_eq_and_eq_of_pos theorem mul_eq_mul_iff_eq_and_eq_of_pos' [PosMulStrictMono α] [MulPosStrictMono α] [PosMulMonoRev α] [MulPosMonoRev α] (hac : a ≤ b) (hbd : c ≤ d) (b0 : 0 < b) (c0 : 0 < c) : @@ -482,11 +545,10 @@ theorem mul_eq_mul_iff_eq_and_eq_of_pos' [PosMulStrictMono α] [MulPosStrictMono refine' ⟨fun h => _, fun h => congr_arg₂ (· * ·) h.1 h.2⟩ rcases hac.eq_or_lt with (rfl | hac) · exact ⟨rfl, (mul_left_cancel_iff_of_pos b0).mp h⟩ - rcases eq_or_lt_of_le hbd with (rfl | hbd) · exact ⟨(mul_right_cancel_iff_of_pos c0).mp h, rfl⟩ - exact ((mul_lt_mul_of_lt_of_lt' hac hbd b0 c0).ne h).elim +#align mul_eq_mul_iff_eq_and_eq_of_pos' mul_eq_mul_iff_eq_and_eq_of_pos' end PartialOrder @@ -503,28 +565,36 @@ theorem pos_and_pos_or_neg_and_neg_of_mul_pos [PosMulMono α] [MulPosMono α] (h exact hab.false.elim · refine' Or.inl ⟨ha, lt_imp_lt_of_le_imp_le (fun hb => _) hab⟩ exact mul_nonpos_of_nonneg_of_nonpos ha.le hb +#align pos_and_pos_or_neg_and_neg_of_mul_pos pos_and_pos_or_neg_and_neg_of_mul_pos theorem neg_of_mul_pos_right [PosMulMono α] [MulPosMono α] (h : 0 < a * b) (ha : a ≤ 0) : b < 0 := ((pos_and_pos_or_neg_and_neg_of_mul_pos h).resolve_left fun h => h.1.not_le ha).2 +#align neg_of_mul_pos_right neg_of_mul_pos_right theorem neg_of_mul_pos_left [PosMulMono α] [MulPosMono α] (h : 0 < a * b) (ha : b ≤ 0) : a < 0 := ((pos_and_pos_or_neg_and_neg_of_mul_pos h).resolve_left fun h => h.2.not_le ha).1 +#align neg_of_mul_pos_left neg_of_mul_pos_left theorem neg_iff_neg_of_mul_pos [PosMulMono α] [MulPosMono α] (hab : 0 < a * b) : a < 0 ↔ b < 0 := ⟨neg_of_mul_pos_right hab ∘ le_of_lt, neg_of_mul_pos_left hab ∘ le_of_lt⟩ +#align neg_iff_neg_of_mul_pos neg_iff_neg_of_mul_pos theorem Left.neg_of_mul_neg_left [PosMulMono α] (h : a * b < 0) (h1 : 0 ≤ a) : b < 0 := lt_of_not_ge fun h2 : b ≥ 0 => (Left.mul_nonneg h1 h2).not_lt h +#align left.neg_of_mul_neg_left Left.neg_of_mul_neg_left theorem Right.neg_of_mul_neg_left [MulPosMono α] (h : a * b < 0) (h1 : 0 ≤ a) : b < 0 := lt_of_not_ge fun h2 : b ≥ 0 => (Right.mul_nonneg h1 h2).not_lt h +#align right.neg_of_mul_neg_left Right.neg_of_mul_neg_left theorem Left.neg_of_mul_neg_right [PosMulMono α] (h : a * b < 0) (h1 : 0 ≤ b) : a < 0 := lt_of_not_ge fun h2 : a ≥ 0 => (Left.mul_nonneg h2 h1).not_lt h +#align left.neg_of_mul_neg_right Left.neg_of_mul_neg_right theorem Right.neg_of_mul_neg_right [MulPosMono α] (h : a * b < 0) (h1 : 0 ≤ b) : a < 0 := lt_of_not_ge fun h2 : a ≥ 0 => (Right.mul_nonneg h2 h1).not_lt h +#align right.neg_of_mul_neg_right Right.neg_of_mul_neg_right end LinearOrder @@ -545,20 +615,24 @@ which assume left covariance. -/ @[simp] theorem le_mul_iff_one_le_right [PosMulMono α] [PosMulMonoRev α] (a0 : 0 < a) : a ≤ a * b ↔ 1 ≤ b := Iff.trans (by rw [mul_one]) (mul_le_mul_left a0) +#align le_mul_iff_one_le_right le_mul_iff_one_le_right @[simp] theorem lt_mul_iff_one_lt_right [PosMulStrictMono α] [PosMulReflectLT α] (a0 : 0 < a) : a < a * b ↔ 1 < b := Iff.trans (by rw [mul_one]) (mul_lt_mul_left a0) +#align lt_mul_iff_one_lt_right lt_mul_iff_one_lt_right @[simp] theorem mul_le_iff_le_one_right [PosMulMono α] [PosMulMonoRev α] (a0 : 0 < a) : a * b ≤ a ↔ b ≤ 1 := Iff.trans (by rw [mul_one]) (mul_le_mul_left a0) +#align mul_le_iff_le_one_right mul_le_iff_le_one_right @[simp] theorem mul_lt_iff_lt_one_right [PosMulStrictMono α] [PosMulReflectLT α] (a0 : 0 < a) : a * b < a ↔ b < 1 := Iff.trans (by rw [mul_one]) (mul_lt_mul_left a0) +#align mul_lt_iff_lt_one_right mul_lt_iff_lt_one_right /-! Lemmas of the form `a ≤ b * a ↔ 1 ≤ b` and `a * b ≤ b ↔ a ≤ 1`, which assume right covariance. -/ @@ -567,47 +641,59 @@ which assume right covariance. -/ @[simp] theorem le_mul_iff_one_le_left [MulPosMono α] [MulPosMonoRev α] (a0 : 0 < a) : a ≤ b * a ↔ 1 ≤ b := Iff.trans (by rw [one_mul]) (mul_le_mul_right a0) +#align le_mul_iff_one_le_left le_mul_iff_one_le_left @[simp] theorem lt_mul_iff_one_lt_left [MulPosStrictMono α] [MulPosReflectLT α] (a0 : 0 < a) : a < b * a ↔ 1 < b := Iff.trans (by rw [one_mul]) (mul_lt_mul_right a0) +#align lt_mul_iff_one_lt_left lt_mul_iff_one_lt_left @[simp] theorem mul_le_iff_le_one_left [MulPosMono α] [MulPosMonoRev α] (b0 : 0 < b) : a * b ≤ b ↔ a ≤ 1 := Iff.trans (by rw [one_mul]) (mul_le_mul_right b0) +#align mul_le_iff_le_one_left mul_le_iff_le_one_left @[simp] theorem mul_lt_iff_lt_one_left [MulPosStrictMono α] [MulPosReflectLT α] (b0 : 0 < b) : a * b < b ↔ a < 1 := Iff.trans (by rw [one_mul]) (mul_lt_mul_right b0) +#align mul_lt_iff_lt_one_left mul_lt_iff_lt_one_left /-! Lemmas of the form `1 ≤ b → a ≤ a * b`. -/ theorem mul_le_of_le_one_left [MulPosMono α] (hb : 0 ≤ b) (h : a ≤ 1) : a * b ≤ b := by simpa only [one_mul] using mul_le_mul_of_nonneg_right h hb +#align mul_le_of_le_one_left mul_le_of_le_one_left theorem le_mul_of_one_le_left [MulPosMono α] (hb : 0 ≤ b) (h : 1 ≤ a) : b ≤ a * b := by simpa only [one_mul] using mul_le_mul_of_nonneg_right h hb +#align le_mul_of_one_le_left le_mul_of_one_le_left theorem mul_le_of_le_one_right [PosMulMono α] (ha : 0 ≤ a) (h : b ≤ 1) : a * b ≤ a := by simpa only [mul_one] using mul_le_mul_of_nonneg_left h ha +#align mul_le_of_le_one_right mul_le_of_le_one_right theorem le_mul_of_one_le_right [PosMulMono α] (ha : 0 ≤ a) (h : 1 ≤ b) : a ≤ a * b := by simpa only [mul_one] using mul_le_mul_of_nonneg_left h ha +#align le_mul_of_one_le_right le_mul_of_one_le_right theorem mul_lt_of_lt_one_left [MulPosStrictMono α] (hb : 0 < b) (h : a < 1) : a * b < b := by simpa only [one_mul] using mul_lt_mul_of_pos_right h hb +#align mul_lt_of_lt_one_left mul_lt_of_lt_one_left theorem lt_mul_of_one_lt_left [MulPosStrictMono α] (hb : 0 < b) (h : 1 < a) : b < a * b := by simpa only [one_mul] using mul_lt_mul_of_pos_right h hb +#align lt_mul_of_one_lt_left lt_mul_of_one_lt_left theorem mul_lt_of_lt_one_right [PosMulStrictMono α] (ha : 0 < a) (h : b < 1) : a * b < a := by simpa only [mul_one] using mul_lt_mul_of_pos_left h ha +#align mul_lt_of_lt_one_right mul_lt_of_lt_one_right theorem lt_mul_of_one_lt_right [PosMulStrictMono α] (ha : 0 < a) (h : 1 < b) : a < a * b := by simpa only [mul_one] using mul_lt_mul_of_pos_left h ha +#align lt_mul_of_one_lt_right lt_mul_of_one_lt_right /-! Lemmas of the form `b ≤ c → a ≤ 1 → b * a ≤ c`. -/ @@ -618,45 +704,55 @@ to find -/ theorem mul_le_of_le_of_le_one_of_nonneg [PosMulMono α] (h : b ≤ c) (ha : a ≤ 1) (hb : 0 ≤ b) : b * a ≤ c := (mul_le_of_le_one_right hb ha).trans h +#align mul_le_of_le_of_le_one_of_nonneg mul_le_of_le_of_le_one_of_nonneg theorem mul_lt_of_le_of_lt_one_of_pos [PosMulStrictMono α] (bc : b ≤ c) (ha : a < 1) (b0 : 0 < b) : b * a < c := (mul_lt_of_lt_one_right b0 ha).trans_le bc +#align mul_lt_of_le_of_lt_one_of_pos mul_lt_of_le_of_lt_one_of_pos theorem mul_lt_of_lt_of_le_one_of_nonneg [PosMulMono α] (h : b < c) (ha : a ≤ 1) (hb : 0 ≤ b) : b * a < c := (mul_le_of_le_one_right hb ha).trans_lt h +#align mul_lt_of_lt_of_le_one_of_nonneg mul_lt_of_lt_of_le_one_of_nonneg /-- Assumes left covariance. -/ theorem Left.mul_le_one_of_le_of_le [PosMulMono α] (ha : a ≤ 1) (hb : b ≤ 1) (a0 : 0 ≤ a) : a * b ≤ 1 := mul_le_of_le_of_le_one_of_nonneg ha hb a0 +#align left.mul_le_one_of_le_of_le Left.mul_le_one_of_le_of_le /-- Assumes left covariance. -/ theorem Left.mul_lt_of_le_of_lt_one_of_pos [PosMulStrictMono α] (ha : a ≤ 1) (hb : b < 1) (a0 : 0 < a) : a * b < 1 := _root_.mul_lt_of_le_of_lt_one_of_pos ha hb a0 +#align left.mul_lt_of_le_of_lt_one_of_pos Left.mul_lt_of_le_of_lt_one_of_pos /-- Assumes left covariance. -/ theorem Left.mul_lt_of_lt_of_le_one_of_nonneg [PosMulMono α] (ha : a < 1) (hb : b ≤ 1) (a0 : 0 ≤ a) : a * b < 1 := _root_.mul_lt_of_lt_of_le_one_of_nonneg ha hb a0 +#align left.mul_lt_of_lt_of_le_one_of_nonneg Left.mul_lt_of_lt_of_le_one_of_nonneg theorem mul_le_of_le_of_le_one' [PosMulMono α] [MulPosMono α] (bc : b ≤ c) (ha : a ≤ 1) (a0 : 0 ≤ a) (c0 : 0 ≤ c) : b * a ≤ c := (mul_le_mul_of_nonneg_right bc a0).trans <| mul_le_of_le_one_right c0 ha +#align mul_le_of_le_of_le_one' mul_le_of_le_of_le_one' theorem mul_lt_of_lt_of_le_one' [PosMulMono α] [MulPosStrictMono α] (bc : b < c) (ha : a ≤ 1) (a0 : 0 < a) (c0 : 0 ≤ c) : b * a < c := (mul_lt_mul_of_pos_right bc a0).trans_le <| mul_le_of_le_one_right c0 ha +#align mul_lt_of_lt_of_le_one' mul_lt_of_lt_of_le_one' theorem mul_lt_of_le_of_lt_one' [PosMulStrictMono α] [MulPosMono α] (bc : b ≤ c) (ha : a < 1) (a0 : 0 ≤ a) (c0 : 0 < c) : b * a < c := (mul_le_mul_of_nonneg_right bc a0).trans_lt <| mul_lt_of_lt_one_right c0 ha +#align mul_lt_of_le_of_lt_one' mul_lt_of_le_of_lt_one' theorem mul_lt_of_lt_of_lt_one_of_pos [PosMulMono α] [MulPosStrictMono α] (bc : b < c) (ha : a ≤ 1) (a0 : 0 < a) (c0 : 0 ≤ c) : b * a < c := (mul_lt_mul_of_pos_right bc a0).trans_le <| mul_le_of_le_one_right c0 ha +#align mul_lt_of_lt_of_lt_one_of_pos mul_lt_of_lt_of_lt_one_of_pos /-! Lemmas of the form `b ≤ c → 1 ≤ a → b ≤ c * a`. -/ @@ -664,45 +760,55 @@ theorem mul_lt_of_lt_of_lt_one_of_pos [PosMulMono α] [MulPosStrictMono α] (bc theorem le_mul_of_le_of_one_le_of_nonneg [PosMulMono α] (h : b ≤ c) (ha : 1 ≤ a) (hc : 0 ≤ c) : b ≤ c * a := h.trans <| le_mul_of_one_le_right hc ha +#align le_mul_of_le_of_one_le_of_nonneg le_mul_of_le_of_one_le_of_nonneg theorem lt_mul_of_le_of_one_lt_of_pos [PosMulStrictMono α] (bc : b ≤ c) (ha : 1 < a) (c0 : 0 < c) : b < c * a := bc.trans_lt <| lt_mul_of_one_lt_right c0 ha +#align lt_mul_of_le_of_one_lt_of_pos lt_mul_of_le_of_one_lt_of_pos theorem lt_mul_of_lt_of_one_le_of_nonneg [PosMulMono α] (h : b < c) (ha : 1 ≤ a) (hc : 0 ≤ c) : b < c * a := h.trans_le <| le_mul_of_one_le_right hc ha +#align lt_mul_of_lt_of_one_le_of_nonneg lt_mul_of_lt_of_one_le_of_nonneg /-- Assumes left covariance. -/ theorem Left.one_le_mul_of_le_of_le [PosMulMono α] (ha : 1 ≤ a) (hb : 1 ≤ b) (a0 : 0 ≤ a) : 1 ≤ a * b := le_mul_of_le_of_one_le_of_nonneg ha hb a0 +#align left.one_le_mul_of_le_of_le Left.one_le_mul_of_le_of_le /-- Assumes left covariance. -/ theorem Left.one_lt_mul_of_le_of_lt_of_pos [PosMulStrictMono α] (ha : 1 ≤ a) (hb : 1 < b) (a0 : 0 < a) : 1 < a * b := lt_mul_of_le_of_one_lt_of_pos ha hb a0 +#align left.one_lt_mul_of_le_of_lt_of_pos Left.one_lt_mul_of_le_of_lt_of_pos /-- Assumes left covariance. -/ theorem Left.lt_mul_of_lt_of_one_le_of_nonneg [PosMulMono α] (ha : 1 < a) (hb : 1 ≤ b) (a0 : 0 ≤ a) : 1 < a * b := _root_.lt_mul_of_lt_of_one_le_of_nonneg ha hb a0 +#align left.lt_mul_of_lt_of_one_le_of_nonneg Left.lt_mul_of_lt_of_one_le_of_nonneg theorem le_mul_of_le_of_one_le' [PosMulMono α] [MulPosMono α] (bc : b ≤ c) (ha : 1 ≤ a) (a0 : 0 ≤ a) (b0 : 0 ≤ b) : b ≤ c * a := (le_mul_of_one_le_right b0 ha).trans <| mul_le_mul_of_nonneg_right bc a0 +#align le_mul_of_le_of_one_le' le_mul_of_le_of_one_le' theorem lt_mul_of_le_of_one_lt' [PosMulStrictMono α] [MulPosMono α] (bc : b ≤ c) (ha : 1 < a) (a0 : 0 ≤ a) (b0 : 0 < b) : b < c * a := (lt_mul_of_one_lt_right b0 ha).trans_le <| mul_le_mul_of_nonneg_right bc a0 +#align lt_mul_of_le_of_one_lt' lt_mul_of_le_of_one_lt' theorem lt_mul_of_lt_of_one_le' [PosMulMono α] [MulPosStrictMono α] (bc : b < c) (ha : 1 ≤ a) (a0 : 0 < a) (b0 : 0 ≤ b) : b < c * a := (le_mul_of_one_le_right b0 ha).trans_lt <| mul_lt_mul_of_pos_right bc a0 +#align lt_mul_of_lt_of_one_le' lt_mul_of_lt_of_one_le' theorem lt_mul_of_lt_of_one_lt_of_pos [PosMulStrictMono α] [MulPosStrictMono α] (bc : b < c) (ha : 1 < a) (a0 : 0 < a) (b0 : 0 < b) : b < c * a := (lt_mul_of_one_lt_right b0 ha).trans <| mul_lt_mul_of_pos_right bc a0 +#align lt_mul_of_lt_of_one_lt_of_pos lt_mul_of_lt_of_one_lt_of_pos /-! Lemmas of the form `a ≤ 1 → b ≤ c → a * b ≤ c`. -/ @@ -710,45 +816,55 @@ theorem lt_mul_of_lt_of_one_lt_of_pos [PosMulStrictMono α] [MulPosStrictMono α theorem mul_le_of_le_one_of_le_of_nonneg [MulPosMono α] (ha : a ≤ 1) (h : b ≤ c) (hb : 0 ≤ b) : a * b ≤ c := (mul_le_of_le_one_left hb ha).trans h +#align mul_le_of_le_one_of_le_of_nonneg mul_le_of_le_one_of_le_of_nonneg theorem mul_lt_of_lt_one_of_le_of_pos [MulPosStrictMono α] (ha : a < 1) (h : b ≤ c) (hb : 0 < b) : a * b < c := (mul_lt_of_lt_one_left hb ha).trans_le h +#align mul_lt_of_lt_one_of_le_of_pos mul_lt_of_lt_one_of_le_of_pos theorem mul_lt_of_le_one_of_lt_of_nonneg [MulPosMono α] (ha : a ≤ 1) (h : b < c) (hb : 0 ≤ b) : a * b < c := (mul_le_of_le_one_left hb ha).trans_lt h +#align mul_lt_of_le_one_of_lt_of_nonneg mul_lt_of_le_one_of_lt_of_nonneg /-- Assumes right covariance. -/ theorem Right.mul_lt_one_of_lt_of_le_of_pos [MulPosStrictMono α] (ha : a < 1) (hb : b ≤ 1) (b0 : 0 < b) : a * b < 1 := mul_lt_of_lt_one_of_le_of_pos ha hb b0 +#align right.mul_lt_one_of_lt_of_le_of_pos Right.mul_lt_one_of_lt_of_le_of_pos /-- Assumes right covariance. -/ theorem Right.mul_lt_one_of_le_of_lt_of_nonneg [MulPosMono α] (ha : a ≤ 1) (hb : b < 1) (b0 : 0 ≤ b) : a * b < 1 := mul_lt_of_le_one_of_lt_of_nonneg ha hb b0 +#align right.mul_lt_one_of_le_of_lt_of_nonneg Right.mul_lt_one_of_le_of_lt_of_nonneg theorem mul_lt_of_lt_one_of_lt_of_pos [PosMulStrictMono α] [MulPosStrictMono α] (ha : a < 1) (bc : b < c) (a0 : 0 < a) (c0 : 0 < c) : a * b < c := (mul_lt_mul_of_pos_left bc a0).trans <| mul_lt_of_lt_one_left c0 ha +#align mul_lt_of_lt_one_of_lt_of_pos mul_lt_of_lt_one_of_lt_of_pos /-- Assumes right covariance. -/ theorem Right.mul_le_one_of_le_of_le [MulPosMono α] (ha : a ≤ 1) (hb : b ≤ 1) (b0 : 0 ≤ b) : a * b ≤ 1 := mul_le_of_le_one_of_le_of_nonneg ha hb b0 +#align right.mul_le_one_of_le_of_le Right.mul_le_one_of_le_of_le theorem mul_le_of_le_one_of_le' [PosMulMono α] [MulPosMono α] (ha : a ≤ 1) (bc : b ≤ c) (a0 : 0 ≤ a) (c0 : 0 ≤ c) : a * b ≤ c := (mul_le_mul_of_nonneg_left bc a0).trans <| mul_le_of_le_one_left c0 ha +#align mul_le_of_le_one_of_le' mul_le_of_le_one_of_le' theorem mul_lt_of_lt_one_of_le' [PosMulMono α] [MulPosStrictMono α] (ha : a < 1) (bc : b ≤ c) (a0 : 0 ≤ a) (c0 : 0 < c) : a * b < c := (mul_le_mul_of_nonneg_left bc a0).trans_lt <| mul_lt_of_lt_one_left c0 ha +#align mul_lt_of_lt_one_of_le' mul_lt_of_lt_one_of_le' theorem mul_lt_of_le_one_of_lt' [PosMulStrictMono α] [MulPosMono α] (ha : a ≤ 1) (bc : b < c) (a0 : 0 < a) (c0 : 0 ≤ c) : a * b < c := (mul_lt_mul_of_pos_left bc a0).trans_le <| mul_le_of_le_one_left c0 ha +#align mul_lt_of_le_one_of_lt' mul_lt_of_le_one_of_lt' /-! Lemmas of the form `1 ≤ a → b ≤ c → b ≤ a * c`. -/ @@ -756,70 +872,86 @@ theorem mul_lt_of_le_one_of_lt' [PosMulStrictMono α] [MulPosMono α] (ha : a theorem lt_mul_of_one_lt_of_le_of_pos [MulPosStrictMono α] (ha : 1 < a) (h : b ≤ c) (hc : 0 < c) : b < a * c := h.trans_lt <| lt_mul_of_one_lt_left hc ha +#align lt_mul_of_one_lt_of_le_of_pos lt_mul_of_one_lt_of_le_of_pos theorem lt_mul_of_one_le_of_lt_of_nonneg [MulPosMono α] (ha : 1 ≤ a) (h : b < c) (hc : 0 ≤ c) : b < a * c := h.trans_le <| le_mul_of_one_le_left hc ha +#align lt_mul_of_one_le_of_lt_of_nonneg lt_mul_of_one_le_of_lt_of_nonneg theorem lt_mul_of_one_lt_of_lt_of_pos [MulPosStrictMono α] (ha : 1 < a) (h : b < c) (hc : 0 < c) : b < a * c := h.trans <| lt_mul_of_one_lt_left hc ha +#align lt_mul_of_one_lt_of_lt_of_pos lt_mul_of_one_lt_of_lt_of_pos /-- Assumes right covariance. -/ theorem Right.one_lt_mul_of_lt_of_le_of_pos [MulPosStrictMono α] (ha : 1 < a) (hb : 1 ≤ b) (b0 : 0 < b) : 1 < a * b := lt_mul_of_one_lt_of_le_of_pos ha hb b0 +#align right.one_lt_mul_of_lt_of_le_of_pos Right.one_lt_mul_of_lt_of_le_of_pos /-- Assumes right covariance. -/ theorem Right.one_lt_mul_of_le_of_lt_of_nonneg [MulPosMono α] (ha : 1 ≤ a) (hb : 1 < b) (b0 : 0 ≤ b) : 1 < a * b := lt_mul_of_one_le_of_lt_of_nonneg ha hb b0 +#align right.one_lt_mul_of_le_of_lt_of_nonneg Right.one_lt_mul_of_le_of_lt_of_nonneg /-- Assumes right covariance. -/ theorem Right.one_lt_mul_of_lt_of_lt [MulPosStrictMono α] (ha : 1 < a) (hb : 1 < b) (b0 : 0 < b) : 1 < a * b := lt_mul_of_one_lt_of_lt_of_pos ha hb b0 +#align right.one_lt_mul_of_lt_of_lt Right.one_lt_mul_of_lt_of_lt theorem lt_mul_of_one_lt_of_lt_of_nonneg [MulPosMono α] (ha : 1 ≤ a) (h : b < c) (hc : 0 ≤ c) : b < a * c := h.trans_le <| le_mul_of_one_le_left hc ha +#align lt_mul_of_one_lt_of_lt_of_nonneg lt_mul_of_one_lt_of_lt_of_nonneg theorem lt_of_mul_lt_of_one_le_of_nonneg_left [PosMulMono α] (h : a * b < c) (hle : 1 ≤ b) (ha : 0 ≤ a) : a < c := (le_mul_of_one_le_right ha hle).trans_lt h +#align lt_of_mul_lt_of_one_le_of_nonneg_left lt_of_mul_lt_of_one_le_of_nonneg_left theorem lt_of_lt_mul_of_le_one_of_nonneg_left [PosMulMono α] (h : a < b * c) (hc : c ≤ 1) (hb : 0 ≤ b) : a < b := h.trans_le <| mul_le_of_le_one_right hb hc +#align lt_of_lt_mul_of_le_one_of_nonneg_left lt_of_lt_mul_of_le_one_of_nonneg_left theorem lt_of_lt_mul_of_le_one_of_nonneg_right [MulPosMono α] (h : a < b * c) (hb : b ≤ 1) (hc : 0 ≤ c) : a < c := h.trans_le <| mul_le_of_le_one_left hc hb +#align lt_of_lt_mul_of_le_one_of_nonneg_right lt_of_lt_mul_of_le_one_of_nonneg_right theorem le_mul_of_one_le_of_le_of_nonneg [MulPosMono α] (ha : 1 ≤ a) (bc : b ≤ c) (c0 : 0 ≤ c) : b ≤ a * c := bc.trans <| le_mul_of_one_le_left c0 ha +#align le_mul_of_one_le_of_le_of_nonneg le_mul_of_one_le_of_le_of_nonneg /-- Assumes right covariance. -/ theorem Right.one_le_mul_of_le_of_le [MulPosMono α] (ha : 1 ≤ a) (hb : 1 ≤ b) (b0 : 0 ≤ b) : 1 ≤ a * b := le_mul_of_one_le_of_le_of_nonneg ha hb b0 +#align right.one_le_mul_of_le_of_le Right.one_le_mul_of_le_of_le theorem le_of_mul_le_of_one_le_of_nonneg_left [PosMulMono α] (h : a * b ≤ c) (hb : 1 ≤ b) (ha : 0 ≤ a) : a ≤ c := (le_mul_of_one_le_right ha hb).trans h +#align le_of_mul_le_of_one_le_of_nonneg_left le_of_mul_le_of_one_le_of_nonneg_left theorem le_of_le_mul_of_le_one_of_nonneg_left [PosMulMono α] (h : a ≤ b * c) (hc : c ≤ 1) (hb : 0 ≤ b) : a ≤ b := h.trans <| mul_le_of_le_one_right hb hc +#align le_of_le_mul_of_le_one_of_nonneg_left le_of_le_mul_of_le_one_of_nonneg_left theorem le_of_mul_le_of_one_le_nonneg_right [MulPosMono α] (h : a * b ≤ c) (ha : 1 ≤ a) (hb : 0 ≤ b) : b ≤ c := (le_mul_of_one_le_left hb ha).trans h +#align le_of_mul_le_of_one_le_nonneg_right le_of_mul_le_of_one_le_nonneg_right theorem le_of_le_mul_of_le_one_of_nonneg_right [MulPosMono α] (h : a ≤ b * c) (hb : b ≤ 1) (hc : 0 ≤ c) : a ≤ c := h.trans <| mul_le_of_le_one_left hc hb +#align le_of_le_mul_of_le_one_of_nonneg_right le_of_le_mul_of_le_one_of_nonneg_right end Preorder @@ -833,6 +965,7 @@ theorem exists_square_le' [PosMulStrictMono α] (a0 : 0 < a) : ∃ b : α, b * b obtain ha | ha := lt_or_le a 1 · exact ⟨a, (mul_lt_of_lt_one_right a0 ha).le⟩ · exact ⟨1, by rwa [mul_one]⟩ +#align exists_square_le' exists_square_le' end LinearOrder diff --git a/Mathlib/Algebra/Order/SMul.lean b/Mathlib/Algebra/Order/SMul.lean index 8ae1934a3735b..06111a273b82b 100644 --- a/Mathlib/Algebra/Order/SMul.lean +++ b/Mathlib/Algebra/Order/SMul.lean @@ -270,6 +270,8 @@ def OrderIso.smulLeft (hc : 0 < c) : M ≃o M where right_inv := smul_inv_smul₀ hc.ne' map_rel_iff' := smul_le_smul_iff_of_pos hc #align order_iso.smul_left OrderIso.smulLeft +#align order_iso.smul_left_symm_apply OrderIso.smulLeft_symmApply +#align order_iso.smul_left_apply OrderIso.smulLeft_apply variable {M} diff --git a/Mathlib/Algebra/Order/Sub/Canonical.lean b/Mathlib/Algebra/Order/Sub/Canonical.lean index 23322c7b5c7f3..016ee08c8119e 100644 --- a/Mathlib/Algebra/Order/Sub/Canonical.lean +++ b/Mathlib/Algebra/Order/Sub/Canonical.lean @@ -329,6 +329,7 @@ theorem tsub_eq_zero_iff_le : a - b = 0 ↔ a ≤ b := by #align tsub_eq_zero_iff_le tsub_eq_zero_iff_le alias tsub_eq_zero_iff_le ↔ _ tsub_eq_zero_of_le +#align tsub_eq_zero_of_le tsub_eq_zero_of_le attribute [simp] tsub_eq_zero_of_le diff --git a/Mathlib/Algebra/Order/Sub/Defs.lean b/Mathlib/Algebra/Order/Sub/Defs.lean index ce64a92ece3d9..548c361dae8fb 100644 --- a/Mathlib/Algebra/Order/Sub/Defs.lean +++ b/Mathlib/Algebra/Order/Sub/Defs.lean @@ -252,6 +252,7 @@ theorem tsub_nonpos : a - b ≤ 0 ↔ a ≤ b := by rw [tsub_le_iff_left, add_ze #align tsub_nonpos tsub_nonpos alias tsub_nonpos ↔ _ tsub_nonpos_of_le +#align tsub_nonpos_of_le tsub_nonpos_of_le end Preorder diff --git a/Mathlib/Algebra/Order/WithZero.lean b/Mathlib/Algebra/Order/WithZero.lean index 4541f18c2046d..1080e85f2dfdc 100644 --- a/Mathlib/Algebra/Order/WithZero.lean +++ b/Mathlib/Algebra/Order/WithZero.lean @@ -252,6 +252,8 @@ Note that `OrderIso.mulLeft₀` refers to the `LinearOrderedField` version. -/ def OrderIso.mulLeft₀' {a : α} (ha : a ≠ 0) : α ≃o α := { Equiv.mulLeft₀ a ha with map_rel_iff' := mul_le_mul_left₀ ha } #align order_iso.mul_left₀' OrderIso.mulLeft₀' +#align order_iso.mul_left₀'_to_equiv OrderIso.mulLeft₀'_toEquiv +#align order_iso.mul_left₀'_apply OrderIso.mulLeft₀'_apply theorem OrderIso.mulLeft₀'_symm {a : α} (ha : a ≠ 0) : (OrderIso.mulLeft₀' ha).symm = OrderIso.mulLeft₀' (inv_ne_zero ha) := by @@ -266,6 +268,8 @@ Note that `OrderIso.mulRight₀` refers to the `LinearOrderedField` version. -/ def OrderIso.mulRight₀' {a : α} (ha : a ≠ 0) : α ≃o α := { Equiv.mulRight₀ a ha with map_rel_iff' := mul_le_mul_right₀ ha } #align order_iso.mul_right₀' OrderIso.mulRight₀' +#align order_iso.mul_right₀'_apply OrderIso.mulRight₀'_apply +#align order_iso.mul_right₀'_to_equiv OrderIso.mulRight₀'_toEquiv theorem OrderIso.mulRight₀'_symm {a : α} (ha : a ≠ 0) : (OrderIso.mulRight₀' ha).symm = OrderIso.mulRight₀' (inv_ne_zero ha) := by diff --git a/Mathlib/Algebra/Order/ZeroLEOne.lean b/Mathlib/Algebra/Order/ZeroLEOne.lean index db4f7f7d9f8eb..91ac18830b899 100644 --- a/Mathlib/Algebra/Order/ZeroLEOne.lean +++ b/Mathlib/Algebra/Order/ZeroLEOne.lean @@ -28,22 +28,27 @@ class ZeroLEOneClass (α : Type _) [Zero α] [One α] [LE α] where /-- `zero_le_one` with the type argument implicit. -/ @[simp] lemma zero_le_one [Zero α] [One α] [LE α] [ZeroLEOneClass α] : (0 : α) ≤ 1 := ZeroLEOneClass.zero_le_one +#align zero_le_one zero_le_one /-- `zero_le_one` with the type argument explicit. -/ lemma zero_le_one' (α) [Zero α] [One α] [LE α] [ZeroLEOneClass α] : (0 : α) ≤ 1 := zero_le_one +#align zero_le_one' zero_le_one' section variable [Zero α] [One α] [PartialOrder α] [ZeroLEOneClass α] [NeZero (1 : α)] /-- See `zero_lt_one'` for a version with the type explicit. -/ @[simp] lemma zero_lt_one : (0 : α) < 1 := zero_le_one.lt_of_ne (NeZero.ne' 1) +#align zero_lt_one zero_lt_one variable (α) /-- See `zero_lt_one` for a version with the type implicit. -/ lemma zero_lt_one' : (0 : α) < 1 := zero_lt_one +#align zero_lt_one' zero_lt_one' end alias zero_lt_one ← one_pos +#align one_pos one_pos diff --git a/Mathlib/Algebra/Parity.lean b/Mathlib/Algebra/Parity.lean index 1984176da8ea5..2899a6fa76861 100644 --- a/Mathlib/Algebra/Parity.lean +++ b/Mathlib/Algebra/Parity.lean @@ -93,11 +93,14 @@ theorem isSquare_iff_exists_sq (m : α) : IsSquare m ↔ ∃ c, m = c ^ 2 := by #align even_iff_exists_two_nsmul even_iff_exists_two_nsmul alias isSquare_iff_exists_sq ↔ IsSquare.exists_sq isSquare_of_exists_sq +#align is_square.exists_sq IsSquare.exists_sq +#align is_square_of_exists_sq isSquare_of_exists_sq attribute [to_additive Even.exists_two_nsmul "Alias of the forwards direction of `even_iff_exists_two_nsmul`."] IsSquare.exists_sq +#align even.exists_two_nsmul Even.exists_two_nsmul @[to_additive Even.nsmul] theorem IsSquare.pow (n : ℕ) : IsSquare a → IsSquare (a ^ n) := by @@ -169,8 +172,10 @@ theorem isSquare_inv : IsSquare a⁻¹ ↔ IsSquare a := by #align even_neg even_neg alias isSquare_inv ↔ _ IsSquare.inv +#align is_square.inv IsSquare.inv attribute [to_additive] IsSquare.inv +#align even.neg Even.neg @[to_additive] theorem IsSquare.zpow (n : ℤ) : IsSquare a → IsSquare (a ^ n) := by @@ -232,6 +237,7 @@ theorem even_iff_exists_bit0 [Add α] {a : α} : Even a ↔ ∃ b, a = bit0 b := #align even_iff_exists_bit0 even_iff_exists_bit0 alias even_iff_exists_bit0 ↔ Even.exists_bit0 _ +#align even.exists_bit0 Even.exists_bit0 section Semiring @@ -307,6 +313,7 @@ theorem odd_iff_exists_bit1 {a : α} : Odd a ↔ ∃ b, a = bit1 b := #align odd_iff_exists_bit1 odd_iff_exists_bit1 alias odd_iff_exists_bit1 ↔ Odd.exists_bit1 _ +#align odd.exists_bit1 Odd.exists_bit1 set_option linter.deprecated false in @[simp] theorem odd_bit1 (a : α) : Odd (bit1 a) := diff --git a/Mathlib/Algebra/Quandle.lean b/Mathlib/Algebra/Quandle.lean index aec56514c3aa8..b88a25d58bd72 100644 --- a/Mathlib/Algebra/Quandle.lean +++ b/Mathlib/Algebra/Quandle.lean @@ -108,6 +108,8 @@ structure ShelfHom (S₁ : Type _) (S₂ : Type _) [Shelf S₁] [Shelf S₂] whe /-- The homomorphism property of a Shelf Homomorphism-/ map_act' : ∀ {x y : S₁}, toFun (Shelf.act x y) = Shelf.act (toFun x) (toFun y) #align shelf_hom ShelfHom +#align shelf_hom.ext_iff ShelfHom.ext_iff +#align shelf_hom.ext ShelfHom.ext /-- A *rack* is an automorphic set (a set with an action on itself by bijections) that is self-distributive. It is a shelf such that each diff --git a/Mathlib/Algebra/Ring/AddAut.lean b/Mathlib/Algebra/Ring/AddAut.lean index 472b7393f496a..943fa5a3a5e91 100644 --- a/Mathlib/Algebra/Ring/AddAut.lean +++ b/Mathlib/Algebra/Ring/AddAut.lean @@ -31,6 +31,8 @@ variable {R : Type _} [Semiring R] def mulLeft : Rˣ →* AddAut R := DistribMulAction.toAddAut _ _ #align add_aut.mul_left AddAut.mulLeft +#align add_aut.mul_left_apply_apply AddAut.mulLeft_apply_apply +#align add_aut.mul_left_apply_symm_apply AddAut.mulLeft_apply_symmApply /-- Right multiplication by a unit of a semiring as an additive automorphism. -/ def mulRight (u : Rˣ) : AddAut R := diff --git a/Mathlib/Algebra/Ring/Aut.lean b/Mathlib/Algebra/Ring/Aut.lean index 1644fc4b7cef8..b201babcce4ae 100644 --- a/Mathlib/Algebra/Ring/Aut.lean +++ b/Mathlib/Algebra/Ring/Aut.lean @@ -147,6 +147,7 @@ def _root_.MulSemiringAction.toRingAut [MulSemiringAction G R] : map_mul' g h := RingEquiv.ext <| mul_smul g h map_one' := RingEquiv.ext <| one_smul _ #align mul_semiring_action.to_ring_aut MulSemiringAction.toRingAut +#align mul_semiring_action.to_ring_aut_apply MulSemiringAction.toRingAut_apply end Semiring diff --git a/Mathlib/Algebra/Ring/Basic.lean b/Mathlib/Algebra/Ring/Basic.lean index 90f102b97f742..774657722dc29 100644 --- a/Mathlib/Algebra/Ring/Basic.lean +++ b/Mathlib/Algebra/Ring/Basic.lean @@ -33,12 +33,14 @@ namespace AddHom def mulLeft [Distrib R] (r : R) : AddHom R R := ⟨(· * ·) r, mul_add r⟩ #align add_hom.mul_left AddHom.mulLeft +#align add_hom.mul_left_apply AddHom.mulLeft_apply /-- Left multiplication by an element of a type with distributive multiplication is an `AddHom`. -/ @[simps (config := { fullyApplied := false })] def mulRight [Distrib R] (r : R) : AddHom R R := ⟨fun a => a * r, fun _ _ => add_mul _ _ r⟩ #align add_hom.mul_right AddHom.mulRight +#align add_hom.mul_right_apply AddHom.mulRight_apply end AddHom diff --git a/Mathlib/Algebra/Ring/Defs.lean b/Mathlib/Algebra/Ring/Defs.lean index 8ebf12b2204c5..1aa44aa6bc95d 100644 --- a/Mathlib/Algebra/Ring/Defs.lean +++ b/Mathlib/Algebra/Ring/Defs.lean @@ -85,6 +85,7 @@ theorem left_distrib [Mul R] [Add R] [LeftDistribClass R] (a b c : R) : #align left_distrib left_distrib alias left_distrib ← mul_add +#align mul_add mul_add theorem right_distrib [Mul R] [Add R] [RightDistribClass R] (a b c : R) : (a + b) * c = a * c + b * c := @@ -92,6 +93,7 @@ theorem right_distrib [Mul R] [Add R] [RightDistribClass R] (a b c : R) : #align right_distrib right_distrib alias right_distrib ← add_mul +#align add_mul add_mul theorem distrib_three_right [Mul R] [Add R] [RightDistribClass R] (a b c d : R) : (a + b + c) * d = a * d + b * d + c * d := by simp [right_distrib] @@ -287,14 +289,17 @@ section MulOneClass variable [MulOneClass α] [HasDistribNeg α] theorem neg_eq_neg_one_mul (a : α) : -a = -1 * a := by simp +#align neg_eq_neg_one_mul neg_eq_neg_one_mul /-- An element of a ring multiplied by the additive inverse of one is the element's additive inverse. -/ theorem mul_neg_one (a : α) : a * -1 = -a := by simp +#align mul_neg_one mul_neg_one /-- The additive inverse of one multiplied by an element of a ring is the element's additive inverse. -/ theorem neg_one_mul (a : α) : -1 * a = -a := by simp +#align neg_one_mul neg_one_mul end MulOneClass @@ -350,12 +355,14 @@ theorem mul_sub_left_distrib (a b c : α) : a * (b - c) = a * b - a * c := by #align mul_sub_left_distrib mul_sub_left_distrib alias mul_sub_left_distrib ← mul_sub +#align mul_sub mul_sub theorem mul_sub_right_distrib (a b c : α) : (a - b) * c = a * c - b * c := by simpa only [sub_eq_add_neg, neg_mul_eq_neg_mul] using add_mul a (-b) c #align mul_sub_right_distrib mul_sub_right_distrib alias mul_sub_right_distrib ← sub_mul +#align sub_mul sub_mul variable {a b c d e : α} @@ -392,12 +399,16 @@ section NonAssocRing variable [NonAssocRing α] theorem sub_one_mul (a b : α) : (a - 1) * b = a * b - b := by rw [sub_mul, one_mul] +#align sub_one_mul sub_one_mul theorem mul_sub_one (a b : α) : a * (b - 1) = a * b - a := by rw [mul_sub, mul_one] +#align mul_sub_one mul_sub_one theorem one_sub_mul (a b : α) : (1 - a) * b = b - a * b := by rw [sub_mul, one_mul] +#align one_sub_mul one_sub_mul theorem mul_one_sub (a b : α) : a * (1 - b) = a - a * b := by rw [mul_sub, mul_one] +#align mul_one_sub mul_one_sub end NonAssocRing diff --git a/Mathlib/Algebra/Ring/Equiv.lean b/Mathlib/Algebra/Ring/Equiv.lean index 95aad9a5d06f5..3c675c6d03823 100644 --- a/Mathlib/Algebra/Ring/Equiv.lean +++ b/Mathlib/Algebra/Ring/Equiv.lean @@ -54,12 +54,15 @@ infixl:25 " ≃+* " => RingEquiv /-- The "plain" equivalence of types underlying an equivalence of (semi)rings. -/ add_decl_doc RingEquiv.toEquiv +#align ring_equiv.to_equiv RingEquiv.toEquiv /-- The equivalence of additive monoids underlying an equivalence of (semi)rings. -/ add_decl_doc RingEquiv.toAddEquiv +#align ring_equiv.to_add_equiv RingEquiv.toAddEquiv /-- The equivalence of multiplicative monoids underlying an equivalence of (semi)rings. -/ add_decl_doc RingEquiv.toMulEquiv +#align ring_equiv.to_mul_equiv RingEquiv.toMulEquiv /-- `RingEquivClass F R S` states that `F` is a type of ring structure preserving equivalences. You should extend this class when you extend `RingEquiv`. -/ @@ -364,6 +367,8 @@ protected def op {α β} [Add α] [Mul α] [Add β] [Mul β] : ext rfl #align ring_equiv.op RingEquiv.op +#align ring_equiv.op_symm_apply_apply RingEquiv.op_symm_apply_apply +#align ring_equiv.op_symm_apply_symm_apply RingEquiv.op_symm_apply_symmApply /-- The 'unopposite' of a ring iso `αᵐᵒᵖ ≃+* βᵐᵒᵖ`. Inverse to `RingEquiv.op`. -/ @[simp] @@ -448,6 +453,7 @@ def piCongrRight {ι : Type _} {R S : ι → Type _} [∀ i, NonUnitalNonAssocSe toFun := fun x j => e j (x j) invFun := fun x j => (e j).symm (x j) } #align ring_equiv.Pi_congr_right RingEquiv.piCongrRight +#align ring_equiv.Pi_congr_right_apply RingEquiv.piCongrRight_apply @[simp] theorem piCongrRight_refl {ι : Type _} {R : ι → Type _} [∀ i, NonUnitalNonAssocSemiring (R i)] : @@ -768,6 +774,8 @@ def ofHomInv' {R S F G : Type _} [NonUnitalNonAssocSemiring R] [NonUnitalNonAsso map_mul' := map_mul hom map_add' := map_add hom #align ring_equiv.of_hom_inv' RingEquiv.ofHomInv' +#align ring_equiv.of_hom_inv'_symm_apply RingEquiv.ofHomInv'_symmApply +#align ring_equiv.of_hom_inv'_apply RingEquiv.ofHomInv'_apply /-- Construct an equivalence of rings from unital homomorphisms in both directions, which are inverses. @@ -785,6 +793,8 @@ def ofHomInv {R S F G : Type _} [NonAssocSemiring R] [NonAssocSemiring S] [RingH map_mul' := map_mul hom map_add' := map_add hom #align ring_equiv.of_hom_inv RingEquiv.ofHomInv +#align ring_equiv.of_hom_inv_apply RingEquiv.ofHomInv_apply +#align ring_equiv.of_hom_inv_symm_apply RingEquiv.ofHomInv_symmApply end SemiringHom diff --git a/Mathlib/Algebra/Ring/Fin.lean b/Mathlib/Algebra/Ring/Fin.lean index 0b61558a8d139..ee54e0c74756c 100644 --- a/Mathlib/Algebra/Ring/Fin.lean +++ b/Mathlib/Algebra/Ring/Fin.lean @@ -33,3 +33,5 @@ def RingEquiv.piFinTwo (R : Fin 2 → Type _) [∀ i, Semiring (R i)] : map_add' := fun _ _ => rfl map_mul' := fun _ _ => rfl } #align ring_equiv.pi_fin_two RingEquiv.piFinTwo +#align ring_equiv.pi_fin_two_apply RingEquiv.piFinTwo_apply +#align ring_equiv.pi_fin_two_symm_apply RingEquiv.piFinTwo_symmApply diff --git a/Mathlib/Algebra/Ring/Opposite.lean b/Mathlib/Algebra/Ring/Opposite.lean index 92bc76fe93ef5..ab581954e57d3 100644 --- a/Mathlib/Algebra/Ring/Opposite.lean +++ b/Mathlib/Algebra/Ring/Opposite.lean @@ -236,6 +236,7 @@ def RingHom.toOpposite {R S : Type _} [Semiring R] [Semiring S] (f : R →+* S) { ((opAddEquiv : S ≃+ Sᵐᵒᵖ).toAddMonoidHom.comp ↑f : R →+ Sᵐᵒᵖ), f.toMonoidHom.toOpposite hf with toFun := MulOpposite.op ∘ f } #align ring_hom.to_opposite RingHom.toOpposite +#align ring_hom.to_opposite_apply RingHom.toOpposite_apply /-- A ring homomorphism `f : R →+* S` such that `f x` commutes with `f y` for all `x, y` defines a ring homomorphism from `Rᵐᵒᵖ`. -/ @@ -245,6 +246,7 @@ def RingHom.fromOpposite {R S : Type _} [Semiring R] [Semiring S] (f : R →+* S { (f.toAddMonoidHom.comp (opAddEquiv : R ≃+ Rᵐᵒᵖ).symm.toAddMonoidHom : Rᵐᵒᵖ →+ S), f.toMonoidHom.fromOpposite hf with toFun := f ∘ MulOpposite.unop } #align ring_hom.from_opposite RingHom.fromOpposite +#align ring_hom.from_opposite_apply RingHom.fromOpposite_apply /-- A ring hom `α →+* β` can equivalently be viewed as a ring hom `αᵐᵒᵖ →+* βᵐᵒᵖ`. This is the action of the (fully faithful) `ᵐᵒᵖ`-functor on morphisms. -/ @@ -256,6 +258,7 @@ def RingHom.op {α β} [NonAssocSemiring α] [NonAssocSemiring β] : left_inv _ := rfl right_inv _ := rfl #align ring_hom.op RingHom.op +#align ring_hom.op_symm_apply_apply RingHom.op_symm_apply_apply /-- The 'unopposite' of a ring hom `αᵐᵒᵖ →+* βᵐᵒᵖ`. Inverse to `RingHom.op`. -/ @[simp] diff --git a/Mathlib/Algebra/Ring/Pi.lean b/Mathlib/Algebra/Ring/Pi.lean index 11df529eca610..b835bef9759ff 100644 --- a/Mathlib/Algebra/Ring/Pi.lean +++ b/Mathlib/Algebra/Ring/Pi.lean @@ -115,6 +115,7 @@ protected def ringHom {γ : Type w} [∀ i, NonAssocSemiring (f i)] [NonAssocSem { Pi.monoidHom fun i => (g i).toMonoidHom, Pi.addMonoidHom fun i => (g i).toAddMonoidHom with toFun := fun x b => g b x } #align pi.ring_hom Pi.ringHom +#align pi.ring_hom_apply Pi.ringHom_apply theorem ringHom_injective {γ : Type w} [Nonempty I] [∀ i, NonAssocSemiring (f i)] [NonAssocSemiring γ] (g : ∀ i, γ →+* f i) (hg : ∀ i, Function.Injective (g i)) : @@ -166,12 +167,14 @@ homomorphism. This is `Function.eval` as a `RingHom`. -/ def Pi.evalRingHom (f : I → Type v) [∀ i, NonAssocSemiring (f i)] (i : I) : (∀ i, f i) →+* f i := { Pi.evalMonoidHom f i, Pi.evalAddMonoidHom f i with } #align pi.eval_ring_hom Pi.evalRingHom +#align pi.eval_ring_hom_apply Pi.evalRingHom_apply /-- `Function.const` as a `RingHom`. -/ @[simps] def Pi.constRingHom (α β : Type _) [NonAssocSemiring β] : β →+* α → β := { Pi.ringHom fun _ => RingHom.id β with toFun := Function.const _ } #align pi.const_ring_hom Pi.constRingHom +#align pi.const_ring_hom_apply Pi.constRingHom_apply /-- Ring homomorphism between the function spaces `I → α` and `I → β`, induced by a ring homomorphism `f` between `α` and `β`. -/ @@ -180,5 +183,6 @@ protected def RingHom.compLeft {α β : Type _} [NonAssocSemiring α] [NonAssocS (f : α →+* β) (I : Type _) : (I → α) →+* I → β := { f.toMonoidHom.compLeft I, f.toAddMonoidHom.compLeft I with toFun := fun h => f ∘ h } #align ring_hom.comp_left RingHom.compLeft +#align ring_hom.comp_left_apply RingHom.compLeft_apply end RingHom diff --git a/Mathlib/Algebra/Ring/Prod.lean b/Mathlib/Algebra/Ring/Prod.lean index 6e7040d37ad05..4937e855c1969 100644 --- a/Mathlib/Algebra/Ring/Prod.lean +++ b/Mathlib/Algebra/Ring/Prod.lean @@ -315,6 +315,8 @@ def prodZeroRing : R ≃+* R × S where left_inv x := rfl right_inv x := by cases x; simp #align ring_equiv.prod_zero_ring RingEquiv.prodZeroRing +#align ring_equiv.prod_zero_ring_symm_apply RingEquiv.prodZeroRing_symmApply +#align ring_equiv.prod_zero_ring_apply RingEquiv.prodZeroRing_apply /-- A ring `R` is isomorphic to `S × R` when `S` is the zero ring -/ @[simps] @@ -326,6 +328,8 @@ def zeroRingProd : R ≃+* S × R where left_inv x := rfl right_inv x := by cases x; simp #align ring_equiv.zero_ring_prod RingEquiv.zeroRingProd +#align ring_equiv.zero_ring_prod_symm_apply RingEquiv.zeroRingProd_symmApply +#align ring_equiv.zero_ring_prod_apply RingEquiv.zeroRingProd_apply end RingEquiv diff --git a/Mathlib/Algebra/SMulWithZero.lean b/Mathlib/Algebra/SMulWithZero.lean index f6915a310bf79..a8c191bbc72be 100644 --- a/Mathlib/Algebra/SMulWithZero.lean +++ b/Mathlib/Algebra/SMulWithZero.lean @@ -232,3 +232,4 @@ def smulMonoidWithZeroHom {α β : Type _} [MonoidWithZero α] [MulZeroOneClass [MulActionWithZero α β] [IsScalarTower α β β] [SMulCommClass α β β] : α × β →*₀ β := { smulMonoidHom with map_zero' := smul_zero _ } #align smul_monoid_with_zero_hom smulMonoidWithZeroHom +#align smul_monoid_with_zero_hom_apply smulMonoidWithZeroHom_apply diff --git a/Mathlib/Algebra/Star/Basic.lean b/Mathlib/Algebra/Star/Basic.lean index a09c35fffd169..7b5f84b3a85cb 100644 --- a/Mathlib/Algebra/Star/Basic.lean +++ b/Mathlib/Algebra/Star/Basic.lean @@ -155,6 +155,7 @@ def starMulEquiv [Semigroup R] [StarSemigroup R] : R ≃* Rᵐᵒᵖ := toFun := fun x => MulOpposite.op (star x) map_mul' := fun x y => by simp only [star_mul, op_mul] } #align star_mul_equiv starMulEquiv +#align star_mul_equiv_apply starMulEquiv_apply /-- `star` as a `MulAut` for commutative `R`. -/ @[simps apply] @@ -165,6 +166,7 @@ def starMulAut [CommSemigroup R] [StarSemigroup R] : MulAut R := toFun := star map_mul' := star_mul' } #align star_mul_aut starMulAut +#align star_mul_aut_apply starMulAut_apply variable (R) @@ -240,6 +242,7 @@ def starAddEquiv [AddMonoid R] [StarAddMonoid R] : R ≃+ R := toFun := star map_add' := star_add } #align star_add_equiv starAddEquiv +#align star_add_equiv_apply starAddEquiv_apply variable (R) @@ -297,6 +300,7 @@ def starRingEquiv [NonUnitalSemiring R] [StarRing R] : R ≃+* Rᵐᵒᵖ := { starAddEquiv.trans (MulOpposite.opAddEquiv : R ≃+ Rᵐᵒᵖ), starMulEquiv with toFun := fun x => MulOpposite.op (star x) } #align star_ring_equiv starRingEquiv +#align star_ring_equiv_apply starRingEquiv_apply @[simp, norm_cast] theorem star_natCast [Semiring R] [StarRing R] (n : ℕ) : star (n : R) = n := @@ -324,6 +328,7 @@ end def starRingAut [CommSemiring R] [StarRing R] : RingAut R := { starAddEquiv, starMulAut with toFun := star } #align star_ring_aut starRingAut +#align star_ring_aut_apply starRingAut_apply variable (R) @@ -386,8 +391,11 @@ theorem RingHom.star_apply {S : Type _} [NonAssocSemiring S] [CommSemiring R] [S -- A more convenient name for complex conjugation alias starRingEnd_self_apply ← Complex.conj_conj +#align complex.conj_conj Complex.conj_conj alias starRingEnd_self_apply ← IsROrC.conj_conj +set_option linter.uppercaseLean3 false in +#align is_R_or_C.conj_conj IsROrC.conj_conj @[simp] theorem star_inv' [DivisionRing R] [StarRing R] (x : R) : star x⁻¹ = (star x)⁻¹ := diff --git a/Mathlib/CategoryTheory/Category/Basic.lean b/Mathlib/CategoryTheory/Category/Basic.lean index 0d5f2df45a48d..5d4989cb28398 100644 --- a/Mathlib/CategoryTheory/Category/Basic.lean +++ b/Mathlib/CategoryTheory/Category/Basic.lean @@ -144,6 +144,9 @@ class Category (obj : Type u) extends CategoryStruct.{v} obj : Type max u (v + 1 assoc : ∀ {W X Y Z : obj} (f : W ⟶ X) (g : X ⟶ Y) (h : Y ⟶ Z), (f ≫ g) ≫ h = f ≫ g ≫ h := by aesop_cat #align category_theory.category CategoryTheory.Category +#align category_theory.category.assoc CategoryTheory.Category.assoc +#align category_theory.category.comp_id CategoryTheory.Category.comp_id +#align category_theory.category.id_comp CategoryTheory.Category.id_comp -- Porting note: `restate_axiom` should not be necessary in lean4 -- Hopefully we can just remove the backticks from field names, diff --git a/Mathlib/CategoryTheory/ConcreteCategory/Bundled.lean b/Mathlib/CategoryTheory/ConcreteCategory/Bundled.lean index 37a8316c6b443..a3b512261544a 100644 --- a/Mathlib/CategoryTheory/ConcreteCategory/Bundled.lean +++ b/Mathlib/CategoryTheory/ConcreteCategory/Bundled.lean @@ -10,6 +10,7 @@ Authors: Scott Morrison, Johannes Hölzl, Reid Barton, Sean Leather -/ import Std.Tactic.Lint.Frontend import Std.Tactic.Lint.Misc +import Mathlib.Mathport.Rename /-! # Bundled types @@ -34,6 +35,7 @@ structure Bundled (c : Type u → Type v) : Type max (u + 1) v where α : Type u /-- The corresponding instance of the bundled type class -/ str : c α := by infer_instance +#align category_theory.bundled CategoryTheory.Bundled namespace Bundled @@ -45,12 +47,14 @@ set_option checkBinderAnnotations false in /-- A generic function for lifting a type equipped with an instance to a bundled object. -/ def of {c : Type u → Type v} (α : Type u) [str : c α] : Bundled c := ⟨α, str⟩ +#align category_theory.bundled.of CategoryTheory.Bundled.of instance : CoeSort (Bundled c) (Type u) := ⟨Bundled.α⟩ theorem coe_mk (α) (str) : (@Bundled.mk c α str : Type u) = α := rfl +#align category_theory.bundled.coe_mk CategoryTheory.Bundled.coe_mk /- `Bundled.map` is reducible so that, if we define a category @@ -67,6 +71,7 @@ Lean 4. /-- Map over the bundled structure -/ def map (f : ∀ {α}, c α → d α) (b : Bundled c) : Bundled d := ⟨b, f b.str⟩ +#align category_theory.bundled.map CategoryTheory.Bundled.map end Bundled diff --git a/Mathlib/CategoryTheory/EssentialImage.lean b/Mathlib/CategoryTheory/EssentialImage.lean index 2d3eb0f3945d0..5ecd54b8df92d 100644 --- a/Mathlib/CategoryTheory/EssentialImage.lean +++ b/Mathlib/CategoryTheory/EssentialImage.lean @@ -94,6 +94,8 @@ instance : Category (EssImageSubcategory F) := def essImageInclusion (F : C ⥤ D) : F.EssImageSubcategory ⥤ D := fullSubcategoryInclusion _ #align category_theory.functor.ess_image_inclusion CategoryTheory.Functor.essImageInclusion +#align category_theory.functor.ess_image_inclusion_obj CategoryTheory.Functor.essImageInclusion_obj +#align category_theory.functor.ess_image_inclusion_map CategoryTheory.Functor.essImageInclusion_map -- Porting note: `deriving Full` is not able to derive this instance instance : Full (essImageInclusion F) := @@ -111,6 +113,8 @@ image of `F`. def toEssImage (F : C ⥤ D) : C ⥤ F.EssImageSubcategory := FullSubcategory.lift _ F (obj_mem_essImage _) #align category_theory.functor.to_ess_image CategoryTheory.Functor.toEssImage +#align category_theory.functor.to_ess_image_map CategoryTheory.Functor.toEssImage_map +#align category_theory.functor.to_ess_image_obj_obj CategoryTheory.Functor.toEssImage_obj_obj /-- The functor `F` factorises through its essential image, where the first functor is essentially surjective and the second is fully faithful. @@ -119,6 +123,8 @@ surjective and the second is fully faithful. def toEssImageCompEssentialImageInclusion (F : C ⥤ D) : F.toEssImage ⋙ F.essImageInclusion ≅ F := FullSubcategory.lift_comp_inclusion _ _ _ #align category_theory.functor.to_ess_image_comp_essential_image_inclusion CategoryTheory.Functor.toEssImageCompEssentialImageInclusion +#align category_theory.functor.to_ess_image_comp_essential_image_inclusion_hom_app CategoryTheory.Functor.toEssImageCompEssentialImageInclusion_hom_app +#align category_theory.functor.to_ess_image_comp_essential_image_inclusion_inv_app CategoryTheory.Functor.toEssImageCompEssentialImageInclusion_inv_app end Functor diff --git a/Mathlib/CategoryTheory/FullSubcategory.lean b/Mathlib/CategoryTheory/FullSubcategory.lean index fc98cf1082b8f..07250e7f744ee 100644 --- a/Mathlib/CategoryTheory/FullSubcategory.lean +++ b/Mathlib/CategoryTheory/FullSubcategory.lean @@ -80,6 +80,8 @@ def inducedFunctor : InducedCategory D F ⥤ D where obj := F map f := f #align category_theory.induced_functor CategoryTheory.inducedFunctor +#align category_theory.induced_functor_map CategoryTheory.inducedFunctor_map +#align category_theory.induced_functor_obj CategoryTheory.inducedFunctor_obj instance InducedCategory.full : Full (inducedFunctor F) where preimage f := f #align category_theory.induced_category.full CategoryTheory.InducedCategory.full @@ -108,6 +110,8 @@ structure FullSubcategory where /--The predicate satisfied by all objects in this subcategory-/ property : Z obj #align category_theory.full_subcategory CategoryTheory.FullSubcategory +#align category_theory.full_subcategory.ext CategoryTheory.FullSubcategory.ext +#align category_theory.full_subcategory.ext_iff CategoryTheory.FullSubcategory.ext_iff instance FullSubcategory.category : Category.{v} (FullSubcategory Z) := InducedCategory.category FullSubcategory.obj @@ -146,6 +150,8 @@ def FullSubcategory.map (h : ∀ ⦃X⦄, Z X → Z' X) : FullSubcategory Z ⥤ obj X := ⟨X.1, h X.2⟩ map f := f #align category_theory.full_subcategory.map CategoryTheory.FullSubcategory.map +#align category_theory.full_subcategory.map_obj_obj CategoryTheory.FullSubcategory.map_obj_obj +#align category_theory.full_subcategory.map_map CategoryTheory.FullSubcategory.map_map instance (h : ∀ ⦃X⦄, Z X → Z' X) : Full (FullSubcategory.map h) where preimage f := f @@ -168,6 +174,8 @@ def FullSubcategory.lift (F : C ⥤ D) (hF : ∀ X, P (F.obj X)) : C ⥤ FullSub obj X := ⟨F.obj X, hF X⟩ map f := F.map f #align category_theory.full_subcategory.lift CategoryTheory.FullSubcategory.lift +#align category_theory.full_subcategory.lift_obj_obj CategoryTheory.FullSubcategory.lift_obj_obj +#align category_theory.full_subcategory.lift_map CategoryTheory.FullSubcategory.lift_map /-- Composing the lift of a functor through a full subcategory with the inclusion yields the original functor. Unfortunately, this is not true by definition, so we only get a natural diff --git a/Mathlib/CategoryTheory/Functor/Basic.lean b/Mathlib/CategoryTheory/Functor/Basic.lean index 42fbd9cf7ea31..fa994fdda6e07 100644 --- a/Mathlib/CategoryTheory/Functor/Basic.lean +++ b/Mathlib/CategoryTheory/Functor/Basic.lean @@ -46,9 +46,12 @@ structure Functor (C : Type u₁) [Category.{v₁} C] (D : Type u₂) [Category. /-- A functor preserves composition. -/ map_comp : ∀ {X Y Z : C} (f : X ⟶ Y) (g : Y ⟶ Z), map (f ≫ g) = map f ≫ map g := by aesop_cat #align category_theory.functor CategoryTheory.Functor +#align category_theory.functor.map_comp CategoryTheory.Functor.map_comp +#align category_theory.functor.map_id CategoryTheory.Functor.map_id /-- The prefunctor between the underlying quivers. -/ add_decl_doc Functor.toPrefunctor +#align category_theory.functor.to_prefunctor CategoryTheory.Functor.toPrefunctor end @@ -61,6 +64,7 @@ attribute [simp] Functor.map_id -- We intentionally don't add `simp` to the `reassoc` lemma, -- which is only useful for rewriting backwards. attribute [reassoc, simp] Functor.map_comp +#align category_theory.functor.map_comp_assoc CategoryTheory.Functor.map_comp_assoc namespace Functor @@ -108,6 +112,7 @@ def comp (F : C ⥤ D) (G : D ⥤ E) : C ⥤ E where obj X := G.obj (F.obj X) map f := G.map (F.map f) #align category_theory.functor.comp CategoryTheory.Functor.comp +#align category_theory.functor.comp_obj CategoryTheory.Functor.comp_obj /-- Notation for composition of functors. -/ infixr:80 " ⋙ " => comp diff --git a/Mathlib/CategoryTheory/Functor/Category.lean b/Mathlib/CategoryTheory/Functor/Category.lean index 88f7029b64e88..70904ae51e0f3 100644 --- a/Mathlib/CategoryTheory/Functor/Category.lean +++ b/Mathlib/CategoryTheory/Functor/Category.lean @@ -112,6 +112,7 @@ def hcomp {H I : D ⥤ E} (α : F ⟶ G) (β : H ⟶ I) : F ⋙ H ⟶ G ⋙ I wh rw [Functor.comp_map, Functor.comp_map, ← assoc, naturality, assoc, ← map_comp I, naturality, map_comp, assoc] #align category_theory.nat_trans.hcomp CategoryTheory.NatTrans.hcomp +#align category_theory.nat_trans.hcomp_app CategoryTheory.NatTrans.hcomp_app /-- Notation for horizontal composition of natural transformations. -/ infixl:80 " ◫ " => hcomp @@ -146,6 +147,9 @@ protected def flip (F : C ⥤ D ⥤ E) : D ⥤ C ⥤ E where map := fun f => (F.map f).app k, } map f := { app := fun j => (F.obj j).map f } #align category_theory.functor.flip CategoryTheory.Functor.flip +#align category_theory.functor.flip_obj_map CategoryTheory.Functor.flip_obj_map +#align category_theory.functor.flip_obj_obj CategoryTheory.Functor.flip_obj_obj +#align category_theory.functor.flip_map_app CategoryTheory.Functor.flip_map_app end Functor @@ -154,11 +158,13 @@ theorem map_hom_inv_app (F : C ⥤ D ⥤ E) {X Y : C} (e : X ≅ Y) (Z : D) : (F.map e.hom).app Z ≫ (F.map e.inv).app Z = 𝟙 _ := by simp [← NatTrans.comp_app, ← Functor.map_comp] #align category_theory.map_hom_inv_app CategoryTheory.map_hom_inv_app +#align category_theory.map_hom_inv_app_assoc CategoryTheory.map_hom_inv_app_assoc @[reassoc (attr := simp)] theorem map_inv_hom_app (F : C ⥤ D ⥤ E) {X Y : C} (e : X ≅ Y) (Z : D) : (F.map e.inv).app Z ≫ (F.map e.hom).app Z = 𝟙 _ := by simp [← NatTrans.comp_app, ← Functor.map_comp] #align category_theory.map_inv_hom_app CategoryTheory.map_inv_hom_app +#align category_theory.map_inv_hom_app_assoc CategoryTheory.map_inv_hom_app_assoc end CategoryTheory diff --git a/Mathlib/CategoryTheory/Functor/FullyFaithful.lean b/Mathlib/CategoryTheory/Functor/FullyFaithful.lean index bf68548899d90..5e20a43f3e285 100644 --- a/Mathlib/CategoryTheory/Functor/FullyFaithful.lean +++ b/Mathlib/CategoryTheory/Functor/FullyFaithful.lean @@ -52,6 +52,7 @@ class Full (F : C ⥤ D) where /-- The property that `Full.preimage f` of maps to `f` via `F.map`. -/ witness : ∀ {X Y : C} (f : F.obj X ⟶ F.obj Y), F.map (preimage f) = f := by aesop_cat #align category_theory.full CategoryTheory.Full +#align category_theory.full.witness CategoryTheory.Full.witness attribute [simp] Full.witness @@ -67,6 +68,7 @@ class Faithful (F : C ⥤ D) : Prop where map_injective : ∀ {X Y : C}, Function.Injective (F.map : (X ⟶ Y) → (F.obj X ⟶ F.obj Y)) := by aesop_cat #align category_theory.faithful CategoryTheory.Faithful +#align category_theory.faithful.map_injective CategoryTheory.Faithful.map_injective namespace Functor @@ -145,6 +147,8 @@ def preimageIso (f : F.obj X ≅ F.obj Y) : hom_inv_id := F.map_injective (by simp) inv_hom_id := F.map_injective (by simp) #align category_theory.functor.preimage_iso CategoryTheory.Functor.preimageIso +#align category_theory.functor.preimage_iso_inv CategoryTheory.Functor.preimageIso_inv +#align category_theory.functor.preimage_iso_hom CategoryTheory.Functor.preimageIso_hom @[simp] theorem preimageIso_mapIso (f : X ≅ Y) : F.preimageIso (F.mapIso f) = f := by @@ -170,6 +174,8 @@ def equivOfFullyFaithful {X Y} : left_inv f := by simp right_inv f := by simp #align category_theory.equiv_of_fully_faithful CategoryTheory.equivOfFullyFaithful +#align category_theory.equiv_of_fully_faithful_apply CategoryTheory.equivOfFullyFaithful_apply +#align category_theory.equiv_of_fully_faithful_symm_apply CategoryTheory.equivOfFullyFaithful_symm_apply /-- If `F` is fully faithful, we have an equivalence of iso-sets `X ≅ Y` and `F X ≅ F Y`. -/ @[simps] @@ -182,6 +188,8 @@ def isoEquivOfFullyFaithful {X Y} : ext simp #align category_theory.iso_equiv_of_fully_faithful CategoryTheory.isoEquivOfFullyFaithful +#align category_theory.iso_equiv_of_fully_faithful_symm_apply CategoryTheory.isoEquivOfFullyFaithful_symm_apply +#align category_theory.iso_equiv_of_fully_faithful_apply CategoryTheory.isoEquivOfFullyFaithful_apply end @@ -200,6 +208,7 @@ def natTransOfCompFullyFaithful (α : F ⋙ H ⟶ G ⋙ H) : apply H.map_injective simpa using α.naturality f #align category_theory.nat_trans_of_comp_fully_faithful CategoryTheory.natTransOfCompFullyFaithful +#align category_theory.nat_trans_of_comp_fully_faithful_app CategoryTheory.natTransOfCompFullyFaithful_app /-- We can construct a natural isomorphism between functors by constructing a natural isomorphism between those functors composed with a fully faithful functor. -/ @@ -210,6 +219,8 @@ def natIsoOfCompFullyFaithful (i : F ⋙ H ≅ G ⋙ H) : F ≅ G := apply H.map_injective simpa using i.hom.naturality f #align category_theory.nat_iso_of_comp_fully_faithful CategoryTheory.natIsoOfCompFullyFaithful +#align category_theory.nat_iso_of_comp_fully_faithful_hom_app CategoryTheory.natIsoOfCompFullyFaithful_hom_app +#align category_theory.nat_iso_of_comp_fully_faithful_inv_app CategoryTheory.natIsoOfCompFullyFaithful_inv_app theorem natIsoOfCompFullyFaithful_hom (i : F ⋙ H ≅ G ⋙ H) : (natIsoOfCompFullyFaithful H i).hom = natTransOfCompFullyFaithful H i.hom := by @@ -233,6 +244,8 @@ def NatTrans.equivOfCompFullyFaithful : left_inv := by aesop_cat right_inv := by aesop_cat #align category_theory.nat_trans.equiv_of_comp_fully_faithful CategoryTheory.NatTrans.equivOfCompFullyFaithful +#align category_theory.nat_trans.equiv_of_comp_fully_faithful_apply CategoryTheory.NatTrans.equivOfCompFullyFaithful_apply +#align category_theory.nat_trans.equiv_of_comp_fully_faithful_symm_apply CategoryTheory.NatTrans.equivOfCompFullyFaithful_symm_apply /-- Horizontal composition with a fully faithful functor induces a bijection on natural isomorphisms. -/ @@ -244,6 +257,8 @@ def NatIso.equivOfCompFullyFaithful : left_inv := by aesop_cat right_inv := by aesop_cat #align category_theory.nat_iso.equiv_of_comp_fully_faithful CategoryTheory.NatIso.equivOfCompFullyFaithful +#align category_theory.nat_iso.equiv_of_comp_fully_faithful_symm_apply CategoryTheory.NatIso.equivOfCompFullyFaithful_symm_apply +#align category_theory.nat_iso.equiv_of_comp_fully_faithful_apply CategoryTheory.NatIso.equivOfCompFullyFaithful_apply end @@ -297,6 +312,7 @@ theorem Faithful.of_comp_iso {H : C ⥤ E} [Faithful H] (h : F ⋙ G ≅ H) : Fa #align category_theory.faithful.of_comp_iso CategoryTheory.Faithful.of_comp_iso alias Faithful.of_comp_iso ← _root_.CategoryTheory.Iso.faithful_of_comp +#align category_theory.iso.faithful_of_comp CategoryTheory.Iso.faithful_of_comp -- We could prove this from `Faithful.of_comp_iso` using `eq_to_iso`, -- but that would introduce a cyclic import. @@ -305,6 +321,7 @@ theorem Faithful.of_comp_eq {H : C ⥤ E} [ℋ : Faithful H] (h : F ⋙ G = H) : #align category_theory.faithful.of_comp_eq CategoryTheory.Faithful.of_comp_eq alias Faithful.of_comp_eq ← _root_.Eq.faithful_of_comp +#align eq.faithful_of_comp Eq.faithful_of_comp variable (F G) /-- “Divide” a functor by a faithful functor. -/ diff --git a/Mathlib/CategoryTheory/Iso.lean b/Mathlib/CategoryTheory/Iso.lean index ec5c82e59e4a7..bc920e91422c8 100644 --- a/Mathlib/CategoryTheory/Iso.lean +++ b/Mathlib/CategoryTheory/Iso.lean @@ -63,8 +63,12 @@ structure Iso {C : Type u} [Category.{v} C] (X Y : C) where is the identity on the target. -/ inv_hom_id : inv ≫ hom = 𝟙 Y := by aesop_cat #align category_theory.iso CategoryTheory.Iso +#align category_theory.iso.inv_hom_id CategoryTheory.Iso.inv_hom_id +#align category_theory.iso.hom_inv_id CategoryTheory.Iso.hom_inv_id attribute [reassoc (attr := simp)] Iso.hom_inv_id Iso.inv_hom_id +#align category_theory.iso.hom_inv_id_assoc CategoryTheory.Iso.hom_inv_id_assoc +#align category_theory.iso.inv_hom_id_assoc CategoryTheory.Iso.inv_hom_id_assoc /-- Notation for an isomorphism in a category. -/ infixr:10 " ≅ " => Iso -- type as \cong or \iso @@ -131,6 +135,8 @@ def refl (X : C) : X ≅ X where hom := 𝟙 X inv := 𝟙 X #align category_theory.iso.refl CategoryTheory.Iso.refl +#align category_theory.iso.refl_inv CategoryTheory.Iso.refl_inv +#align category_theory.iso.refl_hom CategoryTheory.Iso.refl_hom instance : Inhabited (X ≅ X) := ⟨Iso.refl X⟩ @@ -144,6 +150,8 @@ def trans (α : X ≅ Y) (β : Y ≅ Z) : X ≅ Z where hom := α.hom ≫ β.hom inv := β.inv ≫ α.inv #align category_theory.iso.trans CategoryTheory.Iso.trans +#align category_theory.iso.trans_hom CategoryTheory.Iso.trans_hom +#align category_theory.iso.trans_inv CategoryTheory.Iso.trans_inv /-- Notation for composition of isomorphisms. -/ infixr:80 " ≪≫ " => Iso.trans -- type as `\ll \gg`. @@ -276,10 +284,12 @@ theorem inv_hom_id (f : X ⟶ Y) [I : IsIso f] : inv f ≫ f = 𝟙 Y := @[simp] theorem hom_inv_id_assoc (f : X ⟶ Y) [I : IsIso f] {Z} (g : X ⟶ Z) : f ≫ inv f ≫ g = g := by simp [← Category.assoc] +#align category_theory.is_iso.hom_inv_id_assoc CategoryTheory.IsIso.hom_inv_id_assoc @[simp] theorem inv_hom_id_assoc (f : X ⟶ Y) [I : IsIso f] {Z} (g : Y ⟶ Z) : inv f ≫ f ≫ g = g := by simp [← Category.assoc] +#align category_theory.is_iso.inv_hom_id_assoc CategoryTheory.IsIso.inv_hom_id_assoc end IsIso @@ -568,6 +578,8 @@ def mapIso (F : C ⥤ D) {X Y : C} (i : X ≅ Y) : F.obj X ≅ F.obj Y where hom_inv_id := by rw [← map_comp, Iso.hom_inv_id, ← map_id] inv_hom_id := by rw [← map_comp, Iso.inv_hom_id, ← map_id] #align category_theory.functor.map_iso CategoryTheory.Functor.mapIso +#align category_theory.functor.map_iso_inv CategoryTheory.Functor.mapIso_inv +#align category_theory.functor.map_iso_hom CategoryTheory.Functor.mapIso_hom @[simp] theorem mapIso_symm (F : C ⥤ D) {X Y : C} (i : X ≅ Y) : F.mapIso i.symm = (F.mapIso i).symm := diff --git a/Mathlib/CategoryTheory/NatIso.lean b/Mathlib/CategoryTheory/NatIso.lean index 5a283ae5bd9a2..3d87a5b4a1c09 100644 --- a/Mathlib/CategoryTheory/NatIso.lean +++ b/Mathlib/CategoryTheory/NatIso.lean @@ -60,18 +60,22 @@ def app {F G : C ⥤ D} (α : F ≅ G) (X : C) : hom_inv_id := by rw [← comp_app, Iso.hom_inv_id]; rfl inv_hom_id := by rw [← comp_app, Iso.inv_hom_id]; rfl #align category_theory.iso.app CategoryTheory.Iso.app +#align category_theory.iso.app_hom CategoryTheory.Iso.app_hom +#align category_theory.iso.app_inv CategoryTheory.Iso.app_inv @[reassoc (attr := simp)] theorem hom_inv_id_app {F G : C ⥤ D} (α : F ≅ G) (X : C) : α.hom.app X ≫ α.inv.app X = 𝟙 (F.obj X) := congr_fun (congr_arg NatTrans.app α.hom_inv_id) X #align category_theory.iso.hom_inv_id_app CategoryTheory.Iso.hom_inv_id_app +#align category_theory.iso.hom_inv_id_app_assoc CategoryTheory.Iso.hom_inv_id_app_assoc @[reassoc (attr := simp)] theorem inv_hom_id_app {F G : C ⥤ D} (α : F ≅ G) (X : C) : α.inv.app X ≫ α.hom.app X = 𝟙 (G.obj X) := congr_fun (congr_arg NatTrans.app α.inv_hom_id) X #align category_theory.iso.inv_hom_id_app CategoryTheory.Iso.inv_hom_id_app +#align category_theory.iso.inv_hom_id_app_assoc CategoryTheory.Iso.inv_hom_id_app_assoc end Iso @@ -181,6 +185,7 @@ theorem naturality_2' (α : F ⟶ G) (f : X ⟶ Y) [IsIso (α.app Y)] : α.app X ≫ G.map f ≫ inv (α.app Y) = F.map f := by rw [← Category.assoc, ← naturality, Category.assoc, IsIso.hom_inv_id, Category.comp_id] #align category_theory.nat_iso.naturality_2' CategoryTheory.NatIso.naturality_2' +#align category_theory.nat_iso.naturality_2'_assoc CategoryTheory.NatIso.naturality_2'_assoc /-- The components of a natural isomorphism are isomorphisms. -/ @@ -221,6 +226,8 @@ def ofComponents (app : ∀ X : C, F.obj X ≅ G.obj X) simp only [Iso.inv_hom_id_assoc, Iso.hom_inv_id, assoc, comp_id, cancel_mono] at h exact h } #align category_theory.nat_iso.of_components CategoryTheory.NatIso.ofComponents +#align category_theory.nat_iso.of_components_hom_app CategoryTheory.NatIso.ofComponents_hom_app +#align category_theory.nat_iso.of_components_inv_app CategoryTheory.NatIso.ofComponents_inv_app @[simp] theorem ofComponents.app (app' : ∀ X : C, F.obj X ≅ G.obj X) (naturality) (X) : @@ -243,6 +250,8 @@ def hcomp {F G : C ⥤ D} {H I : D ⥤ E} (α : F ≅ G) (β : H ≅ I) : F ⋙ simp ext; rw [← NatTrans.exchange]; simp #align category_theory.nat_iso.hcomp CategoryTheory.NatIso.hcomp +#align category_theory.nat_iso.hcomp_inv CategoryTheory.NatIso.hcomp_inv +#align category_theory.nat_iso.hcomp_hom CategoryTheory.NatIso.hcomp_hom theorem isIso_map_iff {F₁ F₂ : C ⥤ D} (e : F₁ ≅ F₂) {X Y : C} (f : X ⟶ Y) : IsIso (F₁.map f) ↔ IsIso (F₂.map f) := by diff --git a/Mathlib/CategoryTheory/NatTrans.lean b/Mathlib/CategoryTheory/NatTrans.lean index 03c905b1b8734..7a43f7fc94c9c 100644 --- a/Mathlib/CategoryTheory/NatTrans.lean +++ b/Mathlib/CategoryTheory/NatTrans.lean @@ -55,6 +55,9 @@ structure NatTrans (F G : C ⥤ D) : Type max u₁ v₂ where /-- The naturality square for a given morphism. -/ naturality : ∀ ⦃X Y : C⦄ (f : X ⟶ Y), F.map f ≫ app Y = app X ≫ G.map f := by aesop_cat #align category_theory.nat_trans CategoryTheory.NatTrans +#align category_theory.nat_trans.naturality CategoryTheory.NatTrans.naturality +#align category_theory.nat_trans.ext_iff CategoryTheory.NatTrans.ext_iff +#align category_theory.nat_trans.ext CategoryTheory.NatTrans.ext -- TODO Perhaps we should just turn on `ext` in aesop? attribute [aesop safe apply (rule_sets [CategoryTheory])] NatTrans.ext @@ -62,6 +65,7 @@ attribute [aesop safe apply (rule_sets [CategoryTheory])] NatTrans.ext -- Rather arbitrarily, we say that the 'simpler' form is -- components of natural transfomations moving earlier. attribute [reassoc (attr := simp)] NatTrans.naturality +#align category_theory.nat_trans.naturality_assoc CategoryTheory.NatTrans.naturality_assoc theorem congr_app {F G : C ⥤ D} {α β : NatTrans F G} (h : α = β) (X : C) : α.app X = β.app X := by aesop_cat diff --git a/Mathlib/CategoryTheory/Whiskering.lean b/Mathlib/CategoryTheory/Whiskering.lean index 5e52f8fd93c9d..d73625024a89e 100644 --- a/Mathlib/CategoryTheory/Whiskering.lean +++ b/Mathlib/CategoryTheory/Whiskering.lean @@ -53,6 +53,7 @@ def whiskerLeft (F : C ⥤ D) {G H : D ⥤ E} (α : G ⟶ H) : app X := α.app (F.obj X) naturality X Y f := by rw [Functor.comp_map, Functor.comp_map, α.naturality] #align category_theory.whisker_left CategoryTheory.whiskerLeft +#align category_theory.whisker_left_app CategoryTheory.whiskerLeft_app /-- If `α : G ⟶ H` then `whisker_right α F : (G ⋙ F) ⟶ (G ⋙ F)` has components `F.map (α.app X)`. @@ -64,6 +65,7 @@ def whiskerRight {G H : C ⥤ D} (α : G ⟶ H) (F : D ⥤ E) : naturality X Y f := by rw [Functor.comp_map, Functor.comp_map, ← F.map_comp, ← F.map_comp, α.naturality] #align category_theory.whisker_right CategoryTheory.whiskerRight +#align category_theory.whisker_right_app CategoryTheory.whiskerRight_app variable (C D E) @@ -83,6 +85,9 @@ def whiskeringLeft : (C ⥤ D) ⥤ (D ⥤ E) ⥤ C ⥤ E where naturality := fun X Y f => by dsimp; rw [← H.map_comp, ← H.map_comp, ← τ.naturality] } naturality := fun X Y f => by ext; dsimp; rw [f.naturality] } #align category_theory.whiskering_left CategoryTheory.whiskeringLeft +#align category_theory.whiskering_left_obj_map CategoryTheory.whiskeringLeft_obj_map +#align category_theory.whiskering_left_obj_obj CategoryTheory.whiskeringLeft_obj_obj +#align category_theory.whiskering_left_map_app_app CategoryTheory.whiskeringLeft_map_app_app /-- Right-composition gives a functor `(D ⥤ E) ⥤ ((C ⥤ D) ⥤ (C ⥤ E))`. @@ -100,6 +105,9 @@ def whiskeringRight : (D ⥤ E) ⥤ (C ⥤ D) ⥤ C ⥤ E where naturality := fun X Y f => by dsimp; rw [τ.naturality] } naturality := fun X Y f => by ext; dsimp; rw [← NatTrans.naturality] } #align category_theory.whiskering_right CategoryTheory.whiskeringRight +#align category_theory.whiskering_right_map_app_app CategoryTheory.whiskeringRight_map_app_app +#align category_theory.whiskering_right_obj_obj CategoryTheory.whiskeringRight_obj_obj +#align category_theory.whiskering_right_obj_map CategoryTheory.whiskeringRight_obj_map variable {C} {D} {E} @@ -232,6 +240,8 @@ def leftUnitor (F : A ⥤ B) : hom := { app := fun X => 𝟙 (F.obj X) } inv := { app := fun X => 𝟙 (F.obj X) } #align category_theory.functor.left_unitor CategoryTheory.Functor.leftUnitor +#align category_theory.functor.left_unitor_inv_app CategoryTheory.Functor.leftUnitor_inv_app +#align category_theory.functor.left_unitor_hom_app CategoryTheory.Functor.leftUnitor_hom_app /-- The right unitor, a natural isomorphism `(F ⋙ (𝟭 B)) ≅ F`. -/ @@ -241,6 +251,8 @@ def rightUnitor (F : A ⥤ B) : hom := { app := fun X => 𝟙 (F.obj X) } inv := { app := fun X => 𝟙 (F.obj X) } #align category_theory.functor.right_unitor CategoryTheory.Functor.rightUnitor +#align category_theory.functor.right_unitor_hom_app CategoryTheory.Functor.rightUnitor_hom_app +#align category_theory.functor.right_unitor_inv_app CategoryTheory.Functor.rightUnitor_inv_app variable {C : Type u₃} [Category.{v₃} C] @@ -257,6 +269,8 @@ def associator (F : A ⥤ B) (G : B ⥤ C) (H : C ⥤ D) : hom := { app := fun _ => 𝟙 _ } inv := { app := fun _ => 𝟙 _ } #align category_theory.functor.associator CategoryTheory.Functor.associator +#align category_theory.functor.associator_inv_app CategoryTheory.Functor.associator_inv_app +#align category_theory.functor.associator_hom_app CategoryTheory.Functor.associator_hom_app protected theorem assoc (F : A ⥤ B) (G : B ⥤ C) (H : C ⥤ D) : (F ⋙ G) ⋙ H = F ⋙ G ⋙ H := rfl diff --git a/Mathlib/Combinatorics/Quiver/Basic.lean b/Mathlib/Combinatorics/Quiver/Basic.lean index 485697b95ed57..de22e5132db0e 100644 --- a/Mathlib/Combinatorics/Quiver/Basic.lean +++ b/Mathlib/Combinatorics/Quiver/Basic.lean @@ -88,6 +88,8 @@ def id (V : Type _) [Quiver V] : Prefunctor V V where obj := fun X => X map f := f #align prefunctor.id Prefunctor.id +#align prefunctor.id_obj Prefunctor.id_obj +#align prefunctor.id_map Prefunctor.id_map instance (V : Type _) [Quiver V] : Inhabited (Prefunctor V V) := ⟨id V⟩ @@ -99,6 +101,8 @@ def comp {U : Type _} [Quiver U] {V : Type _} [Quiver V] {W : Type _} [Quiver W] obj X := G.obj (F.obj X) map f := G.map (F.map f) #align prefunctor.comp Prefunctor.comp +#align prefunctor.comp_obj Prefunctor.comp_obj +#align prefunctor.comp_map Prefunctor.comp_map @[simp] theorem comp_id {U V : Type _} [Quiver U] [Quiver V] (F : Prefunctor U V) : diff --git a/Mathlib/Combinatorics/Quiver/ConnectedComponent.lean b/Mathlib/Combinatorics/Quiver/ConnectedComponent.lean index aaf967cb49d40..cb728c09e42cb 100644 --- a/Mathlib/Combinatorics/Quiver/ConnectedComponent.lean +++ b/Mathlib/Combinatorics/Quiver/ConnectedComponent.lean @@ -34,12 +34,14 @@ variable (V : Type _) [Quiver.{u+1} V] def zigzagSetoid : Setoid V := ⟨fun a b ↦ Nonempty (@Path (Symmetrify V) _ a b), fun _ ↦ ⟨Path.nil⟩, fun ⟨p⟩ ↦ ⟨p.reverse⟩, fun ⟨p⟩ ⟨q⟩ ↦ ⟨p.comp q⟩⟩ +#align quiver.zigzag_setoid Quiver.zigzagSetoid /-- The type of weakly connected components of a directed graph. Two vertices are in the same weakly connected component if there is a zigzag of arrows from one to the other. -/ def WeaklyConnectedComponent : Type _ := Quotient (zigzagSetoid V) +#align quiver.weakly_connected_component Quiver.WeaklyConnectedComponent namespace WeaklyConnectedComponent @@ -48,6 +50,7 @@ variable {V} /-- The weakly connected component corresponding to a vertex. -/ protected def mk : V → WeaklyConnectedComponent V := @Quotient.mk' _ (zigzagSetoid V) +#align quiver.weakly_connected_component.mk Quiver.WeaklyConnectedComponent.mk instance : CoeTC V (WeaklyConnectedComponent V) := ⟨WeaklyConnectedComponent.mk⟩ @@ -58,6 +61,7 @@ instance [Inhabited V] : Inhabited (WeaklyConnectedComponent V) := protected theorem eq (a b : V) : (a : WeaklyConnectedComponent V) = b ↔ Nonempty (@Path (Symmetrify V) _ a b) := Quotient.eq'' +#align quiver.weakly_connected_component.eq Quiver.WeaklyConnectedComponent.eq end WeaklyConnectedComponent @@ -69,5 +73,6 @@ variable {V} an arrow `e` if either `e` or its reversal is in `H`. -/ def wideSubquiverSymmetrify (H : WideSubquiver (Symmetrify V)) : WideSubquiver V := fun _ _ ↦ { e | H _ _ (Sum.inl e) ∨ H _ _ (Sum.inr e) } +#align quiver.wide_subquiver_symmetrify Quiver.wideSubquiverSymmetrify end Quiver diff --git a/Mathlib/Combinatorics/Quiver/Path.lean b/Mathlib/Combinatorics/Quiver/Path.lean index c08250bc94dc1..edea1a0f49ddf 100644 --- a/Mathlib/Combinatorics/Quiver/Path.lean +++ b/Mathlib/Combinatorics/Quiver/Path.lean @@ -29,10 +29,12 @@ namespace Quiver inductive Path {V : Type u} [Quiver.{v} V] (a : V) : V → Sort max (u + 1) v | nil : Path a a | cons : ∀ {b c : V}, Path a b → (b ⟶ c) → Path a c +#align quiver.path Quiver.Path /-- An arrow viewed as a path of length one. -/ def Hom.toPath {V} [Quiver V] {a b : V} (e : a ⟶ b) : Path a b := Path.nil.cons e +#align quiver.hom.to_path Quiver.Hom.toPath namespace Path @@ -40,23 +42,29 @@ variable {V : Type u} [Quiver V] {a b c d : V} lemma nil_ne_cons (p : Path a b) (e : b ⟶ a) : Path.nil ≠ p.cons e := fun h => by injection h +#align quiver.path.nil_ne_cons Quiver.Path.nil_ne_cons lemma cons_ne_nil (p : Path a b) (e : b ⟶ a) : p.cons e ≠ Path.nil := fun h => by injection h +#align quiver.path.cons_ne_nil Quiver.Path.cons_ne_nil lemma obj_eq_of_cons_eq_cons {p : Path a b} {p' : Path a c} {e : b ⟶ d} {e' : c ⟶ d} (h : p.cons e = p'.cons e') : b = c := by injection h +#align quiver.path.obj_eq_of_cons_eq_cons Quiver.Path.obj_eq_of_cons_eq_cons lemma heq_of_cons_eq_cons {p : Path a b} {p' : Path a c} {e : b ⟶ d} {e' : c ⟶ d} (h : p.cons e = p'.cons e') : HEq p p' := by injection h +#align quiver.path.heq_of_cons_eq_cons Quiver.Path.heq_of_cons_eq_cons lemma hom_heq_of_cons_eq_cons {p : Path a b} {p' : Path a c} {e : b ⟶ d} {e' : c ⟶ d} (h : p.cons e = p'.cons e') : HEq e e' := by injection h +#align quiver.path.hom_heq_of_cons_eq_cons Quiver.Path.hom_heq_of_cons_eq_cons /-- The length of a path is the number of arrows it uses. -/ def length {a : V} : ∀ {b : V}, Path a b → ℕ | _, nil => 0 | _, cons p _ => p.length + 1 +#align quiver.path.length Quiver.Path.length instance {a : V} : Inhabited (Path a a) := ⟨nil⟩ @@ -64,45 +72,54 @@ instance {a : V} : Inhabited (Path a a) := @[simp] theorem length_nil {a : V} : (nil : Path a a).length = 0 := rfl +#align quiver.path.length_nil Quiver.Path.length_nil @[simp] theorem length_cons (a b c : V) (p : Path a b) (e : b ⟶ c) : (p.cons e).length = p.length + 1 := rfl +#align quiver.path.length_cons Quiver.Path.length_cons theorem eq_of_length_zero (p : Path a b) (hzero : p.length = 0) : a = b := by cases p · rfl · cases Nat.succ_ne_zero _ hzero +#align quiver.path.eq_of_length_zero Quiver.Path.eq_of_length_zero /-- Composition of paths. -/ def comp {a b : V} : ∀ {c}, Path a b → Path b c → Path a c | _, p, nil => p | _, p, cons q e => (p.comp q).cons e +#align quiver.path.comp Quiver.Path.comp @[simp] theorem comp_cons {a b c d : V} (p : Path a b) (q : Path b c) (e : c ⟶ d) : p.comp (q.cons e) = (p.comp q).cons e := rfl +#align quiver.path.comp_cons Quiver.Path.comp_cons @[simp] theorem comp_nil {a b : V} (p : Path a b) : p.comp Path.nil = p := rfl +#align quiver.path.comp_nil Quiver.Path.comp_nil @[simp] theorem nil_comp {a : V} : ∀ {b} (p : Path a b), Path.nil.comp p = p | _, nil => rfl | _, cons p _ => by rw [comp_cons, nil_comp p] +#align quiver.path.nil_comp Quiver.Path.nil_comp @[simp] theorem comp_assoc {a b c : V} : ∀ {d} (p : Path a b) (q : Path b c) (r : Path c d), (p.comp q).comp r = p.comp (q.comp r) | _, _, _, nil => rfl | _, p, q, cons r _ => by rw [comp_cons, comp_cons, comp_cons, comp_assoc p q r] +#align quiver.path.comp_assoc Quiver.Path.comp_assoc @[simp] theorem length_comp (p : Path a b) : ∀ {c} (q : Path b c), (p.comp q).length = p.length + q.length | _, nil => rfl | _, cons _ _ => congr_arg Nat.succ (length_comp _ _) +#align quiver.path.length_comp Quiver.Path.length_comp theorem comp_inj {p₁ p₂ : Path a b} {q₁ q₂ : Path b c} (hq : q₁.length = q₂.length) : p₁.comp q₁ = p₂.comp q₂ ↔ p₁ = p₂ ∧ q₁ = q₂ := by @@ -116,6 +133,7 @@ theorem comp_inj {p₁ p₂ : Path a b} {q₁ q₂ : Path b c} (hq : q₁.length obtain ⟨rfl, rfl⟩ := ih (Nat.succ.inj hq) h.2.1.eq rw [h.2.2.eq] exact ⟨rfl, rfl⟩ +#align quiver.path.comp_inj Quiver.Path.comp_inj theorem comp_inj' {p₁ p₂ : Path a b} {q₁ q₂ : Path b c} (h : p₁.length = p₂.length) : p₁.comp q₁ = p₂.comp q₂ ↔ p₁ = p₂ ∧ q₁ = q₂ := @@ -123,26 +141,32 @@ theorem comp_inj' {p₁ p₂ : Path a b} {q₁ q₂ : Path b c} (h : p₁.length by rintro ⟨rfl, rfl⟩ rfl⟩ +#align quiver.path.comp_inj' Quiver.Path.comp_inj' theorem comp_injective_left (q : Path b c) : Injective fun p : Path a b => p.comp q := fun _ _ h => ((comp_inj rfl).1 h).1 +#align quiver.path.comp_injective_left Quiver.Path.comp_injective_left theorem comp_injective_right (p : Path a b) : Injective (p.comp : Path b c → Path a c) := fun _ _ h => ((comp_inj' rfl).1 h).2 +#align quiver.path.comp_injective_right Quiver.Path.comp_injective_right @[simp] theorem comp_inj_left {p₁ p₂ : Path a b} {q : Path b c} : p₁.comp q = p₂.comp q ↔ p₁ = p₂ := q.comp_injective_left.eq_iff +#align quiver.path.comp_inj_left Quiver.Path.comp_inj_left @[simp] theorem comp_inj_right {p : Path a b} {q₁ q₂ : Path b c} : p.comp q₁ = p.comp q₂ ↔ q₁ = q₂ := p.comp_injective_right.eq_iff +#align quiver.path.comp_inj_right Quiver.Path.comp_inj_right /-- Turn a path into a list. The list contains `a` at its head, but not `b` a priori. -/ @[simp] def toList : ∀ {b : V}, Path a b → List V | _, nil => [] | _, @cons _ _ _ c _ p _ => c :: p.toList +#align quiver.path.to_list Quiver.Path.toList /-- `Quiver.Path.toList` is a contravariant functor. The inversion comes from `Quiver.Path` and `List` having different preferred directions for adding elements. -/ diff --git a/Mathlib/Combinatorics/Quiver/SingleObj.lean b/Mathlib/Combinatorics/Quiver/SingleObj.lean index b861b9758039e..1d6ac244ca2b7 100644 --- a/Mathlib/Combinatorics/Quiver/SingleObj.lean +++ b/Mathlib/Combinatorics/Quiver/SingleObj.lean @@ -81,6 +81,8 @@ def hasInvolutiveReverse (rev : α → α) (h : Function.Involutive rev) : def toHom : α ≃ (star α ⟶ star α) := Equiv.refl _ #align quiver.single_obj.to_hom Quiver.SingleObj.toHom +#align quiver.single_obj.to_hom_apply Quiver.SingleObj.toHom_apply +#align quiver.single_obj.to_hom_symm_apply Quiver.SingleObj.toHom_symm_apply /-- Prefunctors between two `SingleObj` quivers correspond to functions between the corresponding arrows types. @@ -92,6 +94,9 @@ def toPrefunctor : (α → β) ≃ SingleObj α ⥤q SingleObj β invFun f a := f.map (toHom a) left_inv _ := rfl right_inv _ := rfl +#align quiver.single_obj.to_prefunctor_symm_apply Quiver.SingleObj.toPrefunctor_symm_apply +#align quiver.single_obj.to_prefunctor_apply_map Quiver.SingleObj.toPrefunctor_apply_map +#align quiver.single_obj.to_prefunctor_apply_obj Quiver.SingleObj.toPrefunctor_apply_obj #align quiver.single_obj.to_prefunctor Quiver.SingleObj.toPrefunctor diff --git a/Mathlib/Combinatorics/Quiver/Subquiver.lean b/Mathlib/Combinatorics/Quiver/Subquiver.lean index 78cfa835377cf..00f17f400cdf2 100644 --- a/Mathlib/Combinatorics/Quiver/Subquiver.lean +++ b/Mathlib/Combinatorics/Quiver/Subquiver.lean @@ -28,6 +28,7 @@ universe v u NB: this does not work for `Prop`-valued quivers. It requires `G : Quiver.{v+1} V`. -/ def WideSubquiver (V) [Quiver.{v + 1} V] := ∀ a b : V, Set (a ⟶ b) +#align wide_subquiver WideSubquiver /-- A type synonym for `V`, when thought of as a quiver having only the arrows from some `WideSubquiver`. -/ @@ -66,6 +67,9 @@ structure Total (V : Type u) [Quiver.{v} V] : Sort max (u + 1) v where right : V /-- an arrow -/ hom : left ⟶ right +#align quiver.total Quiver.Total +#align quiver.total.ext Quiver.Total.ext +#align quiver.total.ext_iff Quiver.Total.ext_iff /-- A wide subquiver of `G` can equivalently be viewed as a total set of arrows. -/ def wideSubquiverEquivSetTotal {V} [Quiver V] : @@ -75,10 +79,12 @@ def wideSubquiverEquivSetTotal {V} [Quiver V] : invFun S a b := { e | Total.mk a b e ∈ S } left_inv _ := rfl right_inv _ := rfl +#align quiver.wide_subquiver_equiv_set_total Quiver.wideSubquiverEquivSetTotal /-- An `L`-labelling of a quiver assigns to every arrow an element of `L`. -/ def Labelling (V : Type u) [Quiver V] (L : Sort _) := ∀ ⦃a b : V⦄, (a ⟶ b) → L +#align quiver.labelling Quiver.Labelling instance {V : Type u} [Quiver V] (L) [Inhabited L] : Inhabited (Labelling V L) := ⟨fun _ _ _ ↦ default⟩ diff --git a/Mathlib/Combinatorics/Quiver/Symmetric.lean b/Mathlib/Combinatorics/Quiver/Symmetric.lean index 6fe19c7523690..b211df8fbce88 100644 --- a/Mathlib/Combinatorics/Quiver/Symmetric.lean +++ b/Mathlib/Combinatorics/Quiver/Symmetric.lean @@ -35,6 +35,7 @@ namespace Quiver NB: this does not work for `Prop`-valued quivers. It requires `[Quiver.{v+1} V]`. -/ -- Porting note: no hasNonemptyInstance linter yet def Symmetrify (V : Type _) := V +#align quiver.symmetrify Quiver.Symmetrify instance symmetrifyQuiver (V : Type u) [Quiver V] : Quiver (Symmetrify V) := ⟨fun a b : V ↦ Sum (a ⟶ b) (b ⟶ a)⟩ @@ -46,15 +47,18 @@ variable (U V W : Type _) [Quiver.{u + 1} U] [Quiver.{v + 1} V] [Quiver.{w + 1} class HasReverse where /-- the map which sends an arrow to its reverse -/ reverse' : ∀ {a b : V}, (a ⟶ b) → (b ⟶ a) +#align quiver.has_reverse Quiver.HasReverse /-- Reverse the direction of an arrow. -/ def reverse {V} [Quiver.{v + 1} V] [HasReverse V] {a b : V} : (a ⟶ b) → (b ⟶ a) := HasReverse.reverse' +#align quiver.reverse Quiver.reverse /-- A quiver `HasInvolutiveReverse` if reversing twice is the identity.`-/ class HasInvolutiveReverse extends HasReverse V where /-- `reverse` is involutive -/ inv' : ∀ {a b : V} (f : a ⟶ b), reverse (reverse f) = f +#align quiver.has_involutive_reverse Quiver.HasInvolutiveReverse variable {U V W} @@ -140,6 +144,7 @@ abbrev Hom.toNeg {X Y : V} (f : X ⟶ Y) : (Quiver.symmetrifyQuiver V).Hom Y X : def Path.reverse [HasReverse V] {a : V} : ∀ {b}, Path a b → Path b a | _, Path.nil => Path.nil | _, Path.cons p e => (Quiver.reverse e).toPath.comp p.reverse +#align quiver.path.reverse Quiver.Path.reverse @[simp] theorem Path.reverse_toPath [HasReverse V] {a b : V} (f : a ⟶ b) : @@ -153,6 +158,7 @@ theorem Path.reverse_comp [HasReverse V] {a b c : V} (p : Path a b) (q : Path b induction' q with _ _ _ _ h · simp · simp [h] +#align quiver.path.reverse_comp Quiver.Path.reverse_comp @[simp] theorem Path.reverse_reverse [h : HasInvolutiveReverse V] {a b : V} (p : Path a b) : @@ -161,6 +167,7 @@ theorem Path.reverse_reverse [h : HasInvolutiveReverse V] {a b : V} (p : Path a · simp · rw [Path.reverse, Path.reverse_comp, h, Path.reverse_toPath, Quiver.reverse_reverse] rfl +#align quiver.path.reverse_reverse Quiver.Path.reverse_reverse end Paths @@ -170,6 +177,7 @@ namespace Symmetrify def of : Prefunctor V (Symmetrify V) where obj := id map := Sum.inl +#align quiver.symmetrify.of Quiver.Symmetrify.of variable {V' : Type _} [Quiver.{v' + 1} V'] @@ -181,6 +189,7 @@ def lift [HasReverse V'] (φ : Prefunctor V V') : map f := match f with | Sum.inl g => φ.map g | Sum.inr g => reverse (φ.map g) +#align quiver.symmetrify.lift Quiver.Symmetrify.lift theorem lift_spec [HasReverse V'] (φ : Prefunctor V V') : Symmetrify.of.comp (Symmetrify.lift φ) = φ := by @@ -189,6 +198,7 @@ theorem lift_spec [HasReverse V'] (φ : Prefunctor V V') : rfl · rintro X Y f rfl +#align quiver.symmetrify.lift_spec Quiver.Symmetrify.lift_spec theorem lift_reverse [h : HasInvolutiveReverse V'] (φ : Prefunctor V V') {X Y : Symmetrify V} (f : X ⟶ Y) : @@ -198,6 +208,7 @@ theorem lift_reverse [h : HasInvolutiveReverse V'] rfl · simp only [reverse_reverse] rfl +#align quiver.symmetrify.lift_reverse Quiver.Symmetrify.lift_reverse /-- `lift φ` is the only prefunctor extending `φ` and preserving reverses. -/ theorem lift_unique [HasReverse V'] (φ : V ⥤q V') (Φ : Symmetrify V ⥤q V') (hΦ : (of ⋙q Φ) = φ) diff --git a/Mathlib/Data/Bool/Basic.lean b/Mathlib/Data/Bool/Basic.lean index 799734c812efe..3ec34b1277dc5 100644 --- a/Mathlib/Data/Bool/Basic.lean +++ b/Mathlib/Data/Bool/Basic.lean @@ -97,12 +97,15 @@ theorem beq_comm [BEq α] [LawfulBEq α] {a b : α} : (a == b) = (b == a) := @[simp] theorem default_bool : default = false := rfl +#align bool.default_bool Bool.default_bool theorem dichotomy (b : Bool) : b = false ∨ b = true := by cases b <;> simp +#align bool.dichotomy Bool.dichotomy @[simp] theorem forall_bool {p : Bool → Prop} : (∀ b, p b) ↔ p false ∧ p true := ⟨fun h ↦ by simp [h], fun ⟨h₁, h₂⟩ b ↦ by cases b <;> assumption⟩ +#align bool.forall_bool Bool.forall_bool @[simp] theorem exists_bool {p : Bool → Prop} : (∃ b, p b) ↔ p false ∨ p true := @@ -110,6 +113,7 @@ theorem exists_bool {p : Bool → Prop} : (∃ b, p b) ↔ p false ∨ p true := fun h ↦ match h with | .inl h => ⟨_, h⟩ | .inr h => ⟨_, h⟩ ⟩ +#align bool.exists_bool Bool.exists_bool /-- If `p b` is decidable for all `b : bool`, then `∀ b, p b` is decidable -/ instance decidableForallBool {p : Bool → Prop} [∀ b, Decidable (p b)] : Decidable (∀ b, p b) := @@ -123,6 +127,7 @@ instance decidableExistsBool {p : Bool → Prop} [∀ b, Decidable (p b)] : Deci theorem cond_eq_ite {α} (b : Bool) (t e : α) : cond b t e = if b then t else e := by cases b <;> simp +#align bool.cond_eq_ite Bool.cond_eq_ite @[simp] theorem cond_decide {α} (p : Prop) [Decidable p] (t e : α) : @@ -138,6 +143,7 @@ theorem not_ne_id : not ≠ id := fun h ↦ ff_ne_tt <| congrFun h true #align bool.bnot_ne_id Bool.not_ne_id theorem coe_bool_iff : ∀ {a b : Bool}, (a ↔ b) ↔ a = b := by decide +#align bool.coe_bool_iff Bool.coe_bool_iff theorem eq_true_of_ne_false : ∀ {a : Bool}, a ≠ false → a = true := by decide #align bool.eq_tt_of_ne_ff Bool.eq_true_of_ne_false @@ -339,6 +345,7 @@ theorem le_true {x : Bool} : x ≤ true := #align bool.le_tt Bool.le_true theorem lt_iff : ∀ {x y : Bool}, x < y ↔ x = false ∧ y = true := by decide +#align bool.lt_iff Bool.lt_iff @[simp] theorem false_lt_true : false < true := @@ -346,6 +353,7 @@ theorem false_lt_true : false < true := #align bool.ff_lt_tt Bool.false_lt_true theorem le_iff_imp : ∀ {x y : Bool}, x ≤ y ↔ x → y := by decide +#align bool.le_iff_imp Bool.le_iff_imp theorem and_le_left : ∀ x y : Bool, (x && y) ≤ x := by decide #align bool.band_le_left Bool.and_le_left @@ -383,23 +391,28 @@ theorem ofNat_le_ofNat {n m : Nat} (h : n ≤ m) : ofNat n ≤ ofNat m := by cases Nat.decEq m 0 with | isFalse hm => rw [decide_eq_false hm]; exact le_true | isTrue hm => subst hm; have h := le_antisymm h (Nat.zero_le n); contradiction +#align bool.of_nat_le_of_nat Bool.ofNat_le_ofNat theorem toNat_le_toNat {b₀ b₁ : Bool} (h : b₀ ≤ b₁) : toNat b₀ ≤ toNat b₁ := by cases h with | inl h => subst h; exact Nat.zero_le _ | inr h => subst h; cases b₀ <;> simp; +#align bool.to_nat_le_to_nat Bool.toNat_le_toNat theorem ofNat_toNat (b : Bool) : ofNat (toNat b) = b := by cases b <;> rfl +#align bool.of_nat_to_nat Bool.ofNat_toNat @[simp] theorem injective_iff {α : Sort _} {f : Bool → α} : Function.Injective f ↔ f false ≠ f true := ⟨fun Hinj Heq ↦ ff_ne_tt (Hinj Heq), fun H x y hxy ↦ by cases x <;> cases y exacts[rfl, (H hxy).elim, (H hxy.symm).elim, rfl]⟩ +#align bool.injective_iff Bool.injective_iff /-- **Kaminski's Equation** -/ theorem apply_apply_apply (f : Bool → Bool) (x : Bool) : f (f (f x)) = f x := by cases x <;> cases h₁ : f true <;> cases h₂ : f false <;> simp only [h₁, h₂] +#align bool.apply_apply_apply Bool.apply_apply_apply end Bool diff --git a/Mathlib/Data/Countable/Defs.lean b/Mathlib/Data/Countable/Defs.lean index 85eed6af0167f..afac1ad19e9cd 100644 --- a/Mathlib/Data/Countable/Defs.lean +++ b/Mathlib/Data/Countable/Defs.lean @@ -53,23 +53,29 @@ protected theorem Function.Injective.countable [Countable β] {f : α → β} (h Countable α := let ⟨g, hg⟩ := exists_injective_nat β ⟨⟨g ∘ f, hg.comp hf⟩⟩ +#align function.injective.countable Function.Injective.countable protected theorem Function.Surjective.countable [Countable α] {f : α → β} (hf : Surjective f) : Countable β := (injective_surjInv hf).countable +#align function.surjective.countable Function.Surjective.countable theorem exists_surjective_nat (α : Sort u) [Nonempty α] [Countable α] : ∃ f : ℕ → α, Surjective f := let ⟨f, hf⟩ := exists_injective_nat α ⟨invFun f, invFun_surjective hf⟩ +#align exists_surjective_nat exists_surjective_nat theorem countable_iff_exists_surjective [Nonempty α] : Countable α ↔ ∃ f : ℕ → α, Surjective f := ⟨@exists_surjective_nat _ _, fun ⟨_, hf⟩ ↦ hf.countable⟩ +#align countable_iff_exists_surjective countable_iff_exists_surjective theorem Countable.of_equiv (α : Sort _) [Countable α] (e : α ≃ β) : Countable β := e.symm.injective.countable +#align countable.of_equiv Countable.of_equiv theorem Equiv.countable_iff (e : α ≃ β) : Countable α ↔ Countable β := ⟨fun h => @Countable.of_equiv _ _ h e, fun h => @Countable.of_equiv _ _ h e.symm⟩ +#align equiv.countable_iff Equiv.countable_iff instance {β : Type v} [Countable β] : Countable (ULift.{u} β) := Countable.of_equiv _ Equiv.ulift.symm diff --git a/Mathlib/Data/Dfinsupp/Basic.lean b/Mathlib/Data/Dfinsupp/Basic.lean index 990e03c4e6c95..28034dc3fa235 100644 --- a/Mathlib/Data/Dfinsupp/Basic.lean +++ b/Mathlib/Data/Dfinsupp/Basic.lean @@ -449,6 +449,7 @@ def filterAddMonoidHom [∀ i, AddZeroClass (β i)] (p : ι → Prop) [Decidable map_zero' := filter_zero p map_add' := filter_add p #align dfinsupp.filter_add_monoid_hom Dfinsupp.filterAddMonoidHom +#align dfinsupp.filter_add_monoid_hom_apply Dfinsupp.filterAddMonoidHom_apply /-- `Dfinsupp.filter` as a `LinearMap`. -/ @[simps] @@ -459,6 +460,7 @@ def filterLinearMap [Semiring γ] [∀ i, AddCommMonoid (β i)] [∀ i, Module map_add' := filter_add p map_smul' := filter_smul p #align dfinsupp.filter_linear_map Dfinsupp.filterLinearMap +#align dfinsupp.filter_linear_map_apply Dfinsupp.filterLinearMap_apply variable {γ β} @@ -522,6 +524,7 @@ def subtypeDomainAddMonoidHom [∀ i, AddZeroClass (β i)] (p : ι → Prop) [De map_zero' := subtypeDomain_zero map_add' := subtypeDomain_add #align dfinsupp.subtype_domain_add_monoid_hom Dfinsupp.subtypeDomainAddMonoidHom +#align dfinsupp.subtype_domain_add_monoid_hom_apply Dfinsupp.subtypeDomainAddMonoidHom_apply /-- `Dfinsupp.subtypeDomain` as a `LinearMap`. -/ @[simps] @@ -532,6 +535,7 @@ def subtypeDomainLinearMap [Semiring γ] [∀ i, AddCommMonoid (β i)] [∀ i, M map_add' := subtypeDomain_add map_smul' := subtypeDomain_smul #align dfinsupp.subtype_domain_linear_map Dfinsupp.subtypeDomainLinearMap +#align dfinsupp.subtype_domain_linear_map_apply Dfinsupp.subtypeDomainLinearMap_apply variable {γ β} @@ -611,6 +615,7 @@ def equivFunOnFintype [Fintype ι] : (Π₀ i, β i) ≃ ∀ i, β i left_inv _ := FunLike.coe_injective rfl right_inv _ := rfl #align dfinsupp.equiv_fun_on_fintype Dfinsupp.equivFunOnFintype +#align dfinsupp.equiv_fun_on_fintype_apply Dfinsupp.equivFunOnFintype_apply @[simp] theorem equivFunOnFintype_symm_coe [Fintype ι] (f : Π₀ i, β i) : equivFunOnFintype.symm f = f := @@ -890,6 +895,7 @@ def singleAddHom (i : ι) : β i →+ Π₀ i, β i map_zero' := single_zero i map_add' := single_add i #align dfinsupp.single_add_hom Dfinsupp.singleAddHom +#align dfinsupp.single_add_hom_apply Dfinsupp.singleAddHom_apply /-- `Dfinsupp.erase` as an `AddMonoidHom`. -/ @[simps] @@ -899,6 +905,7 @@ def eraseAddHom (i : ι) : (Π₀ i, β i) →+ Π₀ i, β i map_zero' := erase_zero i map_add' := erase_add i #align dfinsupp.erase_add_hom Dfinsupp.eraseAddHom +#align dfinsupp.erase_add_hom_apply Dfinsupp.eraseAddHom_apply variable {β} @@ -1415,6 +1422,7 @@ def equivCongrLeft [∀ i, Zero (β i)] (h : ι ≃ κ) : (Π₀ i, β i) ≃ Π rw [comapDomain'_apply, mapRange_apply, comapDomain'_apply, Equiv.cast_eq_iff_heq, h.apply_symm_apply] #align dfinsupp.equiv_congr_left Dfinsupp.equivCongrLeft +#align dfinsupp.equiv_congr_left_apply Dfinsupp.equivCongrLeft_apply section Curry @@ -1676,6 +1684,8 @@ noncomputable def equivProdDfinsupp [∀ i, Zero (α i)] : (Π₀ i, α i) ≃ · exact extendWith_none x.snd _ · rw [comapDomain_apply, extendWith_some] #align dfinsupp.equiv_prod_dfinsupp Dfinsupp.equivProdDfinsupp +#align dfinsupp.equiv_prod_dfinsupp_apply Dfinsupp.equivProdDfinsupp_apply +#align dfinsupp.equiv_prod_dfinsupp_symm_apply Dfinsupp.equivProdDfinsupp_symm_apply theorem equivProdDfinsupp_add [∀ i, AddZeroClass (α i)] (f g : Π₀ i, α i) : equivProdDfinsupp (f + g) = equivProdDfinsupp f + equivProdDfinsupp g := @@ -2024,6 +2034,8 @@ def liftAddHom [∀ i, AddZeroClass (β i)] [AddCommMonoid γ] : (∀ i, β i ext simp [sumAddHom_apply, sum, Finset.sum_add_distrib] #align dfinsupp.lift_add_hom Dfinsupp.liftAddHom +#align dfinsupp.lift_add_hom_apply Dfinsupp.liftAddHom_apply +#align dfinsupp.lift_add_hom_symm_apply Dfinsupp.liftAddHom_symmApply -- Porting note: The elaborator is struggling with `liftAddHom`. Passing it `β` explicitly helps. -- This applies to roughly the remainder of the file. @@ -2165,6 +2177,7 @@ def mapRange.addMonoidHom (f : ∀ i, β₁ i →+ β₂ i) : (Π₀ i, β₁ i) map_zero' := mapRange_zero _ _ map_add' := mapRange_add _ (fun i => (f i).map_zero) fun i => (f i).map_add #align dfinsupp.map_range.add_monoid_hom Dfinsupp.mapRange.addMonoidHom +#align dfinsupp.map_range.add_monoid_hom_apply Dfinsupp.mapRange.addMonoidHom_apply @[simp] theorem mapRange.addMonoidHom_id : @@ -2197,6 +2210,7 @@ def mapRange.addEquiv (e : ∀ i, β₁ i ≃+ β₂ i) : (Π₀ i, β₁ i) ≃ · simp_rw [AddEquiv.self_comp_symm] simp } #align dfinsupp.map_range.add_equiv Dfinsupp.mapRange.addEquiv +#align dfinsupp.map_range.add_equiv_apply Dfinsupp.mapRange.addEquiv_apply @[simp] theorem mapRange.addEquiv_refl : diff --git a/Mathlib/Data/Fin/Basic.lean b/Mathlib/Data/Fin/Basic.lean index 5a9709a8c8e74..d0096d9e6a246 100644 --- a/Mathlib/Data/Fin/Basic.lean +++ b/Mathlib/Data/Fin/Basic.lean @@ -163,6 +163,8 @@ def equivSubtype : Fin n ≃ { i // i < n } where left_inv := fun ⟨_, _⟩ => rfl right_inv := fun ⟨_, _⟩ => rfl #align fin.equiv_subtype Fin.equivSubtype +#align fin.equiv_subtype_symm_apply Fin.equivSubtype_symm_apply +#align fin.equiv_subtype_apply Fin.equivSubtype_apply section coe @@ -330,6 +332,8 @@ theorem val_strictMono : StrictMono (val : Fin n → ℕ) := fun _ _ => id def orderIsoSubtype : Fin n ≃o { i // i < n } := equivSubtype.toOrderIso (by simp [Monotone]) (by simp [Monotone]) #align fin.order_iso_subtype Fin.orderIsoSubtype +#align fin.order_iso_subtype_symm_apply Fin.orderIsoSubtype_symmApply +#align fin.order_iso_subtype_apply Fin.orderIsoSubtype_apply /-- The inclusion map `Fin n → ℕ` is an embedding. -/ @[simps apply] @@ -493,6 +497,8 @@ theorem rev_lt_rev {i j : Fin n} : rev i < rev j ↔ j < i := def revOrderIso {n} : (Fin n)ᵒᵈ ≃o Fin n := ⟨OrderDual.ofDual.trans rev, rev_le_rev⟩ #align fin.rev_order_iso Fin.revOrderIso +#align fin.rev_order_iso_apply Fin.revOrderIso_apply +#align fin.rev_order_iso_to_equiv Fin.revOrderIso_toEquiv @[simp] theorem revOrderIso_symm_apply (i : Fin n) : revOrderIso.symm i = OrderDual.toDual (rev i) := diff --git a/Mathlib/Data/Fin/Fin2.lean b/Mathlib/Data/Fin/Fin2.lean index 6afd2cc7f3af3..444ce6d7e69e7 100644 --- a/Mathlib/Data/Fin/Fin2.lean +++ b/Mathlib/Data/Fin/Fin2.lean @@ -10,6 +10,7 @@ Authors: Mario Carneiro -/ import Std.Tactic.NoMatch import Mathlib.Init.Data.Nat.Notation +import Mathlib.Mathport.Rename /-! # Inductive type variant of `Fin` @@ -42,6 +43,7 @@ inductive Fin2 : ℕ → Type | fz {n} : Fin2 (succ n) /-- `n` as a member of `Fin (succ n)` -/ | fs {n} : Fin2 n → Fin2 (succ n) +#align fin2 Fin2 namespace Fin2 @@ -52,30 +54,36 @@ protected def cases' {n} {C : Fin2 (succ n) → Sort u} (H1 : C fz) (H2 : ∀ n, ∀ i : Fin2 (succ n), C i | fz => H1 | fs n => H2 n +#align fin2.cases' Fin2.cases' /-- Ex falso. The dependent eliminator for the empty `Fin2 0` type. -/ def elim0 {C : Fin2 0 → Sort u} : ∀ i : Fin2 0, C i := fun. +#align fin2.elim0 Fin2.elim0 /-- Converts a `Fin2` into a natural. -/ def toNat : ∀ {n}, Fin2 n → ℕ | _, @fz _ => 0 | _, @fs _ i => succ (toNat i) +#align fin2.to_nat Fin2.toNat /-- Converts a natural into a `Fin2` if it is in range -/ def optOfNat : ∀ {n}, ℕ → Option (Fin2 n) | 0, _ => none | succ _, 0 => some fz | succ m, succ k => fs <$> @optOfNat m k +#align fin2.opt_of_nat Fin2.optOfNat /-- `i + k : Fin2 (n + k)` when `i : Fin2 n` and `k : ℕ` -/ def add {n} (i : Fin2 n) : ∀ k, Fin2 (n + k) | 0 => i | succ k => fs (add i k) +#align fin2.add Fin2.add /-- `left k` is the embedding `Fin2 n → Fin2 (k + n)` -/ def left (k) : ∀ {n}, Fin2 n → Fin2 (k + n) | _, @fz _ => fz | _, @fs _ i => fs (left k i) +#align fin2.left Fin2.left /-- `insertPerm a` is a permutation of `Fin2 n` with the following properties: * `insertPerm a i = i+1` if `i < a` @@ -89,6 +97,7 @@ def insertPerm : ∀ {n}, Fin2 n → Fin2 n → Fin2 n match insertPerm i j with | fz => fz | fs k => fs (fs k) +#align fin2.insert_perm Fin2.insertPerm /-- `remapLeft f k : Fin2 (m + k) → Fin2 (n + k)` applies the function `f : Fin2 m → Fin2 n` to inputs less than `m`, and leaves the right part @@ -97,12 +106,14 @@ def remapLeft {m n} (f : Fin2 m → Fin2 n) : ∀ k, Fin2 (m + k) → Fin2 (n + | 0, i => f i | succ _, @fz _ => fz | succ _, @fs _ i => fs (remapLeft f _ i) +#align fin2.remap_left Fin2.remapLeft /-- This is a simple type class inference prover for proof obligations of the form `m < n` where `m n : ℕ`. -/ class IsLT (m n : ℕ) where /-- The unique field of `Fin2.IsLT`, a proof that `m < n`. -/ h : m < n +#align fin2.is_lt Fin2.IsLT instance IsLT.zero (n) : IsLT 0 (succ n) := ⟨succ_pos _⟩ @@ -116,6 +127,7 @@ def ofNat' : ∀ {n} (m) [IsLT m n], Fin2 n | 0, _, ⟨h⟩ => absurd h (Nat.not_lt_zero _) | succ _, 0, ⟨_⟩ => fz | succ n, succ m, ⟨h⟩ => fs (@ofNat' n m ⟨lt_of_succ_lt_succ h⟩) +#align fin2.of_nat' Fin2.ofNat' @[inherit_doc] local prefix:arg "&" => ofNat' diff --git a/Mathlib/Data/Finite/Defs.lean b/Mathlib/Data/Finite/Defs.lean index 8a6536938211a..257cb418b29df 100644 --- a/Mathlib/Data/Finite/Defs.lean +++ b/Mathlib/Data/Finite/Defs.lean @@ -138,3 +138,5 @@ protected theorem Infinite.false [Finite α] (_ : Infinite α) : False := #align infinite.false Infinite.false alias not_infinite_iff_finite ↔ Finite.of_not_infinite Finite.not_infinite +#align finite.of_not_infinite Finite.of_not_infinite +#align finite.not_infinite Finite.not_infinite diff --git a/Mathlib/Data/Finset/NatAntidiagonal.lean b/Mathlib/Data/Finset/NatAntidiagonal.lean index 3b39ba4424795..d32d54892c44a 100644 --- a/Mathlib/Data/Finset/NatAntidiagonal.lean +++ b/Mathlib/Data/Finset/NatAntidiagonal.lean @@ -143,6 +143,9 @@ def sigmaAntidiagonalEquivProd : (Σn : ℕ, antidiagonal n) ≃ ℕ × ℕ wher exact Sigma.subtype_ext h rfl right_inv x := rfl #align finset.nat.sigma_antidiagonal_equiv_prod Finset.Nat.sigmaAntidiagonalEquivProd +#align finset.nat.sigma_antidiagonal_equiv_prod_symm_apply_fst Finset.Nat.sigmaAntidiagonalEquivProd_symm_apply_fst +#align finset.nat.sigma_antidiagonal_equiv_prod_symm_apply_snd_coe Finset.Nat.sigmaAntidiagonalEquivProd_symm_apply_snd_coe +#align finset.nat.sigma_antidiagonal_equiv_prod_apply Finset.Nat.sigmaAntidiagonalEquivProd_apply end EquivProd diff --git a/Mathlib/Data/Finsupp/Defs.lean b/Mathlib/Data/Finsupp/Defs.lean index 0b7e5c26a0d49..319db63e8a419 100644 --- a/Mathlib/Data/Finsupp/Defs.lean +++ b/Mathlib/Data/Finsupp/Defs.lean @@ -264,6 +264,9 @@ If `α` has a unique term, the type of finitely supported functions `α →₀ noncomputable def _root_.Equiv.finsuppUnique {ι : Type _} [Unique ι] : (ι →₀ M) ≃ M := Finsupp.equivFunOnFinite.trans (Equiv.funUnique ι M) #align equiv.finsupp_unique Equiv.finsuppUnique +#align equiv.finsupp_unique_symm_apply_support_val Equiv.finsuppUnique_symm_apply_support_val +#align equiv.finsupp_unique_symm_apply_to_fun Equiv.finsuppUnique_symm_apply_toFun +#align equiv.finsupp_unique_apply Equiv.finsuppUnique_apply @[ext] theorem unique_ext [Unique α] {f g : α →₀ M} (h : f default = g default) : f = g := @@ -1021,6 +1024,7 @@ def applyAddHom (a : α) : (α →₀ M) →+ M where map_zero' := zero_apply map_add' _ _ := add_apply _ _ _ #align finsupp.apply_add_hom Finsupp.applyAddHom +#align finsupp.apply_add_hom_apply Finsupp.applyAddHom_apply /-- Coercion from a `Finsupp` to a function type is an `AddMonoidHom`. -/ @[simps] @@ -1030,6 +1034,7 @@ noncomputable def coeFnAddHom : (α →₀ M) →+ α → M map_zero' := coe_zero map_add' := coe_add #align finsupp.coe_fn_add_hom Finsupp.coeFnAddHom +#align finsupp.coe_fn_add_hom_apply Finsupp.coeFnAddHom_apply theorem update_eq_single_add_erase (f : α →₀ M) (a : α) (b : M) : f.update a b = single a b + f.erase a := by diff --git a/Mathlib/Data/Fintype/Basic.lean b/Mathlib/Data/Fintype/Basic.lean index 8047553a2332d..57144f9c2e829 100644 --- a/Mathlib/Data/Fintype/Basic.lean +++ b/Mathlib/Data/Fintype/Basic.lean @@ -1027,6 +1027,7 @@ def unitsEquivProdSubtype [Monoid α] : αˣ ≃ { p : α × α // p.1 * p.2 = 1 left_inv _ := Units.ext rfl right_inv _ := Subtype.ext <| Prod.ext rfl rfl #align units_equiv_prod_subtype unitsEquivProdSubtype +#align units_equiv_prod_subtype_apply_coe unitsEquivProdSubtype_apply_coe /-- In a `GroupWithZero` `α`, the unit group `αˣ` is equivalent to the subtype of nonzero elements. -/ @@ -1035,6 +1036,8 @@ def unitsEquivNeZero [GroupWithZero α] : αˣ ≃ { a : α // a ≠ 0 } := ⟨fun a => ⟨a, a.ne_zero⟩, fun a => Units.mk0 _ a.prop, fun _ => Units.ext rfl, fun _ => Subtype.ext rfl⟩ #align units_equiv_ne_zero unitsEquivNeZero +#align units_equiv_ne_zero_apply_coe unitsEquivNeZero_apply_coe +#align units_equiv_ne_zero_symm_apply unitsEquivNeZero_symm_apply end diff --git a/Mathlib/Data/Fintype/Card.lean b/Mathlib/Data/Fintype/Card.lean index c0ff1e2353b5f..447d5fdd88442 100644 --- a/Mathlib/Data/Fintype/Card.lean +++ b/Mathlib/Data/Fintype/Card.lean @@ -706,6 +706,8 @@ def ofLeftInverseOfCardLe (hβα : card β ≤ card α) (f : α → β) (g : β left_inv := h right_inv := h.rightInverse_of_card_le hβα #align equiv.of_left_inverse_of_card_le Equiv.ofLeftInverseOfCardLe +#align equiv.of_left_inverse_of_card_le_symm_apply Equiv.ofLeftInverseOfCardLe_symm_apply +#align equiv.of_left_inverse_of_card_le_apply Equiv.ofLeftInverseOfCardLe_apply /-- Construct an equivalence from functions that are inverse to each other. -/ @[simps] @@ -716,6 +718,8 @@ def ofRightInverseOfCardLe (hαβ : card α ≤ card β) (f : α → β) (g : β left_inv := h.leftInverse_of_card_le hαβ right_inv := h #align equiv.of_right_inverse_of_card_le Equiv.ofRightInverseOfCardLe +#align equiv.of_right_inverse_of_card_le_symm_apply Equiv.ofRightInverseOfCardLe_symm_apply +#align equiv.of_right_inverse_of_card_le_apply Equiv.ofRightInverseOfCardLe_apply end Equiv diff --git a/Mathlib/Data/FunLike/Basic.lean b/Mathlib/Data/FunLike/Basic.lean index 34e416e8cf39c..2d637ef797dc4 100644 --- a/Mathlib/Data/FunLike/Basic.lean +++ b/Mathlib/Data/FunLike/Basic.lean @@ -133,6 +133,7 @@ class FunLike (F : Sort _) (α : outParam (Sort _)) (β : outParam <| α → Sor coe : F → ∀ a : α, β a /-- The coercion to functions must be injective. -/ coe_injective' : Function.Injective coe +#align fun_like FunLike section Dependent @@ -152,38 +153,49 @@ instance (priority := 100) : CoeFun F fun _ ↦ ∀ a : α, β a where coe := Fu -- @[simp] -- porting note: this loops in lean 4 theorem coe_eq_coe_fn : (FunLike.coe (F := F)) = (fun f => ↑f) := rfl +#align fun_like.coe_eq_coe_fn FunLike.coe_eq_coe_fn theorem coe_injective : Function.Injective (fun f : F ↦ (f : ∀ a : α, β a)) := FunLike.coe_injective' +#align fun_like.coe_injective FunLike.coe_injective @[simp] theorem coe_fn_eq {f g : F} : (f : ∀ a : α, β a) = (g : ∀ a : α, β a) ↔ f = g := ⟨fun h ↦ FunLike.coe_injective' h, fun h ↦ by cases h; rfl⟩ +#align fun_like.coe_fn_eq FunLike.coe_fn_eq theorem ext' {f g : F} (h : (f : ∀ a : α, β a) = (g : ∀ a : α, β a)) : f = g := FunLike.coe_injective' h +#align fun_like.ext' FunLike.ext' theorem ext'_iff {f g : F} : f = g ↔ (f : ∀ a : α, β a) = (g : ∀ a : α, β a) := coe_fn_eq.symm +#align fun_like.ext'_iff FunLike.ext'_iff theorem ext (f g : F) (h : ∀ x : α, f x = g x) : f = g := FunLike.coe_injective' (funext h) +#align fun_like.ext FunLike.ext theorem ext_iff {f g : F} : f = g ↔ ∀ x, f x = g x := coe_fn_eq.symm.trans Function.funext_iff +#align fun_like.ext_iff FunLike.ext_iff protected theorem congr_fun {f g : F} (h₁ : f = g) (x : α) : f x = g x := congr_fun (congr_arg _ h₁) x +#align fun_like.congr_fun FunLike.congr_fun theorem ne_iff {f g : F} : f ≠ g ↔ ∃ a, f a ≠ g a := ext_iff.not.trans not_forall +#align fun_like.ne_iff FunLike.ne_iff theorem exists_ne {f g : F} (h : f ≠ g) : ∃ x, f x ≠ g x := ne_iff.mp h +#align fun_like.exists_ne FunLike.exists_ne /-- This is not an instance to avoid slowing down every single `Subsingleton` typeclass search.-/ lemma subsingleton_cod [∀ a, Subsingleton (β a)] : Subsingleton F := ⟨fun _ _ ↦ coe_injective $ Subsingleton.elim _ _⟩ +#align fun_like.subsingleton_cod FunLike.subsingleton_cod end FunLike @@ -199,9 +211,11 @@ namespace FunLike protected theorem congr {f g : F} {x y : α} (h₁ : f = g) (h₂ : x = y) : f x = g y := congr (congr_arg _ h₁) h₂ +#align fun_like.congr FunLike.congr protected theorem congr_arg (f : F) {x y : α} (h₂ : x = y) : f x = f y := congr_arg _ h₂ +#align fun_like.congr_arg FunLike.congr_arg end FunLike diff --git a/Mathlib/Data/FunLike/Embedding.lean b/Mathlib/Data/FunLike/Embedding.lean index 38c9bf707f269..7ecef7620e599 100644 --- a/Mathlib/Data/FunLike/Embedding.lean +++ b/Mathlib/Data/FunLike/Embedding.lean @@ -137,6 +137,7 @@ injective coercion to injective functions `α ↪ β`. class EmbeddingLike (F : Sort _) (α β : outParam (Sort _)) extends FunLike F α fun _ ↦ β where /-- The coercion to functions must produce injective functions. -/ injective' : ∀ f : F, @Function.Injective α β (coe f) +#align embedding_like EmbeddingLike namespace EmbeddingLike @@ -144,14 +145,17 @@ variable {F α β γ : Sort _} [i : EmbeddingLike F α β] protected theorem injective (f : F) : Function.Injective f := injective' f +#align embedding_like.injective EmbeddingLike.injective @[simp] theorem apply_eq_iff_eq (f : F) {x y : α} : f x = f y ↔ x = y := (EmbeddingLike.injective f).eq_iff +#align embedding_like.apply_eq_iff_eq EmbeddingLike.apply_eq_iff_eq @[simp] theorem comp_injective {F : Sort _} [EmbeddingLike F β γ] (f : α → β) (e : F) : Function.Injective (e ∘ f) ↔ Function.Injective f := (EmbeddingLike.injective e).of_comp_iff f +#align embedding_like.comp_injective EmbeddingLike.comp_injective end EmbeddingLike diff --git a/Mathlib/Data/FunLike/Equiv.lean b/Mathlib/Data/FunLike/Equiv.lean index fe54f8b565b18..18e8ca399d9da 100644 --- a/Mathlib/Data/FunLike/Equiv.lean +++ b/Mathlib/Data/FunLike/Equiv.lean @@ -146,6 +146,7 @@ class EquivLike (E : Sort _) (α β : outParam (Sort _)) where coe_injective' : ∀ e g, coe e = coe g → inv e = inv g → e = g -- This is mathematically equivalent to either of the coercions to functions being injective, but -- the `inv` hypothesis makes this easier to prove with `congr'` +#align equiv_like EquivLike namespace EquivLike @@ -153,6 +154,7 @@ variable {E F α β γ : Sort _} [iE : EquivLike E α β] [iF : EquivLike F β theorem inv_injective : Function.Injective (EquivLike.inv : E → β → α) := fun e g h ↦ coe_injective' e g ((right_inv e).eq_rightInverse (h.symm ▸ left_inv g)) h +#align equiv_like.inv_injective EquivLike.inv_injective instance (priority := 100) toEmbeddingLike : EmbeddingLike E α β where coe := (coe : E → α → β) @@ -162,27 +164,34 @@ instance (priority := 100) toEmbeddingLike : EmbeddingLike E α β where protected theorem injective (e : E) : Function.Injective e := EmbeddingLike.injective e +#align equiv_like.injective EquivLike.injective protected theorem surjective (e : E) : Function.Surjective e := (right_inv e).surjective +#align equiv_like.surjective EquivLike.surjective protected theorem bijective (e : E) : Function.Bijective (e : α → β) := ⟨EquivLike.injective e, EquivLike.surjective e⟩ +#align equiv_like.bijective EquivLike.bijective theorem apply_eq_iff_eq (f : E) {x y : α} : f x = f y ↔ x = y := EmbeddingLike.apply_eq_iff_eq f +#align equiv_like.apply_eq_iff_eq EquivLike.apply_eq_iff_eq @[simp] theorem injective_comp (e : E) (f : β → γ) : Function.Injective (f ∘ e) ↔ Function.Injective f := Function.Injective.of_comp_iff' f (EquivLike.bijective e) +#align equiv_like.injective_comp EquivLike.injective_comp @[simp] theorem surjective_comp (e : E) (f : β → γ) : Function.Surjective (f ∘ e) ↔ Function.Surjective f := (EquivLike.surjective e).of_comp_iff f +#align equiv_like.surjective_comp EquivLike.surjective_comp @[simp] theorem bijective_comp (e : E) (f : β → γ) : Function.Bijective (f ∘ e) ↔ Function.Bijective f := (EquivLike.bijective e).of_comp_iff f +#align equiv_like.bijective_comp EquivLike.bijective_comp /-- This lemma is only supposed to be used in the generic context, when working with instances of classes extending `EquivLike`. @@ -193,6 +202,7 @@ TODO: define a generic form of `Equiv.symm`. -/ @[simp] theorem inv_apply_apply (e : E) (a : α) : EquivLike.inv e (e a) = a := left_inv _ _ +#align equiv_like.inv_apply_apply EquivLike.inv_apply_apply /-- This lemma is only supposed to be used in the generic context, when working with instances of classes extending `EquivLike`. @@ -203,20 +213,25 @@ TODO: define a generic form of `Equiv.symm`. -/ @[simp] theorem apply_inv_apply (e : E) (b : β) : e (EquivLike.inv e b) = b := right_inv _ _ +#align equiv_like.apply_inv_apply EquivLike.apply_inv_apply theorem comp_injective (f : α → β) (e : F) : Function.Injective (e ∘ f) ↔ Function.Injective f := EmbeddingLike.comp_injective f e +#align equiv_like.comp_injective EquivLike.comp_injective @[simp] theorem comp_surjective (f : α → β) (e : F) : Function.Surjective (e ∘ f) ↔ Function.Surjective f := Function.Surjective.of_comp_iff' (EquivLike.bijective e) f +#align equiv_like.comp_surjective EquivLike.comp_surjective @[simp] theorem comp_bijective (f : α → β) (e : F) : Function.Bijective (e ∘ f) ↔ Function.Bijective f := (EquivLike.bijective e).of_comp_iff' f +#align equiv_like.comp_bijective EquivLike.comp_bijective /-- This is not an instance to avoid slowing down every single `Subsingleton` typeclass search.-/ lemma subsingleton_dom [Subsingleton β] : Subsingleton F := ⟨fun f g ↦ FunLike.ext f g $ fun _ ↦ (right_inv f).injective $ Subsingleton.elim _ _⟩ +#align equiv_like.subsingleton_dom EquivLike.subsingleton_dom end EquivLike diff --git a/Mathlib/Data/Int/AbsoluteValue.lean b/Mathlib/Data/Int/AbsoluteValue.lean index 82419cad92085..7a10dfbadf965 100644 --- a/Mathlib/Data/Int/AbsoluteValue.lean +++ b/Mathlib/Data/Int/AbsoluteValue.lean @@ -51,3 +51,4 @@ def Int.natAbsHom : ℤ →*₀ ℕ where map_one' := Int.natAbs_one map_zero' := Int.natAbs_zero #align int.nat_abs_hom Int.natAbsHom +#align int.nat_abs_hom_apply Int.natAbsHom_apply diff --git a/Mathlib/Data/Int/Basic.lean b/Mathlib/Data/Int/Basic.lean index 487c9b1df1b95..1e3e4f8a3599a 100644 --- a/Mathlib/Data/Int/Basic.lean +++ b/Mathlib/Data/Int/Basic.lean @@ -86,6 +86,7 @@ theorem _root_.zpow_coe_nat [DivInvMonoid G] (a : G) (n : ℕ) : a ^ (Nat.cast n @[simp] theorem _root_.coe_nat_zsmul [SubNegMonoid G] (a : G) (n : ℕ) : (n : ℤ) • a = n • a := ofNat_zsmul .. attribute [to_additive coe_nat_zsmul] zpow_coe_nat +#align coe_nat_zsmul coe_nat_zsmul /-! ### Extra instances to short-circuit type class resolution @@ -124,6 +125,7 @@ theorem coe_nat_strictMono : StrictMono (· : ℕ → ℤ) := fun _ _ ↦ Int.of #align int.coe_nat_strict_mono Int.coe_nat_strictMono theorem coe_nat_nonneg (n : ℕ) : 0 ≤ (n : ℤ) := ofNat_le.2 (Nat.zero_le _) +#align int.coe_nat_nonneg Int.coe_nat_nonneg #align int.neg_of_nat_ne_zero Int.negSucc_ne_zero #align int.zero_ne_neg_of_nat Int.zero_ne_negSucc @@ -132,33 +134,46 @@ theorem coe_nat_nonneg (n : ℕ) : 0 ≤ (n : ℤ) := ofNat_le.2 (Nat.zero_le _) /-- Immediate successor of an integer: `succ n = n + 1` -/ def succ (a : ℤ) := a + 1 +#align int.succ Int.succ /-- Immediate predecessor of an integer: `pred n = n - 1` -/ def pred (a : ℤ) := a - 1 +#align int.pred Int.pred theorem nat_succ_eq_int_succ (n : ℕ) : (Nat.succ n : ℤ) = Int.succ n := rfl +#align int.nat_succ_eq_int_succ Int.nat_succ_eq_int_succ theorem pred_succ (a : ℤ) : pred (succ a) = a := add_sub_cancel _ _ +#align int.pred_succ Int.pred_succ theorem succ_pred (a : ℤ) : succ (pred a) = a := sub_add_cancel _ _ +#align int.succ_pred Int.succ_pred theorem neg_succ (a : ℤ) : -succ a = pred (-a) := neg_add _ _ +#align int.neg_succ Int.neg_succ theorem succ_neg_succ (a : ℤ) : succ (-succ a) = -a := by rw [neg_succ, succ_pred] +#align int.succ_neg_succ Int.succ_neg_succ theorem neg_pred (a : ℤ) : -pred a = succ (-a) := by rw [eq_neg_of_eq_neg (neg_succ (-a)).symm, neg_neg] +#align int.neg_pred Int.neg_pred theorem pred_neg_pred (a : ℤ) : pred (-pred a) = -a := by rw [neg_pred, pred_succ] +#align int.pred_neg_pred Int.pred_neg_pred theorem pred_nat_succ (n : ℕ) : pred (Nat.succ n) = n := pred_succ n +#align int.pred_nat_succ Int.pred_nat_succ theorem neg_nat_succ (n : ℕ) : -(Nat.succ n : ℤ) = pred (-n) := neg_succ n +#align int.neg_nat_succ Int.neg_nat_succ theorem succ_neg_nat_succ (n : ℕ) : succ (-Nat.succ n) = -n := succ_neg_succ n +#align int.succ_neg_nat_succ Int.succ_neg_nat_succ @[norm_cast] theorem coe_pred_of_pos {n : ℕ} (h : 0 < n) : ((n - 1 : ℕ) : ℤ) = (n : ℤ) - 1 := by cases n; cases h; simp +#align int.coe_pred_of_pos Int.coe_pred_of_pos @[elab_as_elim] protected theorem induction_on {p : ℤ → Prop} (i : ℤ) (hz : p 0) (hp : ∀ i : ℕ, p i → p (i + 1)) (hn : ∀ i : ℕ, p (-i) → p (-i - 1)) : p i := by @@ -172,6 +187,7 @@ theorem succ_neg_nat_succ (n : ℕ) : succ (-Nat.succ n) = -n := succ_neg_succ n intro n; induction n with | zero => simp [hz, Nat.cast_zero] | succ n ih => convert hn _ ih using 1; simp [sub_eq_neg_add] +#align int.induction_on Int.induction_on /-! ### nat abs -/ diff --git a/Mathlib/Data/Int/Cast/Basic.lean b/Mathlib/Data/Int/Cast/Basic.lean index 76de2157b1cc4..27ae8b2daee83 100644 --- a/Mathlib/Data/Int/Cast/Basic.lean +++ b/Mathlib/Data/Int/Cast/Basic.lean @@ -68,6 +68,7 @@ theorem cast_ofNat (n : ℕ) : ((n : ℤ) : R) = n := AddGroupWithOne.intCast_ofNat _ #align int.cast_coe_nat Int.cast_ofNatₓ -- expected `n` to be implicit, and `HasLiftT` +#align int.cast_of_nat Int.cast_ofNatₓ @[simp, norm_cast] theorem cast_one : ((1 : ℤ) : R) = 1 := by diff --git a/Mathlib/Data/Int/GCD.lean b/Mathlib/Data/Int/GCD.lean index 7a9a28df9e301..cd27eae7d5703 100644 --- a/Mathlib/Data/Int/GCD.lean +++ b/Mathlib/Data/Int/GCD.lean @@ -99,6 +99,7 @@ theorem gcdA_zero_right {s : ℕ} (h : s ≠ 0) : gcdA s 0 = 1 := by -- Porting note: `simp [xgcdAux_succ]` crashes Lean here rw [xgcdAux_succ] rfl +#align nat.gcd_a_zero_right Nat.gcdA_zero_right @[simp] theorem gcdB_zero_right {s : ℕ} (h : s ≠ 0) : gcdB s 0 = 0 := by diff --git a/Mathlib/Data/Int/Order/Basic.lean b/Mathlib/Data/Int/Order/Basic.lean index ffc738f0a7c0f..67f5a7c5fe09c 100644 --- a/Mathlib/Data/Int/Order/Basic.lean +++ b/Mathlib/Data/Int/Order/Basic.lean @@ -153,6 +153,7 @@ where refine _root_.cast (by rw [add_sub_assoc]; rfl) (Hp _ (Int.le_of_lt ?_) (neg n)) conv => rhs; apply (add_zero b).symm rw [Int.add_lt_add_iff_left]; apply negSucc_lt_zero +#align int.induction_on' Int.inductionOn' /-- See `Int.inductionOn'` for an induction in both directions. -/ protected theorem le_induction {P : ℤ → Prop} {m : ℤ} (h0 : P m) diff --git a/Mathlib/Data/Int/Order/Units.lean b/Mathlib/Data/Int/Order/Units.lean index 602ec7e7c5804..0d23611c29355 100644 --- a/Mathlib/Data/Int/Order/Units.lean +++ b/Mathlib/Data/Int/Order/Units.lean @@ -32,6 +32,7 @@ theorem units_sq (u : ℤˣ) : u ^ 2 = 1 := by #align int.units_sq Int.units_sq alias units_sq ← units_pow_two +#align int.units_pow_two Int.units_pow_two @[simp] theorem units_mul_self (u : ℤˣ) : u * u = 1 := by rw [← sq, units_sq] diff --git a/Mathlib/Data/List/Chain.lean b/Mathlib/Data/List/Chain.lean index 7daabbe5057a9..0c69f105da9a8 100644 --- a/Mathlib/Data/List/Chain.lean +++ b/Mathlib/Data/List/Chain.lean @@ -31,6 +31,7 @@ namespace List variable {α : Type u} {β : Type v} {R r : α → α → Prop} {l l₁ l₂ : List α} {a b : α} mk_iff_of_inductive_prop List.Chain List.chain_iff +#align list.chain_iff List.chain_iff --Porting note: attribute in Lean3, but not in Lean4 Std so added here instead attribute [simp] Chain.nil diff --git a/Mathlib/Data/List/Forall2.lean b/Mathlib/Data/List/Forall2.lean index 885157385b389..2c0bd4b1db7be 100644 --- a/Mathlib/Data/List/Forall2.lean +++ b/Mathlib/Data/List/Forall2.lean @@ -28,6 +28,7 @@ variable {α β γ δ : Type _} {R S : α → β → Prop} {P : γ → δ → Pr open Relator mk_iff_of_inductive_prop List.Forall₂ List.forall₂_iff +#align list.forall₂_iff List.forall₂_iff #align list.forall₂.nil List.Forall₂.nil #align list.forall₂.cons List.Forall₂.cons diff --git a/Mathlib/Data/List/Infix.lean b/Mathlib/Data/List/Infix.lean index f20df6aedb857..ab6db5d730dad 100644 --- a/Mathlib/Data/List/Infix.lean +++ b/Mathlib/Data/List/Infix.lean @@ -184,10 +184,13 @@ theorem reverse_infix : reverse l₁ <:+: reverse l₂ ↔ l₁ <:+: l₂ := #align list.reverse_infix List.reverse_infix alias reverse_prefix ↔ _ isSuffix.reverse +#align list.is_suffix.reverse List.isSuffix.reverse alias reverse_suffix ↔ _ isPrefix.reverse +#align list.is_prefix.reverse List.isPrefix.reverse alias reverse_infix ↔ _ isInfix.reverse +#align list.is_infix.reverse List.isInfix.reverse theorem isInfix.length_le (h : l₁ <:+: l₂) : l₁.length ≤ l₂.length := h.sublist.length_le @@ -221,8 +224,10 @@ theorem suffix_nil_iff : l <:+ [] ↔ l = [] := #align list.suffix_nil_iff List.suffix_nil_iff alias prefix_nil_iff ↔ eq_nil_of_prefix_nil _ +#align list.eq_nil_of_prefix_nil List.eq_nil_of_prefix_nil alias suffix_nil_iff ↔ eq_nil_of_suffix_nil _ +#align list.eq_nil_of_suffix_nil List.eq_nil_of_suffix_nil theorem infix_iff_prefix_suffix (l₁ l₂ : List α) : l₁ <:+: l₂ ↔ ∃ t, l₁ <+: t ∧ t <:+ l₂ := ⟨fun ⟨s, t, e⟩ => ⟨l₁ ++ t, ⟨_, rfl⟩, by rw [← e, append_assoc]; exact ⟨_, rfl⟩⟩, diff --git a/Mathlib/Data/List/Nodup.lean b/Mathlib/Data/List/Nodup.lean index f938bf3a2d509..5048887453515 100644 --- a/Mathlib/Data/List/Nodup.lean +++ b/Mathlib/Data/List/Nodup.lean @@ -280,6 +280,8 @@ theorem nodup_attach {l : List α} : Nodup (attach l) ↔ Nodup l := #align list.nodup_attach List.nodup_attach alias nodup_attach ↔ Nodup.of_attach Nodup.attach +#align list.nodup.attach List.Nodup.attach +#align list.nodup.of_attach List.Nodup.of_attach --Porting note: commented out --attribute [protected] nodup.attach diff --git a/Mathlib/Data/List/OfFn.lean b/Mathlib/Data/List/OfFn.lean index 4933e7eb876f4..cc31cb5a34a97 100644 --- a/Mathlib/Data/List/OfFn.lean +++ b/Mathlib/Data/List/OfFn.lean @@ -252,6 +252,9 @@ def equivSigmaTuple : List α ≃ Σn, Fin n → α right_inv := fun ⟨_, f⟩ => Fin.sigma_eq_of_eq_comp_cast (length_ofFn _) <| funext fun i => get_ofFn f i #align list.equiv_sigma_tuple List.equivSigmaTuple +#align list.equiv_sigma_tuple_symm_apply List.equivSigmaTuple_symm_apply +#align list.equiv_sigma_tuple_apply_fst List.equivSigmaTuple_apply_fst +#align list.equiv_sigma_tuple_apply_snd List.equivSigmaTuple_apply_snd /-- A recursor for lists that expands a list into a function mapping to its elements. diff --git a/Mathlib/Data/List/Pairwise.lean b/Mathlib/Data/List/Pairwise.lean index 1e14429ced1aa..27e386d263c3f 100644 --- a/Mathlib/Data/List/Pairwise.lean +++ b/Mathlib/Data/List/Pairwise.lean @@ -38,6 +38,7 @@ namespace List variable {α β : Type _} {R S T : α → α → Prop} {a : α} {l : List α} mk_iff_of_inductive_prop List.Pairwise List.pairwise_iff +#align list.pairwise_iff List.pairwise_iff /-! ### Pairwise -/ @@ -422,6 +423,7 @@ theorem pwFilter_eq_self {l : List α} : pwFilter R l = l ↔ Pairwise R l := #align list.pw_filter_eq_self List.pwFilter_eq_self alias pwFilter_eq_self ↔ _ Pairwise.pwFilter +#align list.pairwise.pw_filter List.Pairwise.pwFilter -- Porting note: commented out -- attribute [protected] List.Pairwise.pwFilter diff --git a/Mathlib/Data/List/Perm.lean b/Mathlib/Data/List/Perm.lean index 6b364ed9d3757..26828baeebd0a 100644 --- a/Mathlib/Data/List/Perm.lean +++ b/Mathlib/Data/List/Perm.lean @@ -681,6 +681,8 @@ theorem subperm_cons (a : α) {l₁ l₂ : List α} : a :: l₁ <+~ a :: l₂ #align list.subperm_cons List.subperm_cons alias subperm_cons ↔ subperm.of_cons subperm.cons +#align list.subperm.of_cons List.subperm.of_cons +#align list.subperm.cons List.subperm.cons --Porting note: commented out --attribute [protected] subperm.cons diff --git a/Mathlib/Data/List/Sublists.lean b/Mathlib/Data/List/Sublists.lean index 585d9cdb63c21..8fb00648feab8 100644 --- a/Mathlib/Data/List/Sublists.lean +++ b/Mathlib/Data/List/Sublists.lean @@ -49,6 +49,7 @@ theorem sublists'_singleton (a : α) : sublists' [a] = [[], [a]] := /-- Auxilary helper definiiton for `sublists'` -/ def sublists'Aux (a : α) (r₁ r₂ : List (List α)) : List (List α) := r₁.foldl (init := r₂) fun r l => r ++ [a :: l] +#align list.sublists'_aux List.sublists'Aux theorem sublists'Aux_eq_array_foldl (a : α) : ∀ (r₁ r₂ : List (List α)), sublists'Aux a r₁ r₂ = ((r₁.toArray).foldl (init := r₂.toArray) @@ -117,6 +118,7 @@ theorem sublists_singleton (a : α) : sublists [a] = [[], [a]] := /-- Auxilary helper function for `sublists` -/ def sublistsAux (a : α) (r : List (List α)) : List (List α) := r.foldl (init := []) fun r l => r ++ [l, a :: l] +#align list.sublists_aux List.sublistsAux theorem sublistsAux_eq_array_foldl : sublistsAux = fun (a : α) (r : List (List α)) => diff --git a/Mathlib/Data/List/Zip.lean b/Mathlib/Data/List/Zip.lean index 7baabf67538fe..7dcc3f7a92803 100644 --- a/Mathlib/Data/List/Zip.lean +++ b/Mathlib/Data/List/Zip.lean @@ -412,6 +412,7 @@ theorem nthLe_zipWith {f : α → β → γ} {l : List α} {l' : List β} {i : (zipWith f l l').nthLe i h = f (l.nthLe i (lt_length_left_of_zipWith h)) (l'.nthLe i (lt_length_right_of_zipWith h)) := get_zipWith (i := ⟨i, h⟩) +#align list.nth_le_zip_with List.nthLe_zipWith @[simp] theorem get_zip {l : List α} {l' : List β} {i : Fin (zip l l').length} : diff --git a/Mathlib/Data/Multiset/Basic.lean b/Mathlib/Data/Multiset/Basic.lean index 1fb0a72f8283b..ba8478dd60a66 100644 --- a/Mathlib/Data/Multiset/Basic.lean +++ b/Mathlib/Data/Multiset/Basic.lean @@ -876,6 +876,7 @@ instance is_wellFounded_lt : WellFoundedLT (Multiset α) := /-- `replicate n a` is the multiset containing only `a` with multiplicity `n`. -/ def replicate (n : ℕ) (a : α) : Multiset α := List.replicate n a +#align multiset.replicate Multiset.replicate theorem coe_replicate (n : ℕ) (a : α) : (List.replicate n a : Multiset α) = replicate n a := rfl #align multiset.coe_replicate Multiset.coe_replicate @@ -897,6 +898,7 @@ def replicateAddMonoidHom (a : α) : ℕ →+ Multiset α where map_zero' := replicate_zero a map_add' := fun _ _ => replicate_add _ _ a #align multiset.replicate_add_monoid_hom Multiset.replicateAddMonoidHom +#align multiset.replicate_add_monoid_hom_apply Multiset.replicateAddMonoidHom_apply -- @[simp] -- Porting note: simp can prove this theorem replicate_one (a : α) : replicate 1 a = {a} := rfl @@ -2621,6 +2623,7 @@ image under `f`. -/ def mapEmbedding (f : α ↪ β) : Multiset α ↪o Multiset β := OrderEmbedding.ofMapLeIff (map f) fun _ _ => map_le_map_iff f.inj' #align multiset.map_embedding Multiset.mapEmbedding +#align multiset.map_embedding_apply Multiset.mapEmbedding_apply end Embedding @@ -2657,6 +2660,7 @@ inductive Rel (r : α → β → Prop) : Multiset α → Multiset β → Prop | zero : Rel r 0 0 | cons {a b as bs} : r a b → Rel r as bs → Rel r (a ::ₘ as) (b ::ₘ bs) #align multiset.rel Multiset.Rel +#align multiset.rel_iff Multiset.Rel_iff variable {δ : Type _} {r : α → β → Prop} {p : γ → δ → Prop} diff --git a/Mathlib/Data/Nat/Cast/Basic.lean b/Mathlib/Data/Nat/Cast/Basic.lean index 53a8794a5df96..9a61151d8b6e5 100644 --- a/Mathlib/Data/Nat/Cast/Basic.lean +++ b/Mathlib/Data/Nat/Cast/Basic.lean @@ -42,6 +42,7 @@ def castAddMonoidHom (α : Type _) [AddMonoidWithOne α] : toFun := Nat.cast map_add' := cast_add map_zero' := cast_zero +#align nat.cast_add_monoid_hom Nat.castAddMonoidHom @[simp] theorem coe_castAddMonoidHom [AddMonoidWithOne α] : (castAddMonoidHom α : ℕ → α) = Nat.cast := @@ -51,10 +52,12 @@ theorem coe_castAddMonoidHom [AddMonoidWithOne α] : (castAddMonoidHom α : ℕ @[simp, norm_cast] theorem cast_mul [NonAssocSemiring α] (m n : ℕ) : ((m * n : ℕ) : α) = m * n := by induction n <;> simp [mul_succ, mul_add, *] +#align nat.cast_mul Nat.cast_mul /-- `Nat.cast : ℕ → α` as a `RingHom` -/ def castRingHom (α : Type _) [NonAssocSemiring α] : ℕ →+* α := { castAddMonoidHom α with toFun := Nat.cast, map_one' := cast_one, map_mul' := cast_mul } +#align nat.cast_ring_hom Nat.castRingHom @[simp] theorem coe_castRingHom [NonAssocSemiring α] : (castRingHom α : ℕ → α) = Nat.cast := @@ -65,12 +68,15 @@ theorem cast_commute [NonAssocSemiring α] (n : ℕ) (x : α) : Commute (n : α) induction n with | zero => rw [Nat.cast_zero]; exact Commute.zero_left x | succ n ihn => rw [Nat.cast_succ]; exact ihn.add_left (Commute.one_left x) +#align nat.cast_commute Nat.cast_commute theorem cast_comm [NonAssocSemiring α] (n : ℕ) (x : α) : (n : α) * x = x * n := (cast_commute n x).eq +#align nat.cast_comm Nat.cast_comm theorem commute_cast [NonAssocSemiring α] (x : α) (n : ℕ) : Commute x n := (n.cast_commute x).symm +#align nat.commute_cast Nat.commute_cast section OrderedSemiring @@ -81,10 +87,12 @@ variable [OrderedSemiring α] theorem mono_cast : Monotone (Nat.cast : ℕ → α) := monotone_nat_of_le_succ fun n ↦ by rw [Nat.cast_succ]; exact le_add_of_nonneg_right zero_le_one +#align nat.mono_cast Nat.mono_cast @[simp] theorem cast_nonneg (n : ℕ) : 0 ≤ (n : α) := @Nat.cast_zero α _ ▸ mono_cast (Nat.zero_le n) +#align nat.cast_nonneg Nat.cast_nonneg section Nontrivial @@ -92,9 +100,11 @@ variable [Nontrivial α] theorem cast_add_one_pos (n : ℕ) : 0 < (n : α) + 1 := zero_lt_one.trans_le <| le_add_of_nonneg_left n.cast_nonneg +#align nat.cast_add_one_pos Nat.cast_add_one_pos @[simp] theorem cast_pos {n : ℕ} : (0 : α) < n ↔ 0 < n := by cases n <;> simp [cast_add_one_pos] +#align nat.cast_pos Nat.cast_pos end Nontrivial @@ -108,29 +118,37 @@ theorem StrictMono_cast : StrictMono (Nat.cast : ℕ → α) := @[simps (config := { fullyApplied := false })] def castOrderEmbedding : ℕ ↪o α := OrderEmbedding.ofStrictMono Nat.cast Nat.StrictMono_cast +#align nat.cast_order_embedding Nat.castOrderEmbedding +#align nat.cast_order_embedding_apply Nat.castOrderEmbedding_apply @[simp, norm_cast] theorem cast_le : (m : α) ≤ n ↔ m ≤ n := StrictMono_cast.le_iff_le +#align nat.cast_le Nat.cast_le -- porting note: missing mono attribute -- @[simp, norm_cast, mono] @[simp, norm_cast] theorem cast_lt : (m : α) < n ↔ m < n := StrictMono_cast.lt_iff_lt +#align nat.cast_lt Nat.cast_lt @[simp, norm_cast] theorem one_lt_cast : 1 < (n : α) ↔ 1 < n := by rw [← cast_one, cast_lt] +#align nat.one_lt_cast Nat.one_lt_cast @[simp, norm_cast] theorem one_le_cast : 1 ≤ (n : α) ↔ 1 ≤ n := by rw [← cast_one, cast_le] +#align nat.one_le_cast Nat.one_le_cast @[simp, norm_cast] theorem cast_lt_one : (n : α) < 1 ↔ n = 0 := by rw [← cast_one, cast_lt, lt_succ_iff, ← bot_eq_zero, le_bot_iff] +#align nat.cast_lt_one Nat.cast_lt_one @[simp, norm_cast] theorem cast_le_one : (n : α) ≤ 1 ↔ n ≤ 1 := by rw [← cast_one, cast_le] +#align nat.cast_le_one Nat.cast_le_one end OrderedSemiring @@ -144,21 +162,26 @@ theorem cast_tsub [CanonicallyOrderedCommSemiring α] [Sub α] [OrderedSub α] exact mono_cast h · rcases le_iff_exists_add'.mp h with ⟨m, rfl⟩ rw [add_tsub_cancel_right, cast_add, add_tsub_cancel_right] +#align nat.cast_tsub Nat.cast_tsub @[simp, norm_cast] theorem cast_min [LinearOrderedSemiring α] {a b : ℕ} : ((min a b : ℕ) : α) = min (a : α) b := (@mono_cast α _).map_min +#align nat.cast_min Nat.cast_min @[simp, norm_cast] theorem cast_max [LinearOrderedSemiring α] {a b : ℕ} : ((max a b : ℕ) : α) = max (a : α) b := (@mono_cast α _).map_max +#align nat.cast_max Nat.cast_max @[simp, norm_cast] theorem abs_cast [LinearOrderedRing α] (a : ℕ) : |(a : α)| = a := abs_of_nonneg (cast_nonneg a) +#align nat.abs_cast Nat.abs_cast theorem coe_nat_dvd [Semiring α] {m n : ℕ} (h : m ∣ n) : (m : α) ∣ (n : α) := map_dvd (Nat.castRingHom α) h +#align nat.coe_nat_dvd Nat.coe_nat_dvd alias coe_nat_dvd ← _root_.Dvd.dvd.natCast @@ -178,10 +201,12 @@ theorem ext_nat' [AddMonoid A] [AddMonoidHomClass F ℕ A] (f g : F) (h : f 1 = | zero => simp_rw [Nat.zero_eq, map_zero f, map_zero g] | succ n ihn => simp [Nat.succ_eq_add_one, h, ihn] +#align ext_nat' ext_nat' @[ext] theorem AddMonoidHom.ext_nat [AddMonoid A] {f g : ℕ →+ A} : f 1 = g 1 → f = g := ext_nat' f g +#align add_monoid_hom.ext_nat AddMonoidHom.ext_nat variable [AddMonoidWithOne A] @@ -189,12 +214,14 @@ variable [AddMonoidWithOne A] theorem eq_natCast' [AddMonoidHomClass F ℕ A] (f : F) (h1 : f 1 = 1) : ∀ n : ℕ, f n = n | 0 => by simp [map_zero f] | n + 1 => by rw [map_add, h1, eq_natCast' f h1 n, Nat.cast_add_one] +#align eq_nat_cast' eq_natCast' theorem map_natCast' {A} [AddMonoidWithOne A] [AddMonoidHomClass F A B] (f : F) (h : f 1 = 1) : ∀ n : ℕ, f n = n | 0 => by simp [map_zero f] | n + 1 => by rw [Nat.cast_add, map_add, Nat.cast_add, map_natCast' f h n, Nat.cast_one, h, Nat.cast_one] +#align map_nat_cast' map_natCast' end AddMonoidHomClass @@ -209,10 +236,12 @@ theorem ext_nat'' [MonoidWithZeroHomClass F ℕ A] (f g : F) (h_pos : ∀ {n : rintro (_ | n) · simp [map_zero f, map_zero g] · exact h_pos n.succ_pos +#align ext_nat'' ext_nat'' @[ext] theorem MonoidWithZeroHom.ext_nat {f g : ℕ →*₀ A} : (∀ {n : ℕ}, 0 < n → f n = g n) → f = g := ext_nat'' f g +#align monoid_with_zero_hom.ext_nat MonoidWithZeroHom.ext_nat end MonoidWithZeroHomClass @@ -223,21 +252,26 @@ variable {R S F : Type _} [NonAssocSemiring R] [NonAssocSemiring S] @[simp] theorem eq_natCast [RingHomClass F ℕ R] (f : F) : ∀ n, f n = n := eq_natCast' f <| map_one f +#align eq_nat_cast eq_natCast @[simp] theorem map_natCast [RingHomClass F R S] (f : F) : ∀ n : ℕ, f (n : R) = n := map_natCast' f <| map_one f +#align map_nat_cast map_natCast theorem ext_nat [RingHomClass F ℕ R] (f g : F) : f = g := ext_nat' f g <| by simp only [map_one f, map_one g] +#align ext_nat ext_nat theorem NeZero.nat_of_injective {n : ℕ} [h : NeZero (n : R)] [RingHomClass F R S] {f : F} (hf : Function.Injective f) : NeZero (n : S) := ⟨fun h ↦ NeZero.natCast_ne n R <| hf <| by simpa only [map_natCast, map_zero f] ⟩ +#align ne_zero.nat_of_injective NeZero.nat_of_injective theorem NeZero.nat_of_neZero {R S} [Semiring R] [Semiring S] {F} [RingHomClass F R S] (f : F) {n : ℕ} [hn : NeZero (n : S)] : NeZero (n : R) := .of_map (f := f) (neZero := by simp only [map_natCast, hn]) +#align ne_zero.nat_of_ne_zero NeZero.nat_of_neZero end RingHomClass @@ -253,10 +287,12 @@ end RingHom @[simp, norm_cast] theorem Nat.cast_id (n : ℕ) : n.cast = n := rfl +#align nat.cast_id Nat.cast_id @[simp] theorem Nat.castRingHom_nat : Nat.castRingHom ℕ = RingHom.id ℕ := rfl +#align nat.cast_ring_hom_nat Nat.castRingHom_nat /-- We don't use `RingHomClass` here, since that might cause type-class slowdown for `Subsingleton`-/ @@ -271,10 +307,12 @@ variable [AddMonoidWithOne α] @[simp, norm_cast] theorem op_natCast (n : ℕ) : op (n : α) = n := rfl +#align mul_opposite.op_nat_cast MulOpposite.op_natCast @[simp, norm_cast] theorem unop_natCast (n : ℕ) : unop (n : αᵐᵒᵖ) = n := rfl +#align mul_opposite.unop_nat_cast MulOpposite.unop_natCast end MulOpposite @@ -288,16 +326,19 @@ instance natCast : NatCast (∀ a, π a) := { natCast := fun n _ ↦ n } theorem nat_apply (n : ℕ) (a : α) : (n : ∀ a, π a) a = n := rfl +#align pi.nat_apply Pi.nat_apply @[simp] theorem coe_nat (n : ℕ) : (n : ∀ a, π a) = fun _ ↦ ↑n := rfl +#align pi.coe_nat Pi.coe_nat end Pi theorem Sum.elim_natCast_natCast {α β γ : Type _} [NatCast γ] (n : ℕ) : Sum.elim (n : α → γ) (n : β → γ) = n := @Sum.elim_lam_const_lam_const α β γ n +#align sum.elim_nat_cast_nat_cast Sum.elim_natCast_natCast /-! ### Order dual -/ diff --git a/Mathlib/Data/Nat/Cast/Defs.lean b/Mathlib/Data/Nat/Cast/Defs.lean index f8ced014066a7..d4ef8570379ee 100644 --- a/Mathlib/Data/Nat/Cast/Defs.lean +++ b/Mathlib/Data/Nat/Cast/Defs.lean @@ -233,6 +233,7 @@ theorem one_add_one_eq_two [AddMonoidWithOne α] : 1 + 1 = (2 : α) := by rw [←Nat.cast_one, ←Nat.cast_add] apply congrArg decide +#align one_add_one_eq_two one_add_one_eq_two theorem two_add_one_eq_three [AddMonoidWithOne α] : 2 + 1 = (3 : α) := by rw [←one_add_one_eq_two, ←Nat.cast_one, ←Nat.cast_add, ←Nat.cast_add] diff --git a/Mathlib/Data/Nat/Factorial/Basic.lean b/Mathlib/Data/Nat/Factorial/Basic.lean index 4896f75aa150b..b06eb8dc2d353 100644 --- a/Mathlib/Data/Nat/Factorial/Basic.lean +++ b/Mathlib/Data/Nat/Factorial/Basic.lean @@ -403,6 +403,7 @@ theorem descFactorial_eq_zero_iff_lt {n : ℕ} : ∀ {k : ℕ}, n.descFactorial #align nat.desc_factorial_eq_zero_iff_lt Nat.descFactorial_eq_zero_iff_lt alias descFactorial_eq_zero_iff_lt ↔ _ descFactorial_of_lt +#align nat.desc_factorial_of_lt Nat.descFactorial_of_lt theorem add_descFactorial_eq_ascFactorial (n : ℕ) : ∀ k : ℕ, (n + k).descFactorial k = n.ascFactorial k diff --git a/Mathlib/Data/Nat/Pairing.lean b/Mathlib/Data/Nat/Pairing.lean index 3bd7d9af4c73a..85326a7955909 100644 --- a/Mathlib/Data/Nat/Pairing.lean +++ b/Mathlib/Data/Nat/Pairing.lean @@ -82,6 +82,8 @@ theorem unpair_mkpair (a b : ℕ) : unpair (mkpair a b) = (a, b) := by def mkpairEquiv : ℕ × ℕ ≃ ℕ := ⟨uncurry mkpair, unpair, fun ⟨a, b⟩ => unpair_mkpair a b, mkpair_unpair⟩ #align nat.mkpair_equiv Nat.mkpairEquiv +#align nat.mkpair_equiv_apply Nat.mkpairEquiv_apply +#align nat.mkpair_equiv_symm_apply Nat.mkpairEquiv_symm_apply theorem surjective_unpair : Surjective unpair := mkpairEquiv.symm.surjective diff --git a/Mathlib/Data/Nat/Pow.lean b/Mathlib/Data/Nat/Pow.lean index b70a74fb5eb17..edbea5d4f0515 100644 --- a/Mathlib/Data/Nat/Pow.lean +++ b/Mathlib/Data/Nat/Pow.lean @@ -154,6 +154,7 @@ theorem sq_sub_sq (a b : ℕ) : a ^ 2 - b ^ 2 = (a + b) * (a - b) := by #align nat.sq_sub_sq Nat.sq_sub_sq alias sq_sub_sq ← pow_two_sub_pow_two +#align nat.pow_two_sub_pow_two Nat.pow_two_sub_pow_two /-! ### `pow` and `mod` / `dvd` -/ diff --git a/Mathlib/Data/Nat/Prime.lean b/Mathlib/Data/Nat/Prime.lean index 190280d2fcd90..c70a25479b43f 100644 --- a/Mathlib/Data/Nat/Prime.lean +++ b/Mathlib/Data/Nat/Prime.lean @@ -476,6 +476,7 @@ theorem dvd_of_forall_prime_mul_dvd {a b : ℕ} · apply one_dvd obtain ⟨p, hp⟩ := exists_prime_and_dvd ha exact _root_.trans (dvd_mul_left a p) (hdvd p hp.1 hp.2) +#align nat.dvd_of_forall_prime_mul_dvd Nat.dvd_of_forall_prime_mul_dvd /-- Euclid's theorem on the **infinitude of primes**. Here given in the form: for every `n`, there exists a prime number `p ≥ n`. -/ @@ -582,6 +583,8 @@ theorem prime_iff {p : ℕ} : p.Prime ↔ _root_.Prime p := #align nat.prime_iff Nat.prime_iff alias prime_iff ↔ Prime.prime _root_.Prime.nat_prime +#align nat.prime.prime Nat.Prime.prime +#align prime.nat_prime Prime.nat_prime -- Porting note: attributes `protected`, `nolint dup_namespace` removed diff --git a/Mathlib/Data/Nat/WithBot.lean b/Mathlib/Data/Nat/WithBot.lean index 81616735902d0..c429702670290 100644 --- a/Mathlib/Data/Nat/WithBot.lean +++ b/Mathlib/Data/Nat/WithBot.lean @@ -89,6 +89,7 @@ theorem add_one_le_of_lt {n m : WithBot ℕ} (h : n < m) : n + 1 ≤ m := by · exact bot_le cases m exacts [(not_lt_bot h).elim, WithBot.some_le_some.2 (WithBot.some_lt_some.1 h)] +#align nat.with_bot.add_one_le_of_lt Nat.WithBot.add_one_le_of_lt end WithBot diff --git a/Mathlib/Data/Opposite.lean b/Mathlib/Data/Opposite.lean index e1722232a6f0c..d06bd9cc22ec9 100644 --- a/Mathlib/Data/Opposite.lean +++ b/Mathlib/Data/Opposite.lean @@ -64,14 +64,17 @@ variable {α} --@[pp_nodot] def op : α → αᵒᵖ := id +#align opposite.op Opposite.op /-- The canonical map `αᵒᵖ → α`. -/ -- Porting note: pp_nodot has not been implemented. --@[pp_nodot] def unop : αᵒᵖ → α := id +#align opposite.unop Opposite.unop theorem op_injective : Function.Injective (op : α → αᵒᵖ) := fun _ _ => id +#align opposite.op_injective Opposite.op_injective theorem unop_injective : Function.Injective (unop : αᵒᵖ → α) := fun _ _ => id #align opposite.unop_injective Opposite.unop_injective @@ -79,20 +82,24 @@ theorem unop_injective : Function.Injective (unop : αᵒᵖ → α) := fun _ _ @[simp] theorem op_unop (x : αᵒᵖ) : op (unop x) = x := rfl +#align opposite.op_unop Opposite.op_unop @[simp] theorem unop_op (x : α) : unop (op x) = x := rfl +#align opposite.unop_op Opposite.unop_op -- We could prove these by `Iff.rfl`, but that would make these eligible for `dsimp`. That would be -- a bad idea because `Opposite` is irreducible. @[simp] theorem op_inj_iff (x y : α) : op x = op y ↔ x = y := op_injective.eq_iff +#align opposite.op_inj_iff Opposite.op_inj_iff @[simp] theorem unop_inj_iff (x y : αᵒᵖ) : unop x = unop y ↔ x = y := unop_injective.eq_iff +#align opposite.unop_inj_iff Opposite.unop_inj_iff /-- The type-level equivalence between a type and its opposite. -/ def equivToOpposite : α ≃ αᵒᵖ where @@ -100,6 +107,7 @@ def equivToOpposite : α ≃ αᵒᵖ where invFun := unop left_inv := unop_op right_inv := op_unop +#align opposite.equiv_to_opposite Opposite.equivToOpposite @[simp] theorem equivToOpposite_coe : (equivToOpposite : α → αᵒᵖ) = op := @@ -113,9 +121,11 @@ theorem equivToOpposite_symm_coe : (equivToOpposite.symm : αᵒᵖ → α) = un theorem op_eq_iff_eq_unop {x : α} {y} : op x = y ↔ x = unop y := equivToOpposite.apply_eq_iff_eq_symm_apply +#align opposite.op_eq_iff_eq_unop Opposite.op_eq_iff_eq_unop theorem unop_eq_iff_eq_op {x} {y : α} : unop x = y ↔ x = op y := equivToOpposite.symm.apply_eq_iff_eq_symm_apply +#align opposite.unop_eq_iff_eq_op Opposite.unop_eq_iff_eq_op instance [Inhabited α] : Inhabited αᵒᵖ := ⟨op default⟩ @@ -123,5 +133,6 @@ instance [Inhabited α] : Inhabited αᵒᵖ := /-- A recursor for `Opposite`. Use as `induction x using Opposite.rec`. -/ @[simp] protected def rec {F : αᵒᵖ → Sort v} (h : ∀ X, F (op X)) : ∀ X, F X := fun X => h (unop X) +#align opposite.rec Opposite.rec end Opposite diff --git a/Mathlib/Data/Option/Basic.lean b/Mathlib/Data/Option/Basic.lean index 7fbff1ae82faf..a843be14131e9 100644 --- a/Mathlib/Data/Option/Basic.lean +++ b/Mathlib/Data/Option/Basic.lean @@ -41,6 +41,7 @@ variable {α β γ δ : Type _} theorem coe_def : (fun a ↦ ↑a : α → Option α) = some := rfl +#align option.coe_def Option.coe_def #align option.get_or_else Option.getD @@ -51,40 +52,49 @@ theorem getD_coe (x y : α) : Option.getD (↑x) y = x := theorem coe_get {o : Option α} (h : o.isSome) : ((Option.get _ h : α) : Option α) = o := Option.some_get h +#align option.coe_get Option.coe_get theorem eq_of_mem_of_mem {a : α} {o1 o2 : Option α} (h1 : a ∈ o1) (h2 : a ∈ o2) : o1 = o2 := h1.trans h2.symm +#align option.eq_of_mem_of_mem Option.eq_of_mem_of_mem theorem Mem.leftUnique : Relator.LeftUnique ((· ∈ ·) : α → Option α → Prop) := fun _ _ _=> mem_unique #align option.mem.left_unique Option.Mem.leftUnique theorem some_injective (α : Type _) : Function.Injective (@some α) := fun _ _ ↦ some_inj.mp +#align option.some_injective Option.some_injective /-- `Option.map f` is injective if `f` is injective. -/ theorem map_injective {f : α → β} (Hf : Function.Injective f) : Function.Injective (Option.map f) | none, none, _ => rfl | some a₁, some a₂, H => by rw [Hf (Option.some.inj H)] +#align option.map_injective Option.map_injective @[simp] theorem map_comp_some (f : α → β) : Option.map f ∘ some = some ∘ f := rfl +#align option.map_comp_some Option.map_comp_some @[simp] theorem none_bind' (f : α → Option β) : none.bind f = none := rfl +#align option.none_bind' Option.none_bind' @[simp] theorem some_bind' (a : α) (f : α → Option β) : (some a).bind f = f a := rfl +#align option.some_bind' Option.some_bind' theorem bind_eq_some' {x : Option α} {f : α → Option β} {b : β} : x.bind f = some b ↔ ∃ a, x = some a ∧ f a = some b := by cases x <;> simp +#align option.bind_eq_some' Option.bind_eq_some' theorem bind_eq_none' {o : Option α} {f : α → Option β} : o.bind f = none ↔ ∀ b a, a ∈ o → b ∉ f a := by simp only [eq_none_iff_forall_not_mem, mem_def, bind_eq_some, not_exists, not_and] +#align option.bind_eq_none' Option.bind_eq_none' theorem joinM_eq_join : joinM = @join α := funext fun _ ↦ rfl @@ -92,6 +102,7 @@ theorem joinM_eq_join : joinM = @join α := theorem bind_eq_bind {α β : Type _} {f : α → Option β} {x : Option α} : x >>= f = x.bind f := rfl +#align option.bind_eq_bind Option.bind_eq_bind --Porting note: New lemma used to prove a theorem in Data.List.Basic theorem map_eq_bind (f : α → β) (o : Option α) : @@ -100,28 +111,34 @@ theorem map_eq_bind (f : α → β) (o : Option α) : theorem map_coe {α β} {a : α} {f : α → β} : f <$> (a : Option α) = ↑(f a) := rfl +#align option.map_coe Option.map_coe @[simp] theorem map_coe' {a : α} {f : α → β} : Option.map f (a : Option α) = ↑(f a) := rfl +#align option.map_coe' Option.map_coe' /-- `Option.map` as a function between functions is injective. -/ theorem map_injective' : Function.Injective (@Option.map α β) := fun f g h ↦ funext fun x ↦ some_injective _ <| by simp only [← map_some', h] +#align option.map_injective' Option.map_injective' @[simp] theorem map_inj {f g : α → β} : Option.map f = Option.map g ↔ f = g := map_injective'.eq_iff +#align option.map_inj Option.map_inj attribute [simp] map_id @[simp] theorem map_eq_id {f : α → α} : Option.map f = id ↔ f = id := map_injective'.eq_iff' map_id +#align option.map_eq_id Option.map_eq_id theorem map_comm {f₁ : α → β} {f₂ : α → γ} {g₁ : β → δ} {g₂ : γ → δ} (h : g₁ ∘ f₁ = g₂ ∘ f₂) (a : α) : (Option.map f₁ a).map g₁ = (Option.map f₂ a).map g₂ := by rw [map_map, h, ← map_map] +#align option.map_comm Option.map_comm section pmap @@ -131,57 +148,70 @@ variable {p : α → Prop} (f : ∀ a : α, p a → β) (x : Option α) -- @[simp] theorem pbind_eq_bind (f : α → Option β) (x : Option α) : (x.pbind fun a _ ↦ f a) = x.bind f := by cases x <;> simp only [pbind, none_bind', some_bind'] +#align option.pbind_eq_bind Option.pbind_eq_bind theorem map_bind {α β γ} (f : β → γ) (x : Option α) (g : α → Option β) : Option.map f (x >>= g) = x >>= fun a ↦ Option.map f (g a) := by simp only [← map_eq_map, ← bind_pure_comp, LawfulMonad.bind_assoc] +#align option.map_bind Option.map_bind theorem map_bind' (f : β → γ) (x : Option α) (g : α → Option β) : Option.map f (x.bind g) = x.bind fun a ↦ Option.map f (g a) := by cases x <;> simp +#align option.map_bind' Option.map_bind' theorem map_pbind (f : β → γ) (x : Option α) (g : ∀ a, a ∈ x → Option β) : Option.map f (x.pbind g) = x.pbind fun a H ↦ Option.map f (g a H) := by cases x <;> simp only [pbind, map_none'] +#align option.map_pbind Option.map_pbind theorem pbind_map (f : α → β) (x : Option α) (g : ∀ b : β, b ∈ x.map f → Option γ) : pbind (Option.map f x) g = x.pbind fun a h ↦ g (f a) (mem_map_of_mem _ h) := by cases x <;> rfl +#align option.pbind_map Option.pbind_map @[simp] theorem pmap_none (f : ∀ a : α, p a → β) {H} : pmap f (@none α) H = none := rfl +#align option.pmap_none Option.pmap_none @[simp] theorem pmap_some (f : ∀ a : α, p a → β) {x : α} (h : p x) : pmap f (some x) = fun _ ↦ some (f x h) := rfl +#align option.pmap_some Option.pmap_some theorem mem_pmem {a : α} (h : ∀ a ∈ x, p a) (ha : a ∈ x) : f a (h a ha) ∈ pmap f x h := by rw [mem_def] at ha ⊢ subst ha rfl +#align option.mem_pmem Option.mem_pmem theorem pmap_map (g : γ → α) (x : Option γ) (H) : pmap f (x.map g) H = pmap (fun a h ↦ f (g a) h) x fun a h ↦ H _ (mem_map_of_mem _ h) := by cases x <;> simp only [map_none', map_some', pmap] +#align option.pmap_map Option.pmap_map theorem map_pmap (g : β → γ) (f : ∀ a, p a → β) (x H) : Option.map g (pmap f x H) = pmap (fun a h ↦ g (f a h)) x H := by cases x <;> simp only [map_none', map_some', pmap] +#align option.map_pmap Option.map_pmap -- Porting note: Can't simp tag this anymore because `pmap` simplifies -- @[simp] theorem pmap_eq_map (p : α → Prop) (f : α → β) (x H) : @pmap _ _ p (fun a _ ↦ f a) x H = Option.map f x := by cases x <;> simp only [map_none', map_some', pmap] +#align option.pmap_eq_map Option.pmap_eq_map theorem pmap_bind {α β γ} {x : Option α} {g : α → Option β} {p : β → Prop} {f : ∀ b, p b → γ} (H) (H' : ∀ (a : α), ∀ b ∈ g a, b ∈ x >>= g) : pmap f (x >>= g) H = x >>= fun a ↦ pmap f (g a) fun b h ↦ H _ (H' a _ h) := by cases x <;> simp only [pmap, bind_eq_bind, none_bind, some_bind] +#align option.pmap_bind Option.pmap_bind theorem bind_pmap {α β γ} {p : α → Prop} (f : ∀ a, p a → β) (x : Option α) (g : β → Option γ) (H) : pmap f x H >>= g = x.pbind fun a h ↦ g (f a (H _ h)) := by cases x <;> simp only [pmap, bind_eq_bind, none_bind, some_bind, pbind] +#align option.bind_pmap Option.bind_pmap variable {f x} @@ -192,6 +222,7 @@ theorem pbind_eq_none {f : ∀ a : α, a ∈ x → Option β} · simp only [pbind, iff_false] intro h cases h' _ rfl h +#align option.pbind_eq_none Option.pbind_eq_none theorem pbind_eq_some {f : ∀ a : α, a ∈ x → Option β} {y : β} : x.pbind f = some y ↔ ∃ (z : α) (H : z ∈ x), f z H = some y := by @@ -204,10 +235,12 @@ theorem pbind_eq_some {f : ∀ a : α, a ∈ x → Option β} {y : β} : rintro ⟨z, H, hz⟩ simp only [mem_def, Option.some_inj] at H simpa [H] using hz +#align option.pbind_eq_some Option.pbind_eq_some -- Porting note: Can't simp tag this anymore because `pmap` simplifies -- @[simp] theorem pmap_eq_none_iff {h} : pmap f x h = none ↔ x = none := by cases x <;> simp +#align option.pmap_eq_none_iff Option.pmap_eq_none_iff -- Porting note: Can't simp tag this anymore because `pmap` simplifies -- @[simp] @@ -222,18 +255,21 @@ theorem pmap_eq_some_iff {hf} {y : β} : · rintro ⟨a, H, rfl⟩ simp only [mem_def, Option.some_inj] at H simp only [H, pmap] +#align option.pmap_eq_some_iff Option.pmap_eq_some_iff -- Porting note: Can't simp tag this anymore because `join` and `pmap` simplify -- @[simp] theorem join_pmap_eq_pmap_join {f : ∀ a, p a → β} {x : Option (Option α)} (H) : (pmap (pmap f) x H).join = pmap f x.join fun a h ↦ H (some a) (mem_of_mem_join h) _ rfl := by rcases x with (_ | _ | x) <;> simp +#align option.join_pmap_eq_pmap_join Option.join_pmap_eq_pmap_join end pmap @[simp] theorem seq_some {α β} {a : α} {f : α → β} : some f <*> some a = some (f a) := rfl +#align option.seq_some Option.seq_some @[simp] theorem some_orElse' (a : α) (x : Option α) : (some a).orElse (fun _ ↦ x) = some a := @@ -273,9 +309,11 @@ theorem orElse_none' (x : Option α) : x.orElse (fun _ ↦ none) = x := by cases theorem iget_mem [Inhabited α] : ∀ {o : Option α}, isSome o → o.iget ∈ o | some _, _ => rfl +#align option.iget_mem Option.iget_mem theorem iget_of_mem [Inhabited α] {a : α} : ∀ {o : Option α}, a ∈ o → o.iget = a | _, rfl => rfl +#align option.iget_of_mem Option.iget_of_mem theorem getD_default_eq_iget [Inhabited α] (o : Option α) : o.getD default = o.iget := by cases o <;> rfl @@ -285,6 +323,7 @@ theorem getD_default_eq_iget [Inhabited α] (o : Option α) : theorem guard_eq_some' {p : Prop} [Decidable p] (u) : _root_.guard p = some u ↔ p := by cases u by_cases p <;> simp [_root_.guard, h] +#align option.guard_eq_some' Option.guard_eq_some' theorem liftOrGet_choice {f : α → α → α} (h : ∀ a b, f a b = a ∨ f a b = b) : ∀ o₁ o₂, liftOrGet f o₁ o₂ = o₁ ∨ liftOrGet f o₁ o₂ = o₂ @@ -305,44 +344,53 @@ function to `a` if it comes from `α`, and return `b` otherwise. -/ def casesOn' : Option α → β → (α → β) → β | none, n, _ => n | some a, _, s => s a +#align option.cases_on' Option.casesOn' @[simp] theorem casesOn'_none (x : β) (f : α → β) : casesOn' none x f = x := rfl +#align option.cases_on'_none Option.casesOn'_none @[simp] theorem casesOn'_some (x : β) (f : α → β) (a : α) : casesOn' (some a) x f = f a := rfl +#align option.cases_on'_some Option.casesOn'_some @[simp] theorem casesOn'_coe (x : β) (f : α → β) (a : α) : casesOn' (a : Option α) x f = f a := rfl +#align option.cases_on'_coe Option.casesOn'_coe -- Porting note: Left-hand side does not simplify. -- @[simp] theorem casesOn'_none_coe (f : Option α → β) (o : Option α) : casesOn' o (f none) (f ∘ (fun a ↦ ↑a)) = f o := by cases o <;> rfl +#align option.cases_on'_none_coe Option.casesOn'_none_coe theorem orElse_eq_some (o o' : Option α) (x : α) : (o <|> o') = some x ↔ o = some x ∨ o = none ∧ o' = some x := by cases o · simp only [true_and, false_or, eq_self_iff_true, none_orElse] · simp only [some_orElse, or_false, false_and] +#align option.orelse_eq_some Option.orElse_eq_some theorem orElse_eq_some' (o o' : Option α) (x : α) : o.orElse (fun _ ↦ o') = some x ↔ o = some x ∨ o = none ∧ o' = some x := Option.orElse_eq_some o o' x +#align option.orelse_eq_some' Option.orElse_eq_some' @[simp] theorem orElse_eq_none (o o' : Option α) : (o <|> o') = none ↔ o = none ∧ o' = none := by cases o · simp only [true_and, none_orElse, eq_self_iff_true] · simp only [some_orElse, false_and] +#align option.orelse_eq_none Option.orElse_eq_none @[simp] theorem orElse_eq_none' (o o' : Option α) : o.orElse (fun _ ↦ o') = none ↔ o = none ∧ o' = none := Option.orElse_eq_none o o' +#align option.orelse_eq_none' Option.orElse_eq_none' section @@ -350,6 +398,7 @@ open Classical theorem choice_eq_none (α : Type _) [IsEmpty α] : choice α = none := dif_neg (not_nonempty_iff_imp_false.mpr isEmptyElim) +#align option.choice_eq_none Option.choice_eq_none #align option.choice_is_some_iff_nonempty Option.choice_isSome_iff_nonempty @@ -359,5 +408,6 @@ end -- @[simp] theorem elim_none_some (f : Option α → β) : (fun x ↦ Option.elim x (f none) (f ∘ some)) = f := funext fun o ↦ by cases o <;> rfl +#align option.elim_none_some Option.elim_none_some end Option diff --git a/Mathlib/Data/Option/Defs.lean b/Mathlib/Data/Option/Defs.lean index d07424fa870f6..29b454fe5e2fb 100644 --- a/Mathlib/Data/Option/Defs.lean +++ b/Mathlib/Data/Option/Defs.lean @@ -29,12 +29,14 @@ inductive rel (r : α → β → Prop) : Option α → Option β → Prop some {a b} : r a b → rel r (some a) (some b) | /-- `none ~ none` -/ none : rel r none none +#align option.rel Option.rel /-- Traverse an object of `Option α` with a function `f : α → F β` for an applicative `F`. -/ protected def traverse.{u, v} {F : Type u → Type v} [Applicative F] {α β : Type _} (f : α → F β) : Option α → F (Option β) | none => pure none | some x => some <$> f x +#align option.traverse Option.traverse /-- If you maybe have a monadic computation in a `[Monad m]` which produces a term of type `α`, then there is a naturally associated way to always perform a computation in `m` which maybe @@ -42,6 +44,7 @@ produces a result. -/ def maybe.{u, v} {m : Type u → Type v} [Monad m] {α : Type u} : Option (m α) → m (Option α) | none => pure none | some fn => some <$> fn +#align option.maybe Option.maybe #align option.mmap Option.mapM #align option.melim Option.elimM @@ -61,6 +64,7 @@ protected def elim' (b : β) (f : α → β) : Option α → β #align option.elim Option.elim' theorem mem_some_iff {α : Type _} {a b : α} : a ∈ some b ↔ b = a := by simp +#align option.mem_some_iff Option.mem_some_iff /-- `o = none` is decidable even if the wrapped type does not have decidable equality. This is not an instance because it is not definitionally equal to `Option.decidableEq`. @@ -69,6 +73,7 @@ Try to use `o.isNone` or `o.isSome` instead. @[inline] def decidableEqNone {o : Option α} : Decidable (o = none) := decidable_of_decidable_of_iff isNone_iff_eq_none +#align option.decidable_eq_none Option.decidableEqNone instance decidableForallMem {p : α → Prop} [DecidablePred p] : ∀ o : Option α, Decidable (∀ a ∈ o, p a) @@ -87,9 +92,11 @@ instance decidableExistsMem {p : α → Prop} [DecidablePred p] : def iget [Inhabited α] : Option α → α | some x => x | none => default +#align option.iget Option.iget theorem iget_some [Inhabited α] {a : α} : (some a).iget = a := rfl +#align option.iget_some Option.iget_some @[simp] theorem mem_toList {a : α} {o : Option α} : a ∈ toList o ↔ a ∈ o := by diff --git a/Mathlib/Data/PNat/Basic.lean b/Mathlib/Data/PNat/Basic.lean index f83dd83882a51..3110b8901918f 100644 --- a/Mathlib/Data/PNat/Basic.lean +++ b/Mathlib/Data/PNat/Basic.lean @@ -149,6 +149,8 @@ def _root_.Equiv.pnatEquivNat : ℕ+ ≃ ℕ where left_inv := succPNat_natPred right_inv := Nat.natPred_succPNat #align equiv.pnat_equiv_nat Equiv.pnatEquivNat +#align equiv.pnat_equiv_nat_symm_apply Equiv.pnatEquivNat_symm_apply +#align equiv.pnat_equiv_nat_apply Equiv.pnatEquivNat_apply /-- The order isomorphism between ℕ and ℕ+ given by `succ`. -/ @[simps (config := { fullyApplied := false }) apply] @@ -156,6 +158,7 @@ def _root_.OrderIso.pnatIsoNat : ℕ+ ≃o ℕ where toEquiv := Equiv.pnatEquivNat map_rel_iff' := natPred_le_natPred #align order_iso.pnat_iso_nat OrderIso.pnatIsoNat +#align order_iso.pnat_iso_nat_apply OrderIso.pnatIsoNat_apply @[simp] theorem _root_.OrderIso.pnatIsoNat_symm_apply : OrderIso.pnatIsoNat.symm = Nat.succPNat := diff --git a/Mathlib/Data/Part.lean b/Mathlib/Data/Part.lean index ad53f31f7ee81..c3c93336a9afe 100644 --- a/Mathlib/Data/Part.lean +++ b/Mathlib/Data/Part.lean @@ -435,6 +435,8 @@ protected def bind (f : Part α) (g : α → Part β) : Part β := def map (f : α → β) (o : Part α) : Part β := ⟨o.Dom, f ∘ o.get⟩ #align part.map Part.map +#align part.map_dom Part.map_Dom +#align part.map_get Part.map_get theorem mem_map (f : α → β) {o : Part α} : ∀ {a}, a ∈ o → f a ∈ map f o | _, ⟨_, rfl⟩ => ⟨_, rfl⟩ diff --git a/Mathlib/Data/Pi/Algebra.lean b/Mathlib/Data/Pi/Algebra.lean index e796ca5dc3978..0cb38b7316372 100644 --- a/Mathlib/Data/Pi/Algebra.lean +++ b/Mathlib/Data/Pi/Algebra.lean @@ -51,22 +51,32 @@ instance instOne [∀ i, One <| f i] : One (∀ i : I, f i) := @[to_additive (attr := simp)] theorem one_apply [∀ i, One <| f i] : (1 : ∀ i, f i) i = 1 := rfl +#align pi.one_apply Pi.one_apply +#align pi.zero_apply Pi.zero_apply @[to_additive] theorem one_def [∀ i, One <| f i] : (1 : ∀ i, f i) = fun _ => 1 := rfl +#align pi.one_def Pi.one_def +#align pi.zero_def Pi.zero_def @[to_additive (attr := simp)] theorem const_one [One β] : const α (1 : β) = 1 := rfl +#align pi.const_one Pi.const_one +#align pi.const_zero Pi.const_zero @[to_additive (attr := simp)] theorem one_comp [One γ] (x : α → β) : (1 : β → γ) ∘ x = 1 := rfl +#align pi.one_comp Pi.one_comp +#align pi.zero_comp Pi.zero_comp @[to_additive (attr := simp)] theorem comp_one [One β] (x : β → γ) : x ∘ (1 : α → β) = const α (x 1) := rfl +#align pi.comp_one Pi.comp_one +#align pi.comp_zero Pi.comp_zero @[to_additive] instance instMul [∀ i, Mul <| f i] : Mul (∀ i : I, f i) := @@ -78,18 +88,26 @@ instance instMul [∀ i, Mul <| f i] : Mul (∀ i : I, f i) := @[to_additive (attr := simp)] theorem mul_apply [∀ i, Mul <| f i] : (x * y) i = x i * y i := rfl +#align pi.mul_apply Pi.mul_apply +#align pi.add_apply Pi.add_apply @[to_additive] theorem mul_def [∀ i, Mul <| f i] : x * y = fun i => x i * y i := rfl +#align pi.mul_def Pi.mul_def +#align pi.add_def Pi.add_def @[to_additive (attr := simp)] theorem const_mul [Mul β] (a b : β) : const α a * const α b = const α (a * b) := rfl +#align pi.const_mul Pi.const_mul +#align pi.const_add Pi.const_add @[to_additive] theorem mul_comp [Mul γ] (x y : β → γ) (z : α → β) : (x * y) ∘ z = x ∘ z * y ∘ z := rfl +#align pi.mul_comp Pi.mul_comp +#align pi.add_comp Pi.add_comp @[to_additive] instance instSMul [∀ i, SMul α <| f i] : SMul α (∀ i : I, f i) := @@ -105,20 +123,32 @@ instance instPow [∀ i, Pow (f i) β] : Pow (∀ i, f i) β := @[to_additive (attr := simp) (reorder := 5) smul_apply] theorem pow_apply [∀ i, Pow (f i) β] (x : ∀ i, f i) (b : β) (i : I) : (x ^ b) i = x i ^ b := rfl +#align pi.pow_apply Pi.pow_apply +#align pi.smul_apply Pi.smul_apply @[to_additive (reorder := 5) smul_def] theorem pow_def [∀ i, Pow (f i) β] (x : ∀ i, f i) (b : β) : x ^ b = fun i => x i ^ b := rfl +#align pi.pow_def Pi.pow_def +#align pi.smul_def Pi.smul_def @[to_additive (attr := simp) (reorder := 2 5) smul_const] theorem const_pow [Pow α β] (a : α) (b : β) : const I a ^ b = const I (a ^ b) := rfl +#align pi.const_pow Pi.const_pow +#align pi.smul_const Pi.smul_const @[to_additive (reorder := 6) smul_comp] theorem pow_comp [Pow γ α] (x : β → γ) (a : α) (y : I → β) : (x ^ a) ∘ y = x ∘ y ^ a := rfl +#align pi.pow_comp Pi.pow_comp +#align pi.smul_comp Pi.smul_comp attribute [to_additive] smul_apply smul_def smul_const smul_comp +#align pi.vadd_def Pi.vadd_def +#align pi.vadd_const Pi.vadd_const +#align pi.vadd_apply Pi.vadd_apply +#align pi.vadd_comp Pi.vadd_comp /-! Porting note: `bit0` and `bit1` are deprecated. This section can be removed entirely @@ -131,10 +161,12 @@ set_option linter.deprecated false @[simp, deprecated] theorem bit0_apply [∀ i, Add <| f i] : (bit0 x) i = bit0 (x i) := rfl +#align pi.bit0_apply Pi.bit0_apply @[simp, deprecated] theorem bit1_apply [∀ i, Add <| f i] [∀ i, One <| f i] : (bit1 x) i = bit1 (x i) := rfl +#align pi.bit1_apply Pi.bit1_apply end deprecated @@ -148,18 +180,26 @@ instance instInv [∀ i, Inv <| f i] : Inv (∀ i : I, f i) := @[to_additive (attr := simp)] theorem inv_apply [∀ i, Inv <| f i] : x⁻¹ i = (x i)⁻¹ := rfl +#align pi.inv_apply Pi.inv_apply +#align pi.neg_apply Pi.neg_apply @[to_additive] theorem inv_def [∀ i, Inv <| f i] : x⁻¹ = fun i => (x i)⁻¹ := rfl +#align pi.inv_def Pi.inv_def +#align pi.neg_def Pi.neg_def @[to_additive] theorem const_inv [Inv β] (a : β) : (const α a)⁻¹ = const α a⁻¹ := rfl +#align pi.const_inv Pi.const_inv +#align pi.const_neg Pi.const_neg @[to_additive] theorem inv_comp [Inv γ] (x : β → γ) (y : α → β) : x⁻¹ ∘ y = (x ∘ y)⁻¹ := rfl +#align pi.inv_comp Pi.inv_comp +#align pi.neg_comp Pi.neg_comp @[to_additive] instance instDiv [∀ i, Div <| f i] : Div (∀ i : I, f i) := @@ -171,18 +211,26 @@ instance instDiv [∀ i, Div <| f i] : Div (∀ i : I, f i) := @[to_additive (attr := simp)] theorem div_apply [∀ i, Div <| f i] : (x / y) i = x i / y i := rfl +#align pi.div_apply Pi.div_apply +#align pi.sub_apply Pi.sub_apply @[to_additive] theorem div_def [∀ i, Div <| f i] : x / y = fun i => x i / y i := rfl +#align pi.div_def Pi.div_def +#align pi.sub_def Pi.sub_def @[to_additive] theorem div_comp [Div γ] (x y : β → γ) (z : α → β) : (x / y) ∘ z = x ∘ z / y ∘ z := rfl +#align pi.div_comp Pi.div_comp +#align pi.sub_comp Pi.sub_comp @[to_additive (attr := simp)] theorem const_div [Div β] (a b : β) : const α a / const α b = const α (a / b) := rfl +#align pi.const_div Pi.const_div +#align pi.const_sub Pi.const_sub section @@ -311,16 +359,19 @@ end @[simp] protected def prod (f' : ∀ i, f i) (g' : ∀ i, g i) (i : I) : f i × g i := (f' i, g' i) +#align pi.prod Pi.prod -- Porting note : simp now unfolds the lhs, so we are not marking these as simp. -- @[simp] theorem prod_fst_snd : Pi.prod (Prod.fst : α × β → α) (Prod.snd : α × β → β) = id := funext fun _ => Prod.mk.eta +#align pi.prod_fst_snd Pi.prod_fst_snd -- Porting note : simp now unfolds the lhs, so we are not marking these as simp. -- @[simp] theorem prod_snd_fst : Pi.prod (Prod.snd : α × β → β) (Prod.fst : α × β → α) = Prod.swap := rfl +#align pi.prod_snd_fst Pi.prod_snd_fst end Pi @@ -331,6 +382,8 @@ section Extend @[to_additive] theorem extend_one [One γ] (f : α → β) : Function.extend f (1 : α → γ) (1 : β → γ) = 1 := funext fun _ => by apply ite_self +#align function.extend_one Function.extend_one +#align function.extend_zero Function.extend_zero @[to_additive] theorem extend_mul [Mul γ] (f : α → β) (g₁ g₂ : α → γ) (e₁ e₂ : β → γ): @@ -343,6 +396,8 @@ theorem extend_mul [Mul γ] (f : α → β) (g₁ g₂ : α → γ) (e₁ e₂ : -- which converts to -- `funext fun _ => by convert (apply_dite₂ (· * ·) _ _ _ _ _).symm` -- However this does not work, and we're not sure why. +#align function.extend_mul Function.extend_mul +#align function.extend_add Function.extend_add @[to_additive] theorem extend_inv [Inv γ] (f : α → β) (g : α → γ) (e : β → γ) : @@ -355,6 +410,8 @@ theorem extend_inv [Inv γ] (f : α → β) (g : α → γ) (e : β → γ) : -- which converts to -- `funext fun _ => by convert (apply_dite Inv.inv _ _ _).symm` -- However this does not work, and we're not sure why. +#align function.extend_inv Function.extend_inv +#align function.extend_neg Function.extend_neg @[to_additive] theorem extend_div [Div γ] (f : α → β) (g₁ g₂ : α → γ) (e₁ e₂ : β → γ) : @@ -367,20 +424,25 @@ theorem extend_div [Div γ] (f : α → β) (g₁ g₂ : α → γ) (e₁ e₂ : -- which converts to -- `funext fun _ => by convert (apply_dite₂ (· / ·) _ _ _ _ _).symm` -- However this does not work, and we're not sure why. +#align function.extend_div Function.extend_div +#align function.extend_sub Function.extend_sub end Extend theorem surjective_pi_map {F : ∀ i, f i → g i} (hF : ∀ i, Surjective (F i)) : Surjective fun x : ∀ i, f i => fun i => F i (x i) := fun y => ⟨fun i => (hF i (y i)).choose, funext fun i => (hF i (y i)).choose_spec⟩ +#align function.surjective_pi_map Function.surjective_pi_map theorem injective_pi_map {F : ∀ i, f i → g i} (hF : ∀ i, Injective (F i)) : Injective fun x : ∀ i, f i => fun i => F i (x i) := fun _ _ h => funext fun i => hF i <| (congr_fun h i : _) +#align function.injective_pi_map Function.injective_pi_map theorem bijective_pi_map {F : ∀ i, f i → g i} (hF : ∀ i, Bijective (F i)) : Bijective fun x : ∀ i, f i => fun i => F i (x i) := ⟨injective_pi_map fun i => (hF i).injective, surjective_pi_map fun i => (hF i).surjective⟩ +#align function.bijective_pi_map Function.bijective_pi_map end Function @@ -389,6 +451,8 @@ end Function def uniqueOfSurjectiveOne (α : Type _) {β : Type _} [One β] (h : Function.Surjective (1 : α → β)) : Unique β := h.uniqueOfSurjectiveConst α (1 : β) +#align unique_of_surjective_one uniqueOfSurjectiveOne +#align unique_of_surjective_zero uniqueOfSurjectiveZero @[to_additive] theorem Subsingleton.pi_mulSingle_eq {α : Type _} [DecidableEq I] [Subsingleton I] [One α] @@ -405,6 +469,8 @@ variable (a a' : α → γ) (b b' : β → γ) @[to_additive (attr := simp)] theorem elim_one_one [One γ] : Sum.elim (1 : α → γ) (1 : β → γ) = 1 := Sum.elim_const_const 1 +#align sum.elim_one_one Sum.elim_one_one +#align sum.elim_zero_zero Sum.elim_zero_zero @[to_additive (attr := simp)] theorem elim_mulSingle_one [DecidableEq α] [DecidableEq β] [One γ] (i : α) (c : γ) : @@ -424,15 +490,21 @@ theorem elim_one_mulSingle [DecidableEq α] [DecidableEq β] [One γ] (i : β) ( @[to_additive] theorem elim_inv_inv [Inv γ] : Sum.elim a⁻¹ b⁻¹ = (Sum.elim a b)⁻¹ := (Sum.comp_elim Inv.inv a b).symm +#align sum.elim_inv_inv Sum.elim_inv_inv +#align sum.elim_neg_neg Sum.elim_neg_neg @[to_additive] theorem elim_mul_mul [Mul γ] : Sum.elim (a * a') (b * b') = Sum.elim a b * Sum.elim a' b' := by ext x cases x <;> rfl +#align sum.elim_mul_mul Sum.elim_mul_mul +#align sum.elim_add_add Sum.elim_add_add @[to_additive] theorem elim_div_div [Div γ] : Sum.elim (a / a') (b / b') = Sum.elim a b / Sum.elim a' b' := by ext x cases x <;> rfl +#align sum.elim_div_div Sum.elim_div_div +#align sum.elim_sub_sub Sum.elim_sub_sub end Sum diff --git a/Mathlib/Data/Prod/Basic.lean b/Mathlib/Data/Prod/Basic.lean index 3d9667e184e1f..06abca272fb32 100644 --- a/Mathlib/Data/Prod/Basic.lean +++ b/Mathlib/Data/Prod/Basic.lean @@ -31,40 +31,51 @@ namespace Prod @[simp] theorem «forall» {p : α × β → Prop} : (∀ x, p x) ↔ ∀ a b, p (a, b) := ⟨fun h a b ↦ h (a, b), fun h ⟨a, b⟩ ↦ h a b⟩ +#align prod.forall Prod.forall @[simp] theorem «exists» {p : α × β → Prop} : (∃ x, p x) ↔ ∃ a b, p (a, b) := ⟨fun ⟨⟨a, b⟩, h⟩ ↦ ⟨a, b, h⟩, fun ⟨a, b, h⟩ ↦ ⟨⟨a, b⟩, h⟩⟩ +#align prod.exists Prod.exists theorem forall' {p : α → β → Prop} : (∀ x : α × β, p x.1 x.2) ↔ ∀ a b, p a b := Prod.forall +#align prod.forall' Prod.forall' theorem exists' {p : α → β → Prop} : (∃ x : α × β, p x.1 x.2) ↔ ∃ a b, p a b := Prod.exists +#align prod.exists' Prod.exists' @[simp] theorem snd_comp_mk (x : α) : Prod.snd ∘ (Prod.mk x : β → α × β) = id := rfl +#align prod.snd_comp_mk Prod.snd_comp_mk @[simp] theorem fst_comp_mk (x : α) : Prod.fst ∘ (Prod.mk x : β → α × β) = Function.const β x := rfl +#align prod.fst_comp_mk Prod.fst_comp_mk @[simp] theorem map_mk (f : α → γ) (g : β → δ) (a : α) (b : β) : map f g (a, b) = (f a, g b) := rfl +#align prod.map_mk Prod.map_mk theorem map_fst (f : α → γ) (g : β → δ) (p : α × β) : (map f g p).1 = f p.1 := rfl +#align prod.map_fst Prod.map_fst theorem map_snd (f : α → γ) (g : β → δ) (p : α × β) : (map f g p).2 = g p.2 := rfl +#align prod.map_snd Prod.map_snd theorem map_fst' (f : α → γ) (g : β → δ) : Prod.fst ∘ map f g = f ∘ Prod.fst := funext <| map_fst f g +#align prod.map_fst' Prod.map_fst' theorem map_snd' (f : α → γ) (g : β → δ) : Prod.snd ∘ map f g = g ∘ Prod.snd := funext <| map_snd f g +#align prod.map_snd' Prod.map_snd' /-- Composing a `Prod.map` with another `Prod.map` is equal to a single `Prod.map` of composed functions. @@ -72,6 +83,7 @@ a single `Prod.map` of composed functions. theorem map_comp_map {ε ζ : Type _} (f : α → β) (f' : γ → δ) (g : β → ε) (g' : δ → ζ) : Prod.map g g' ∘ Prod.map f f' = Prod.map (g ∘ f) (g' ∘ f') := rfl +#align prod.map_comp_map Prod.map_comp_map /-- Composing a `Prod.map` with another `Prod.map` is equal to a single `Prod.map` of composed functions, fully applied. @@ -79,27 +91,34 @@ a single `Prod.map` of composed functions, fully applied. theorem map_map {ε ζ : Type _} (f : α → β) (f' : γ → δ) (g : β → ε) (g' : δ → ζ) (x : α × γ) : Prod.map g g' (Prod.map f f' x) = Prod.map (g ∘ f) (g' ∘ f') x := rfl +#align prod.map_map Prod.map_map -- Porting note: mathlib3 proof uses `by cc` for the mpr direction -- Porting note: `@[simp]` tag removed because auto-generated `mk.injEq` simplifies LHS -- @[simp] theorem mk.inj_iff {a₁ a₂ : α} {b₁ b₂ : β} : (a₁, b₁) = (a₂, b₂) ↔ a₁ = a₂ ∧ b₁ = b₂ := Iff.of_eq (mk.injEq _ _ _ _) +#align prod.mk.inj_iff Prod.mk.inj_iff theorem mk.inj_left {α β : Type _} (a : α) : Function.Injective (Prod.mk a : β → α × β) := by intro b₁ b₂ h simpa only [true_and, Prod.mk.inj_iff, eq_self_iff_true] using h +#align prod.mk.inj_left Prod.mk.inj_left theorem mk.inj_right {α β : Type _} (b : β) : Function.Injective (fun a ↦ Prod.mk a b : α → α × β) := by intro b₁ b₂ h simpa only [and_true, eq_self_iff_true, mk.inj_iff] using h +#align prod.mk.inj_right Prod.mk.inj_right lemma mk_inj_left : (a, b₁) = (a, b₂) ↔ b₁ = b₂ := (mk.inj_left _).eq_iff lemma mk_inj_right : (a₁, b) = (a₂, b) ↔ a₁ = a₂ := (mk.inj_right _).eq_iff +#align prod.mk_inj_right Prod.mk_inj_right +#align prod.mk_inj_left Prod.mk_inj_left theorem ext_iff {p q : α × β} : p = q ↔ p.1 = q.1 ∧ p.2 = q.2 := by rw [← @mk.eta _ _ p, ← @mk.eta _ _ q, mk.inj_iff] +#align prod.ext_iff Prod.ext_iff @[ext] theorem ext {α β} {p q : α × β} (h₁ : p.1 = q.1) (h₂ : p.2 = q.2) : p = q := @@ -108,47 +127,60 @@ theorem ext {α β} {p q : α × β} (h₁ : p.1 = q.1) (h₂ : p.2 = q.2) : p = theorem map_def {f : α → γ} {g : β → δ} : Prod.map f g = fun p : α × β ↦ (f p.1, g p.2) := funext fun p ↦ ext (map_fst f g p) (map_snd f g p) +#align prod.map_def Prod.map_def theorem id_prod : (fun p : α × β ↦ (p.1, p.2)) = id := rfl +#align prod.id_prod Prod.id_prod theorem map_id : Prod.map (@id α) (@id β) = id := id_prod +#align prod.map_id Prod.map_id theorem fst_surjective [h : Nonempty β] : Function.Surjective (@fst α β) := fun x ↦ h.elim fun y ↦ ⟨⟨x, y⟩, rfl⟩ +#align prod.fst_surjective Prod.fst_surjective theorem snd_surjective [h : Nonempty α] : Function.Surjective (@snd α β) := fun y ↦ h.elim fun x ↦ ⟨⟨x, y⟩, rfl⟩ +#align prod.snd_surjective Prod.snd_surjective theorem fst_injective [Subsingleton β] : Function.Injective (@fst α β) := fun _ _ h ↦ ext h (Subsingleton.elim _ _) +#align prod.fst_injective Prod.fst_injective theorem snd_injective [Subsingleton α] : Function.Injective (@snd α β) := fun _ _ h ↦ ext (Subsingleton.elim _ _) h +#align prod.snd_injective Prod.snd_injective /-- Swap the factors of a product. `swap (a, b) = (b, a)` -/ def swap : α × β → β × α := fun p ↦ (p.2, p.1) +#align prod.swap Prod.swap @[simp] theorem swap_swap : ∀ x : α × β, swap (swap x) = x | ⟨_, _⟩ => rfl +#align prod.swap_swap Prod.swap_swap @[simp] theorem fst_swap {p : α × β} : (swap p).1 = p.2 := rfl +#align prod.fst_swap Prod.fst_swap @[simp] theorem snd_swap {p : α × β} : (swap p).2 = p.1 := rfl +#align prod.snd_swap Prod.snd_swap @[simp] theorem swap_prod_mk {a : α} {b : β} : swap (a, b) = (b, a) := rfl +#align prod.swap_prod_mk Prod.swap_prod_mk @[simp] theorem swap_swap_eq : swap ∘ swap = @id (α × β) := funext swap_swap +#align prod.swap_swap_eq Prod.swap_swap_eq @[simp] theorem swap_leftInverse : Function.LeftInverse (@swap α β) swap := @@ -162,25 +194,32 @@ theorem swap_rightInverse : Function.RightInverse (@swap α β) swap := theorem swap_injective : Function.Injective (@swap α β) := swap_leftInverse.injective +#align prod.swap_injective Prod.swap_injective theorem swap_surjective : Function.Surjective (@swap α β) := swap_leftInverse.surjective +#align prod.swap_surjective Prod.swap_surjective theorem swap_bijective : Function.Bijective (@swap α β) := ⟨swap_injective, swap_surjective⟩ +#align prod.swap_bijective Prod.swap_bijective @[simp] theorem swap_inj {p q : α × β} : swap p = swap q ↔ p = q := swap_injective.eq_iff +#align prod.swap_inj Prod.swap_inj theorem eq_iff_fst_eq_snd_eq : ∀ {p q : α × β}, p = q ↔ p.1 = q.1 ∧ p.2 = q.2 | ⟨p₁, p₂⟩, ⟨q₁, q₂⟩ => by simp +#align prod.eq_iff_fst_eq_snd_eq Prod.eq_iff_fst_eq_snd_eq theorem fst_eq_iff : ∀ {p : α × β} {x : α}, p.1 = x ↔ p = (x, p.2) | ⟨a, b⟩, x => by simp +#align prod.fst_eq_iff Prod.fst_eq_iff theorem snd_eq_iff : ∀ {p : α × β} {x : β}, p.2 = x ↔ p = (p.1, x) | ⟨a, b⟩, x => by simp +#align prod.snd_eq_iff Prod.snd_eq_iff variable {r : α → α → Prop} {s : β → β → Prop} {x y : α × β} @@ -190,8 +229,10 @@ theorem lex_def (r : α → α → Prop) (s : β → β → Prop) {p q : α × match p, q, h with | (a, b), (c, d), Or.inl h => Lex.left _ _ h | (a, b), (c, d), Or.inr ⟨e, h⟩ => by subst e; exact Lex.right _ h⟩ +#align prod.lex_def Prod.lex_def lemma lex_iff : Prod.Lex r s x y ↔ r x.1 y.1 ∨ x.1 = y.1 ∧ s x.2 y.2 := lex_def _ _ +#align prod.lex_iff Prod.lex_iff instance Lex.decidable [DecidableEq α] (r : α → α → Prop) (s : β → β → Prop) [DecidableRel r] [DecidableRel s] : @@ -201,6 +242,7 @@ instance Lex.decidable [DecidableEq α] @[refl] theorem Lex.refl_left (r : α → α → Prop) (s : β → β → Prop) [IsRefl α r] : ∀ x, Prod.Lex r s x x | (_, _) => Lex.left _ _ (refl _) +#align prod.lex.refl_left Prod.Lex.refl_left instance {r : α → α → Prop} {s : β → β → Prop} [IsRefl α r] : IsRefl (α × β) (Prod.Lex r s) := ⟨Lex.refl_left _ _⟩ @@ -208,6 +250,7 @@ instance {r : α → α → Prop} {s : β → β → Prop} [IsRefl α r] : IsRef @[refl] theorem Lex.refl_right (r : α → α → Prop) (s : β → β → Prop) [IsRefl β s] : ∀ x, Prod.Lex r s x x | (_, _) => Lex.right _ (refl _) +#align prod.lex.refl_right Prod.Lex.refl_right instance {r : α → α → Prop} {s : β → β → Prop} [IsRefl β s] : IsRefl (α × β) (Prod.Lex r s) := ⟨Lex.refl_right _ _⟩ @@ -222,6 +265,7 @@ theorem Lex.trans {r : α → α → Prop} {s : β → β → Prop} [IsTrans α | (_, _), (_, _), (_, _), left _ _ hxy₁, right _ _ => left _ _ hxy₁ | (_, _), (_, _), (_, _), right _ _, left _ _ hyz₁ => left _ _ hyz₁ | (_, _), (_, _), (_, _), right _ hxy₂, right _ hyz₂ => right _ (_root_.trans hxy₂ hyz₂) +#align prod.lex.trans Prod.Lex.trans instance {r : α → α → Prop} {s : β → β → Prop} [IsTrans α r] [IsTrans β s] : IsTrans (α × β) (Prod.Lex r s) := diff --git a/Mathlib/Data/Prod/PProd.lean b/Mathlib/Data/Prod/PProd.lean index 5f82d0ee1873a..8127ea342d4ad 100644 --- a/Mathlib/Data/Prod/PProd.lean +++ b/Mathlib/Data/Prod/PProd.lean @@ -24,20 +24,25 @@ namespace PProd @[simp] theorem mk.eta {p : PProd α β} : PProd.mk p.1 p.2 = p := PProd.casesOn p fun _ _ ↦ rfl +#align pprod.mk.eta PProd.mk.eta @[simp] theorem «forall» {p : PProd α β → Prop} : (∀ x, p x) ↔ ∀ a b, p ⟨a, b⟩ := ⟨fun h a b ↦ h ⟨a, b⟩, fun h ⟨a, b⟩ ↦ h a b⟩ +#align pprod.forall PProd.forall @[simp] theorem «exists» {p : PProd α β → Prop} : (∃ x, p x) ↔ ∃ a b, p ⟨a, b⟩ := ⟨fun ⟨⟨a, b⟩, h⟩ ↦ ⟨a, b, h⟩, fun ⟨a, b, h⟩ ↦ ⟨⟨a, b⟩, h⟩⟩ +#align pprod.exists PProd.exists theorem forall' {p : α → β → Prop} : (∀ x : PProd α β, p x.1 x.2) ↔ ∀ a b, p a b := PProd.forall +#align pprod.forall' PProd.forall' theorem exists' {p : α → β → Prop} : (∃ x : PProd α β, p x.1 x.2) ↔ ∃ a b, p a b := PProd.exists +#align pprod.exists' PProd.exists' end PProd @@ -46,3 +51,4 @@ theorem Function.Injective.pprod_map {f : α → β} {g : γ → δ} (hf : Injec have A := congr_arg PProd.fst h have B := congr_arg PProd.snd h congr_arg₂ PProd.mk (hf A) (hg B) +#align function.injective.pprod_map Function.Injective.pprod_map diff --git a/Mathlib/Data/Quot.lean b/Mathlib/Data/Quot.lean index d71a9045062aa..b17c511f06b2e 100644 --- a/Mathlib/Data/Quot.lean +++ b/Mathlib/Data/Quot.lean @@ -31,6 +31,7 @@ theorem ext {α : Sort _} : ∀ {s t : Setoid α}, by have : r = p := funext fun a ↦ funext fun b ↦ propext <| Eq a b subst this rfl +#align setoid.ext Setoid.ext end Setoid @@ -68,24 +69,29 @@ protected def hrecOn₂ (qa : Quot ra) (qb : Quot rb) (f : ∀ a b, φ ⟦a⟧ have h₂ : HEq (f a₂ b) (@Quot.hrecOn _ _ (φ _) ⟦b⟧ (f a₂) (@cb _)) := by simp [heq_self_iff_true] (h₁.trans (ca pa)).trans h₂ +#align quot.hrec_on₂ Quot.hrecOn₂ /-- Map a function `f : α → β` such that `ra x y` implies `rb (f x) (f y)` to a map `Quot ra → Quot rb`. -/ protected def map (f : α → β) (h : (ra ⇒ rb) f f) : Quot ra → Quot rb := (Quot.lift fun x ↦ ⟦f x⟧) fun x y (h₁ : ra x y) ↦ Quot.sound <| h h₁ +#align quot.map Quot.map /-- If `ra` is a subrelation of `ra'`, then we have a natural map `Quot ra → Quot ra'`. -/ protected def mapRight {ra' : α → α → Prop} (h : ∀ a₁ a₂, ra a₁ a₂ → ra' a₁ a₂) : Quot ra → Quot ra' := Quot.map id h +#align quot.map_right Quot.mapRight /-- Weaken the relation of a quotient. This is the same as `Quot.map id`. -/ def factor {α : Type _} (r s : α → α → Prop) (h : ∀ x y, r x y → s x y) : Quot r → Quot s := Quot.lift (Quot.mk s) fun x y rxy ↦ Quot.sound (h x y rxy) +#align quot.factor Quot.factor theorem factor_mk_eq {α : Type _} (r s : α → α → Prop) (h : ∀ x y, r x y → s x y) : factor r s h ∘ Quot.mk _ = Quot.mk _ := rfl +#align quot.factor_mk_eq Quot.factor_mk_eq variable {γ : Sort _} {r : α → α → Prop} {s : β → β → Prop} @@ -94,6 +100,7 @@ theorem lift_mk (f : α → γ) (h : ∀ a₁ a₂, r a₁ a₂ → f a₁ = f a Quot.lift f h (Quot.mk r a) = f a := rfl #align quot.lift_beta Quot.lift_mk +#align quot.lift_mk Quot.lift_mk theorem liftOn_mk (a : α) (f : α → γ) (h : ∀ a₁ a₂, r a₁ a₂ → f a₁ = f a₂) : Quot.liftOn (Quot.mk r a) f h = f a := @@ -112,12 +119,14 @@ protected def lift₂ (f : α → β → γ) (hr : ∀ a b₁ b₂, s b₁ b₂ (hs : ∀ a₁ a₂ b, r a₁ a₂ → f a₁ b = f a₂ b) (q₁ : Quot r) (q₂ : Quot s) : γ := Quot.lift (fun a ↦ Quot.lift (f a) (hr a)) (fun a₁ a₂ ha ↦ funext fun q ↦ Quot.induction_on q fun b ↦ hs a₁ a₂ b ha) q₁ q₂ +#align quot.lift₂ Quot.lift₂ @[simp] theorem lift₂_mk (f : α → β → γ) (hr : ∀ a b₁ b₂, s b₁ b₂ → f a b₁ = f a b₂) (hs : ∀ a₁ a₂ b, r a₁ a₂ → f a₁ b = f a₂ b) (a : α) (b : β) : Quot.lift₂ f hr hs (Quot.mk r a) (Quot.mk s b) = f a b := rfl +#align quot.lift₂_mk Quot.lift₂_mk /-- Descends a function `f : α → β → γ` to quotients of `α` and `β` and applies it. -/ -- porting note: removed `@[elab_as_elim]`, gave "unexpected resulting type γ" @@ -125,6 +134,7 @@ theorem lift₂_mk (f : α → β → γ) (hr : ∀ a b₁ b₂, s b₁ b₂ → protected def liftOn₂ (p : Quot r) (q : Quot s) (f : α → β → γ) (hr : ∀ a b₁ b₂, s b₁ b₂ → f a b₁ = f a b₂) (hs : ∀ a₁ a₂ b, r a₁ a₂ → f a₁ b = f a₂ b) : γ := Quot.lift₂ f hr hs p q +#align quot.lift_on₂ Quot.liftOn₂ @[simp] theorem liftOn₂_mk (a : α) (b : β) (f : α → β → γ) (hr : ∀ a b₁ b₂, s b₁ b₂ → f a b₁ = f a b₂) @@ -141,12 +151,14 @@ protected def map₂ (f : α → β → γ) (hr : ∀ a b₁ b₂, s b₁ b₂ (hs : ∀ a₁ a₂ b, r a₁ a₂ → t (f a₁ b) (f a₂ b)) (q₁ : Quot r) (q₂ : Quot s) : Quot t := Quot.lift₂ (fun a b ↦ Quot.mk t <| f a b) (fun a b₁ b₂ hb ↦ Quot.sound (hr a b₁ b₂ hb)) (fun a₁ a₂ b ha ↦ Quot.sound (hs a₁ a₂ b ha)) q₁ q₂ +#align quot.map₂ Quot.map₂ @[simp] theorem map₂_mk (f : α → β → γ) (hr : ∀ a b₁ b₂, s b₁ b₂ → t (f a b₁) (f a b₂)) (hs : ∀ a₁ a₂ b, r a₁ a₂ → t (f a₁ b) (f a₂ b)) (a : α) (b : β) : Quot.map₂ f hr hs (Quot.mk r a) (Quot.mk s b) = Quot.mk t (f a b) := rfl +#align quot.map₂_mk Quot.map₂_mk /-- A binary version of `Quot.recOnSubsingleton`. -/ -- porting note: removed `@[reducible]` because it caused extremely slow `simp` @@ -157,11 +169,13 @@ protected def recOnSubsingleton₂ {φ : Quot r → Quot s → Sort _} @Quot.recOnSubsingleton _ r (fun q ↦ φ q q₂) (fun a ↦ Quot.ind (β := λ b => Subsingleton (φ (mk r a) b)) (h a) q₂) q₁ fun a ↦ Quot.recOnSubsingleton q₂ fun b ↦ f a b +#align quot.rec_on_subsingleton₂ Quot.recOnSubsingleton₂ @[elab_as_elim] protected theorem induction_on₂ {δ : Quot r → Quot s → Prop} (q₁ : Quot r) (q₂ : Quot s) (h : ∀ a b, δ (Quot.mk r a) (Quot.mk s b)) : δ q₁ q₂ := Quot.ind (β := λ a => δ a q₂) (fun a₁ ↦ Quot.ind (fun a₂ ↦ h a₁ a₂) q₂) q₁ +#align quot.induction_on₂ Quot.induction_on₂ @[elab_as_elim] protected theorem induction_on₃ {δ : Quot r → Quot s → Quot t → Prop} (q₁ : Quot r) @@ -169,6 +183,7 @@ protected theorem induction_on₃ {δ : Quot r → Quot s → Quot t → Prop} ( δ q₁ q₂ q₃ := Quot.ind (β := λ a => δ a q₂ q₃) (fun a₁ ↦ Quot.ind (β := λ b => δ _ b q₃) (fun a₂ ↦ Quot.ind (fun a₃ ↦ h a₁ a₂ a₃) q₃) q₂) q₁ +#align quot.induction_on₃ Quot.induction_on₃ instance lift.decidablePred (r : α → α → Prop) (f : α → Prop) (h : ∀ a b, r a b → f a = f b) [hf : DecidablePred f] : @@ -219,16 +234,19 @@ instance {α : Type _} [Setoid α] : IsEquiv α (· ≈ ·) where protected def hrecOn₂ (qa : Quotient sa) (qb : Quotient sb) (f : ∀ a b, φ ⟦a⟧ ⟦b⟧) (c : ∀ a₁ b₁ a₂ b₂, a₁ ≈ a₂ → b₁ ≈ b₂ → HEq (f a₁ b₁) (f a₂ b₂)) : φ qa qb := Quot.hrecOn₂ qa qb f (fun p ↦ c _ _ _ _ p (Setoid.refl _)) fun p ↦ c _ _ _ _ (Setoid.refl _) p +#align quotient.hrec_on₂ Quotient.hrecOn₂ /-- Map a function `f : α → β` that sends equivalent elements to equivalent elements to a function `Quotient sa → Quotient sb`. Useful to define unary operations on quotients. -/ protected def map (f : α → β) (h : ((· ≈ ·) ⇒ (· ≈ ·)) f f) : Quotient sa → Quotient sb := Quot.map f h +#align quotient.map Quotient.map @[simp] theorem map_mk (f : α → β) (h : ((· ≈ ·) ⇒ (· ≈ ·)) f f) (x : α) : Quotient.map f h (⟦x⟧ : Quotient sa) = (⟦f x⟧ : Quotient sb) := rfl +#align quotient.map_mk Quotient.map_mk variable {γ : Sort _} [sc : Setoid γ] @@ -238,11 +256,13 @@ Useful to define binary operations on quotients. -/ protected def map₂ (f : α → β → γ) (h : ((· ≈ ·) ⇒ (· ≈ ·) ⇒ (· ≈ ·)) f f) : Quotient sa → Quotient sb → Quotient sc := Quotient.lift₂ (fun x y ↦ ⟦f x y⟧) fun _ _ _ _ h₁ h₂ ↦ Quot.sound <| h h₁ h₂ +#align quotient.map₂ Quotient.map₂ @[simp] theorem map₂_mk (f : α → β → γ) (h : ((· ≈ ·) ⇒ (· ≈ ·) ⇒ (· ≈ ·)) f f) (x : α) (y : β) : Quotient.map₂ f h (⟦x⟧ : Quotient sa) (⟦y⟧ : Quotient sb) = (⟦f x y⟧ : Quotient sc) := rfl +#align quotient.map₂_mk Quotient.map₂_mk instance lift.decidablePred (f : α → Prop) (h : ∀ a b, a ≈ b → f a = f b) [DecidablePred f] : DecidablePred (Quotient.lift f h) := @@ -269,6 +289,7 @@ end Quotient theorem Quot.eq {α : Type _} {r : α → α → Prop} {x y : α} : Quot.mk r x = Quot.mk r y ↔ EqvGen r x y := ⟨Quot.exact r, Quot.EqvGen_sound⟩ +#align quot.eq Quot.eq @[simp] theorem Quotient.eq [r : Setoid α] {x y : α} : Quotient.mk r x = ⟦y⟧ ↔ x ≈ y := @@ -277,16 +298,19 @@ theorem Quotient.eq [r : Setoid α] {x y : α} : Quotient.mk r x = ⟦y⟧ ↔ x theorem forall_quotient_iff {α : Type _} [r : Setoid α] {p : Quotient r → Prop} : (∀ a : Quotient r, p a) ↔ ∀ a : α, p ⟦a⟧ := ⟨fun h _ ↦ h _, fun h a ↦ a.induction_on h⟩ +#align forall_quotient_iff forall_quotient_iff @[simp] theorem Quotient.lift_mk [s : Setoid α] (f : α → β) (h : ∀ a b : α, a ≈ b → f a = f b) (x : α) : Quotient.lift f h (Quotient.mk s x) = f x := rfl +#align quotient.lift_mk Quotient.lift_mk @[simp] theorem Quotient.lift_comp_mk [Setoid α] (f : α → β) (h : ∀ a b : α, a ≈ b → f a = f b) : Quotient.lift f h ∘ Quotient.mk _ = f := rfl +#align quotient.lift_comp_mk Quotient.lift_comp_mk @[simp] theorem Quotient.lift₂_mk {α : Sort _} {β : Sort _} {γ : Sort _} [Setoid α] [Setoid β] @@ -295,6 +319,7 @@ theorem Quotient.lift₂_mk {α : Sort _} {β : Sort _} {γ : Sort _} [Setoid α (a : α) (b : β) : Quotient.lift₂ f h (Quotient.mk _ a) (Quotient.mk _ b) = f a b := rfl +#align quotient.lift₂_mk Quotient.lift₂_mk theorem Quotient.liftOn_mk [s : Setoid α] (f : α → β) (h : ∀ a b : α, a ≈ b → f a = f b) (x : α) : Quotient.liftOn (Quotient.mk s x) f h = f x := @@ -311,16 +336,19 @@ theorem Quotient.liftOn₂_mk {α : Sort _} {β : Sort _} [Setoid α] (f : α /-- `Quot.mk r` is a surjective function. -/ theorem surjective_quot_mk (r : α → α → Prop) : Function.Surjective (Quot.mk r) := Quot.exists_rep +#align surjective_quot_mk surjective_quot_mk /-- `Quotient.mk` is a surjective function. -/ theorem surjective_quotient_mk (α : Sort _) [s : Setoid α] : Function.Surjective (Quotient.mk _ : α → Quotient s) := Quot.exists_rep +#align surjective_quotient_mk surjective_quotient_mk /-- Choose an element of the equivalence class using the axiom of choice. Sound but noncomputable. -/ noncomputable def Quot.out {r : α → α → Prop} (q : Quot r) : α := Classical.choose (Quot.exists_rep q) +#align quot.out Quot.out /-- Unwrap the VM representation of a quotient to obtain an element of the equivalence class. Computable but unsound. -/ @@ -330,39 +358,48 @@ unsafe def Quot.unquot {r : α → α → Prop} : Quot r → α := @[simp] theorem Quot.out_eq {r : α → α → Prop} (q : Quot r) : Quot.mk r q.out = q := Classical.choose_spec (Quot.exists_rep q) +#align quot.out_eq Quot.out_eq /-- Choose an element of the equivalence class using the axiom of choice. Sound but noncomputable. -/ noncomputable def Quotient.out [s : Setoid α] : Quotient s → α := Quot.out +#align quotient.out Quotient.out @[simp] theorem Quotient.out_eq [s : Setoid α] (q : Quotient s) : ⟦q.out⟧ = q := Quot.out_eq q +#align quotient.out_eq Quotient.out_eq theorem Quotient.mk_out [Setoid α] (a : α) : ⟦a⟧.out ≈ a := Quotient.exact (Quotient.out_eq _) +#align quotient.mk_out Quotient.mk_out theorem Quotient.mk_eq_iff_out [s : Setoid α] {x : α} {y : Quotient s} : ⟦x⟧ = y ↔ x ≈ Quotient.out y := by refine' Iff.trans _ Quotient.eq rw [Quotient.out_eq y] +#align quotient.mk_eq_iff_out Quotient.mk_eq_iff_out theorem Quotient.eq_mk_iff_out [s : Setoid α] {x : Quotient s} {y : α} : x = ⟦y⟧ ↔ Quotient.out x ≈ y := by refine' Iff.trans _ Quotient.eq rw [Quotient.out_eq x] +#align quotient.eq_mk_iff_out Quotient.eq_mk_iff_out @[simp] theorem Quotient.out_equiv_out {s : Setoid α} {x y : Quotient s} : x.out ≈ y.out ↔ x = y := by rw [← Quotient.eq_mk_iff_out, Quotient.out_eq] +#align quotient.out_equiv_out Quotient.out_equiv_out theorem Quotient.out_injective {s : Setoid α} : Function.Injective (@Quotient.out α s) := fun _ _ h ↦ Quotient.out_equiv_out.1 <| h ▸ Setoid.refl _ +#align quotient.out_injective Quotient.out_injective @[simp] theorem Quotient.out_inj {s : Setoid α} {x y : Quotient s} : x.out = y.out ↔ x = y := ⟨fun h ↦ Quotient.out_injective h, fun h ↦ h ▸ rfl⟩ +#align quotient.out_inj Quotient.out_inj section Pi @@ -378,11 +415,13 @@ noncomputable def Quotient.choice {ι : Type _} {α : ι → Type _} [S : ∀ i, (f : ∀ i, Quotient (S i)) : @Quotient (∀ i, α i) (by infer_instance) := ⟦fun i ↦ (f i).out⟧ +#align quotient.choice Quotient.choice @[simp] theorem Quotient.choice_eq {ι : Type _} {α : ι → Type _} [∀ i, Setoid (α i)] (f : ∀ i, α i) : (Quotient.choice fun i ↦ ⟦f i⟧) = ⟦f⟧ := Quotient.sound fun _ ↦ Quotient.mk_out _ +#align quotient.choice_eq Quotient.choice_eq @[elab_as_elim] theorem Quotient.induction_on_pi {ι : Type _} {α : ι → Sort _} [s : ∀ i, Setoid (α i)] @@ -390,11 +429,13 @@ theorem Quotient.induction_on_pi {ι : Type _} {α : ι → Sort _} [s : ∀ i, (h : ∀ a : ∀ i, α i, p fun i ↦ ⟦a i⟧) : p f := by rw [← (funext fun i ↦ Quotient.out_eq (f i) : (fun i ↦ ⟦(f i).out⟧) = f)] apply h +#align quotient.induction_on_pi Quotient.induction_on_pi end Pi theorem nonempty_quotient_iff (s : Setoid α) : Nonempty (Quotient s) ↔ Nonempty α := ⟨fun ⟨a⟩ ↦ Quotient.inductionOn a Nonempty.intro, fun ⟨a⟩ ↦ ⟨⟦a⟧⟩⟩ +#align nonempty_quotient_iff nonempty_quotient_iff /-! ### Truncation -/ @@ -406,15 +447,18 @@ theorem nonempty_quotient_iff (s : Setoid α) : Nonempty (Quotient s) ↔ Nonemp maintain computability. -/ def Trunc.{u} (α : Sort u) : Sort u := @Quot α fun _ _ ↦ True +#align trunc Trunc theorem true_equivalence : @Equivalence α fun _ _ ↦ True := ⟨fun _ ↦ trivial, fun _ ↦ trivial, fun _ _ ↦ trivial⟩ +#align true_equivalence true_equivalence namespace Trunc /-- Constructor for `Trunc α` -/ def mk (a : α) : Trunc α := Quot.mk _ a +#align trunc.mk Trunc.mk instance [Inhabited α] : Inhabited (Trunc α) := ⟨mk default⟩ @@ -422,33 +466,41 @@ instance [Inhabited α] : Inhabited (Trunc α) := /-- Any constant function lifts to a function out of the truncation -/ def lift (f : α → β) (c : ∀ a b : α, f a = f b) : Trunc α → β := Quot.lift f fun a b _ ↦ c a b +#align trunc.lift Trunc.lift theorem ind {β : Trunc α → Prop} : (∀ a : α, β (mk a)) → ∀ q : Trunc α, β q := Quot.ind +#align trunc.ind Trunc.ind protected theorem lift_mk (f : α → β) (c) (a : α) : lift f c (mk a) = f a := rfl +#align trunc.lift_mk Trunc.lift_mk /-- Lift a constant function on `q : Trunc α`. -/ -- porting note: removed `@[elab_as_elim]` because it gave "unexpected eliminator resulting type" -- porting note: removed `@[reducible]` because it caused extremely slow `simp` protected def liftOn (q : Trunc α) (f : α → β) (c : ∀ a b : α, f a = f b) : β := lift f c q +#align trunc.lift_on Trunc.liftOn @[elab_as_elim] protected theorem induction_on {β : Trunc α → Prop} (q : Trunc α) (h : ∀ a, β (mk a)) : β q := ind h q +#align trunc.induction_on Trunc.induction_on theorem exists_rep (q : Trunc α) : ∃ a : α, mk a = q := Quot.exists_rep q +#align trunc.exists_rep Trunc.exists_rep @[elab_as_elim] protected theorem induction_on₂ {C : Trunc α → Trunc β → Prop} (q₁ : Trunc α) (q₂ : Trunc β) (h : ∀ a b, C (mk a) (mk b)) : C q₁ q₂ := Trunc.induction_on q₁ fun a₁ ↦ Trunc.induction_on q₂ (h a₁) +#align trunc.induction_on₂ Trunc.induction_on₂ protected theorem eq (a b : Trunc α) : a = b := Trunc.induction_on₂ a b fun _ _ ↦ Quot.sound trivial +#align trunc.eq Trunc.eq instance : Subsingleton (Trunc α) := ⟨Trunc.eq⟩ @@ -456,10 +508,12 @@ instance : Subsingleton (Trunc α) := /-- The `bind` operator for the `Trunc` monad. -/ def bind (q : Trunc α) (f : α → Trunc β) : Trunc β := Trunc.liftOn q f fun _ _ ↦ Trunc.eq _ _ +#align trunc.bind Trunc.bind /-- A function `f : α → β` defines a function `map f : Trunc α → Trunc β`. -/ def map (f : α → β) (q : Trunc α) : Trunc β := bind q (Trunc.mk ∘ f) +#align trunc.map Trunc.map instance : Monad Trunc where pure := @Trunc.mk @@ -486,6 +540,7 @@ protected def rec (f : ∀ a, C (mk a)) (h : ∀ a b : α, (Eq.ndrec (f a) (Trunc.eq (mk a) (mk b)) : C (mk b)) = f b) (q : Trunc α) : C q := Quot.rec f (fun a b _ ↦ h a b) q +#align trunc.rec Trunc.rec /-- A version of `Trunc.rec` taking `q : Trunc α` as the first argument. -/ -- porting note: removed `@[reducible]` because it caused extremely slow `simp` @@ -493,6 +548,7 @@ protected def rec (f : ∀ a, C (mk a)) protected def recOn (q : Trunc α) (f : ∀ a, C (mk a)) (h : ∀ a b : α, (Eq.ndrec (f a) (Trunc.eq (mk a) (mk b)) : C (mk b)) = f b) : C q := Trunc.rec f h q +#align trunc.rec_on Trunc.recOn /-- A version of `Trunc.recOn` assuming the codomain is a `Subsingleton`. -/ -- porting note: removed `@[reducible]` because it caused extremely slow `simp` @@ -500,17 +556,21 @@ protected def recOn (q : Trunc α) (f : ∀ a, C (mk a)) protected def recOnSubsingleton [∀ a, Subsingleton (C (mk a))] (q : Trunc α) (f : ∀ a, C (mk a)) : C q := Trunc.rec f (fun _ b ↦ Subsingleton.elim _ (f b)) q +#align trunc.rec_on_subsingleton Trunc.recOnSubsingleton /-- Noncomputably extract a representative of `Trunc α` (using the axiom of choice). -/ noncomputable def out : Trunc α → α := Quot.out +#align trunc.out Trunc.out @[simp] theorem out_eq (q : Trunc α) : mk q.out = q := Trunc.eq _ _ +#align trunc.out_eq Trunc.out_eq protected theorem nonempty (q : Trunc α) : Nonempty α := nonempty_of_exists q.exists_rep +#align trunc.nonempty Trunc.nonempty end Trunc @@ -547,6 +607,7 @@ instance argument. -/ protected def liftOn' (q : Quotient s₁) (f : α → φ) (h : ∀ a b, @Setoid.r α s₁ a b → f a = f b) : φ := Quotient.liftOn q f h +#align quotient.lift_on' Quotient.liftOn' @[simp] protected theorem liftOn'_mk'' (f : α → φ) (h) (x : α) : @@ -565,6 +626,7 @@ instead of instance arguments. -/ protected def liftOn₂' (q₁ : Quotient s₁) (q₂ : Quotient s₂) (f : α → β → γ) (h : ∀ a₁ a₂ b₁ b₂, @Setoid.r α s₁ a₁ b₁ → @Setoid.r β s₂ a₂ b₂ → f a₁ a₂ = f b₁ b₂) : γ := Quotient.liftOn₂ q₁ q₂ f h +#align quotient.lift_on₂' Quotient.liftOn₂' @[simp] protected theorem liftOn₂'_mk'' (f : α → β → γ) (h) (a : α) (b : β) : @@ -577,6 +639,7 @@ instance argument. -/ protected theorem ind' {p : Quotient s₁ → Prop} (h : ∀ a, p (Quotient.mk'' a)) (q : Quotient s₁) : p q := Quotient.ind h q +#align quotient.ind' Quotient.ind' /-- A version of `Quotient.ind₂` taking `{s₁ : Setoid α} {s₂ : Setoid β}` as implicit arguments instead of instance arguments. -/ @@ -585,6 +648,7 @@ protected theorem ind₂' {p : Quotient s₁ → Quotient s₂ → Prop} (h : ∀ a₁ a₂, p (Quotient.mk'' a₁) (Quotient.mk'' a₂)) (q₁ : Quotient s₁) (q₂ : Quotient s₂) : p q₁ q₂ := Quotient.ind₂ h q₁ q₂ +#align quotient.ind₂' Quotient.ind₂' /-- A version of `Quotient.inductionOn` taking `{s : Setoid α}` as an implicit argument instead of an instance argument. -/ @@ -620,6 +684,7 @@ protected def recOnSubsingleton' {φ : Quotient s₁ → Sort _} [∀ a, Subsing (q : Quotient s₁) (f : ∀ a, φ (Quotient.mk'' a)) : φ q := Quotient.recOnSubsingleton q f +#align quotient.rec_on_subsingleton' Quotient.recOnSubsingleton' /-- A version of `Quotient.recOnSubsingleton₂` taking `{s₁ : Setoid α} {s₂ : Setoid α}` as implicit arguments instead of instance arguments. -/ @@ -630,11 +695,13 @@ protected def recOnSubsingleton₂' {φ : Quotient s₁ → Quotient s₂ → So (q₁ : Quotient s₁) (q₂ : Quotient s₂) (f : ∀ a₁ a₂, φ (Quotient.mk'' a₁) (Quotient.mk'' a₂)) : φ q₁ q₂ := Quotient.recOnSubsingleton₂ q₁ q₂ f +#align quotient.rec_on_subsingleton₂' Quotient.recOnSubsingleton₂' /-- Recursion on a `Quotient` argument `a`, result type depends on `⟦a⟧`. -/ protected def hrecOn' {φ : Quotient s₁ → Sort _} (qa : Quotient s₁) (f : ∀ a, φ (Quotient.mk'' a)) (c : ∀ a₁ a₂, a₁ ≈ a₂ → HEq (f a₁) (f a₂)) : φ qa := Quot.hrecOn qa f c +#align quotient.hrec_on' Quotient.hrecOn' @[simp] theorem hrecOn'_mk'' {φ : Quotient s₁ → Sort _} (f : ∀ a, φ (Quotient.mk'' a)) @@ -649,6 +716,7 @@ protected def hrecOn₂' {φ : Quotient s₁ → Quotient s₂ → Sort _} (qa : (c : ∀ a₁ b₁ a₂ b₂, a₁ ≈ a₂ → b₁ ≈ b₂ → HEq (f a₁ b₁) (f a₂ b₂)) : φ qa qb := Quotient.hrecOn₂ qa qb f c +#align quotient.hrec_on₂' Quotient.hrecOn₂' @[simp] theorem hrecOn₂'_mk'' {φ : Quotient s₁ → Quotient s₂ → Sort _} @@ -662,6 +730,7 @@ theorem hrecOn₂'_mk'' {φ : Quotient s₁ → Quotient s₂ → Sort _} to a function `Quotient sa → Quotient sb`. Useful to define unary operations on quotients. -/ protected def map' (f : α → β) (h : (s₁.r ⇒ s₂.r) f f) : Quotient s₁ → Quotient s₂ := Quot.map f h +#align quotient.map' Quotient.map' @[simp] theorem map'_mk'' (f : α → β) (h) (x : α) : @@ -672,6 +741,7 @@ theorem map'_mk'' (f : α → β) (h) (x : α) : protected def map₂' (f : α → β → γ) (h : (s₁.r ⇒ s₂.r ⇒ s₃.r) f f) : Quotient s₁ → Quotient s₂ → Quotient s₃ := Quotient.map₂ f h +#align quotient.map₂' Quotient.map₂' @[simp] theorem map₂'_mk'' (f : α → β → γ) (h) (x : α) : @@ -682,9 +752,11 @@ theorem map₂'_mk'' (f : α → β → γ) (h) (x : α) : theorem exact' {a b : α} : (Quotient.mk'' a : Quotient s₁) = Quotient.mk'' b → @Setoid.r _ s₁ a b := Quotient.exact +#align quotient.exact' Quotient.exact' theorem sound' {a b : α} : @Setoid.r _ s₁ a b → @Quotient.mk'' α s₁ a = Quotient.mk'' b := Quotient.sound +#align quotient.sound' Quotient.sound' @[simp] protected theorem eq' [s₁ : Setoid α] {a b : α} : @@ -701,13 +773,16 @@ protected theorem eq'' {a b : α} : @Quotient.mk'' α s₁ a = Quotient.mk'' b instance argument. -/ noncomputable def out' (a : Quotient s₁) : α := Quotient.out a +#align quotient.out' Quotient.out' @[simp] theorem out_eq' (q : Quotient s₁) : Quotient.mk'' q.out' = q := q.out_eq +#align quotient.out_eq' Quotient.out_eq' theorem mk_out' (a : α) : @Setoid.r α s₁ (Quotient.mk'' a : Quotient s₁).out' a := Quotient.exact (Quotient.out_eq _) +#align quotient.mk_out' Quotient.mk_out' section @@ -731,6 +806,7 @@ protected theorem liftOn₂'_mk [t : Setoid β] (f : α → β → γ) (h) (a : theorem map'_mk [t : Setoid β] (f : α → β) (h) (x : α) : (Quotient.mk s x).map' f h = (Quotient.mk t (f x)) := rfl +#align quotient.map'_mk Quotient.map'_mk end diff --git a/Mathlib/Data/Rat/Cast.lean b/Mathlib/Data/Rat/Cast.lean index d9b26c03bc772..efa9de7f3c4bd 100644 --- a/Mathlib/Data/Rat/Cast.lean +++ b/Mathlib/Data/Rat/Cast.lean @@ -323,6 +323,7 @@ theorem cast_mono : Monotone ((↑) : ℚ → K) := def castOrderEmbedding : ℚ ↪o K := OrderEmbedding.ofStrictMono (↑) cast_strictMono #align rat.cast_order_embedding Rat.castOrderEmbedding +#align rat.cast_order_embedding_apply Rat.castOrderEmbedding_apply @[simp, norm_cast] theorem cast_le {m n : ℚ} : (m : K) ≤ n ↔ m ≤ n := diff --git a/Mathlib/Data/Real/Basic.lean b/Mathlib/Data/Real/Basic.lean index 1099d6fa952de..f3d2f30ddf0cc 100644 --- a/Mathlib/Data/Real/Basic.lean +++ b/Mathlib/Data/Real/Basic.lean @@ -243,6 +243,10 @@ def ringEquivCauchy : ℝ ≃+* CauSeq.Completion.Cauchy (abs : ℚ → ℚ) := map_mul' := cauchy_mul } set_option linter.uppercaseLean3 false in #align real.ring_equiv_Cauchy Real.ringEquivCauchy +set_option linter.uppercaseLean3 false in +#align real.ring_equiv_Cauchy_apply Real.ringEquivCauchy_apply +set_option linter.uppercaseLean3 false in +#align real.ring_equiv_Cauchy_symm_apply_cauchy Real.ringEquivCauchy_symmApply_cauchy /-! Extra instances to short-circuit type class resolution. diff --git a/Mathlib/Data/Real/CauSeqCompletion.lean b/Mathlib/Data/Real/CauSeqCompletion.lean index 470a62f943893..5276d0334ea6d 100644 --- a/Mathlib/Data/Real/CauSeqCompletion.lean +++ b/Mathlib/Data/Real/CauSeqCompletion.lean @@ -179,6 +179,7 @@ def ofRatRingHom : β →+* (Cauchy abv) where map_add' := ofRat_add map_mul' := ofRat_mul #align cau_seq.completion.of_rat_ring_hom CauSeq.Completion.ofRatRingHom +#align cau_seq.completion.of_rat_ring_hom_apply CauSeq.Completion.ofRatRingHom_apply theorem ofRat_sub (x y : β) : ofRat (x - y) = (ofRat x - ofRat y : Cauchy abv) := congr_arg mk (const_sub _ _) diff --git a/Mathlib/Data/Set/Basic.lean b/Mathlib/Data/Set/Basic.lean index 60ec642bcda07..d2f1ef1827b49 100644 --- a/Mathlib/Data/Set/Basic.lean +++ b/Mathlib/Data/Set/Basic.lean @@ -134,8 +134,10 @@ theorem lt_iff_ssubset : s < t ↔ s ⊂ t := #align set.lt_iff_ssubset Set.lt_iff_ssubset alias le_iff_subset ↔ _root_.LE.le.subset _root_.HasSubset.Subset.le +#align has_subset.subset.le HasSubset.Subset.le alias lt_iff_ssubset ↔ _root_.LT.lt.ssubset _root_.HasSSubset.SSubset.lt +#align has_ssubset.ssubset.lt HasSSubset.SSubset.lt -- Porting note: I've introduced this abbreviation, with the `@[coe]` attribute, -- so that `norm_cast` has something to index on. @@ -446,6 +448,7 @@ theorem nonempty_coe_sort {s : Set α} : Nonempty (↥s) ↔ s.Nonempty := #align set.nonempty_coe_sort Set.nonempty_coe_sort alias nonempty_coe_sort ↔ _ Nonempty.coe_sort +#align set.nonempty.coe_sort Set.Nonempty.coe_sort theorem nonempty_def : s.Nonempty ↔ ∃ x, x ∈ s := Iff.rfl @@ -605,6 +608,7 @@ theorem nonempty_iff_ne_empty : s.Nonempty ↔ s ≠ ∅ := #align set.nonempty_iff_ne_empty Set.nonempty_iff_ne_empty alias nonempty_iff_ne_empty ↔ Nonempty.ne_empty _ +#align set.nonempty.ne_empty Set.Nonempty.ne_empty @[simp] theorem not_nonempty_empty : ¬(∅ : Set α).Nonempty := fun ⟨_, hx⟩ => hx @@ -637,6 +641,7 @@ theorem empty_ssubset : ∅ ⊂ s ↔ s.Nonempty := #align set.empty_ssubset Set.empty_ssubset alias empty_ssubset ↔ _ Nonempty.empty_ssubset +#align set.nonempty.empty_ssubset Set.Nonempty.empty_ssubset /-! @@ -677,6 +682,7 @@ theorem univ_subset_iff {s : Set α} : univ ⊆ s ↔ s = univ := #align set.univ_subset_iff Set.univ_subset_iff alias univ_subset_iff ↔ eq_univ_of_univ_subset _ +#align set.eq_univ_of_univ_subset Set.eq_univ_of_univ_subset theorem eq_univ_iff_forall {s : Set α} : s = univ ↔ ∀ x, x ∈ s := univ_subset_iff.symm.trans <| forall_congr' fun _ => imp_iff_right trivial @@ -1517,6 +1523,7 @@ lemma not_disjoint_iff_nonempty_inter : ¬ Disjoint s t ↔ (s ∩ t).Nonempty : #align set.not_disjoint_iff_nonempty_inter Set.not_disjoint_iff_nonempty_inter alias not_disjoint_iff_nonempty_inter ↔ _ Nonempty.not_disjoint +#align set.nonempty.not_disjoint Set.Nonempty.not_disjoint lemma disjoint_or_nonempty_inter (s t : Set α) : Disjoint s t ∨ (s ∩ t).Nonempty := (em _).imp_right not_disjoint_iff_nonempty_inter.1 @@ -1558,6 +1565,8 @@ lemma disjoint_union_right : Disjoint s (t ∪ u) ↔ Disjoint s t ∧ Disjoint lemma disjoint_sdiff_left : Disjoint (t \ s) s := disjoint_sdiff_self_left lemma disjoint_sdiff_right : Disjoint s (t \ s) := disjoint_sdiff_self_right +#align set.disjoint_sdiff_right Set.disjoint_sdiff_right +#align set.disjoint_sdiff_left Set.disjoint_sdiff_left @[simp default+1] lemma disjoint_singleton_left : Disjoint {a} s ↔ a ∉ s := by simp [Set.disjoint_iff, subset_def] @@ -1710,12 +1719,16 @@ theorem disjoint_compl_right_iff_subset : Disjoint s (tᶜ) ↔ s ⊆ t := #align set.disjoint_compl_right_iff_subset Set.disjoint_compl_right_iff_subset alias subset_compl_iff_disjoint_right ↔ _ _root_.Disjoint.subset_compl_right +#align disjoint.subset_compl_right Disjoint.subset_compl_right alias subset_compl_iff_disjoint_left ↔ _ _root_.Disjoint.subset_compl_left +#align disjoint.subset_compl_left Disjoint.subset_compl_left alias disjoint_compl_left_iff_subset ↔ _ _root_.HasSubset.Subset.disjoint_compl_left +#align has_subset.subset.disjoint_compl_left HasSubset.Subset.disjoint_compl_left alias disjoint_compl_right_iff_subset ↔ _ _root_.HasSubset.Subset.disjoint_compl_right +#align has_subset.subset.disjoint_compl_right HasSubset.Subset.disjoint_compl_right theorem subset_union_compl_iff_inter_subset {s t u : Set α} : s ⊆ t ∪ uᶜ ↔ s ∩ u ⊆ t := (@isCompl_compl _ u _).le_sup_right_iff_inf_left_le @@ -2526,6 +2539,7 @@ theorem nontrivial_coe_sort {s : Set α} : Nontrivial s ↔ s.Nontrivial := by #align set.nontrivial_coe_sort Set.nontrivial_coe_sort alias nontrivial_coe_sort ↔ _ Nontrivial.coe_sort +#align set.nontrivial.coe_sort Set.Nontrivial.coe_sort /-- A type with a set `s` whose `coe_sort` is a nontrivial type is nontrivial. For the corresponding result for `Subtype`, see `Subtype.nontrivial_iff_exists_ne`. -/ @@ -2549,8 +2563,10 @@ theorem not_nontrivial_iff : ¬s.Nontrivial ↔ s.Subsingleton := #align set.not_nontrivial_iff Set.not_nontrivial_iff alias not_nontrivial_iff ↔ _ Subsingleton.not_nontrivial +#align set.subsingleton.not_nontrivial Set.Subsingleton.not_nontrivial alias not_subsingleton_iff ↔ _ Nontrivial.not_subsingleton +#align set.nontrivial.not_subsingleton Set.Nontrivial.not_subsingleton protected lemma subsingleton_or_nontrivial (s : Set α) : s.Subsingleton ∨ s.Nontrivial := by simp [or_iff_not_imp_right] diff --git a/Mathlib/Data/Set/Function.lean b/Mathlib/Data/Set/Function.lean index 8f9143df76a64..e40e2054f949f 100644 --- a/Mathlib/Data/Set/Function.lean +++ b/Mathlib/Data/Set/Function.lean @@ -170,6 +170,7 @@ theorem injective_codRestrict {f : ι → α} {s : Set α} (h : ∀ x, f x ∈ s #align set.injective_cod_restrict Set.injective_codRestrict alias injective_codRestrict ↔ _ _root_.Function.Injective.codRestrict +#align function.injective.cod_restrict Function.Injective.codRestrict variable {s s₁ s₂ : Set α} {t t₁ t₂ : Set β} {p : Set γ} {f f₁ f₂ f₃ : α → β} {g g₁ g₂ : β → γ} {f' f₁' f₂' : β → α} {g' : γ → β} {a : α} {b : β} @@ -240,6 +241,7 @@ theorem eqOn_range {ι : Sort _} {f : ι → α} {g₁ g₂ : α → β} : #align set.eq_on_range Set.eqOn_range alias eqOn_range ↔ eqOn.comp_eq _ +#align set.eq_on.comp_eq Set.eqOn.comp_eq /-! ### Congruence lemmas -/ @@ -552,6 +554,7 @@ variable (t f) def restrictPreimage : f ⁻¹' t → t := (Set.mapsTo_preimage f t).restrict _ _ _ #align set.restrict_preimage Set.restrictPreimage +#align set.restrict_preimage_coe Set.restrictPreimage_coe theorem range_restrictPreimage : range (t.restrictPreimage f) = Subtype.val ⁻¹' range f := by delta Set.restrictPreimage @@ -576,6 +579,9 @@ lemma restrictPreimage_bijective (hf : Bijective f) : Bijective (t.restrictPreim alias Set.restrictPreimage_injective ← _root_.Function.Injective.restrictPreimage alias Set.restrictPreimage_surjective ← _root_.Function.Surjective.restrictPreimage alias Set.restrictPreimage_bijective ← _root_.Function.Bijective.restrictPreimage +#align function.bijective.restrict_preimage Function.Bijective.restrictPreimage +#align function.surjective.restrict_preimage Function.Surjective.restrictPreimage +#align function.injective.restrict_preimage Function.Injective.restrictPreimage end @@ -609,6 +615,7 @@ theorem InjOn.ne_iff {x y} (h : InjOn f s) (hx : x ∈ s) (hy : y ∈ s) : f x #align set.inj_on.ne_iff Set.InjOn.ne_iff alias InjOn.ne_iff ↔ _ InjOn.ne +#align set.inj_on.ne Set.InjOn.ne theorem InjOn.congr (h₁ : InjOn f₁ s) (h : EqOn f₁ f₂ s) : InjOn f₂ s := fun _ hx _ hy => h hx ▸ h hy ▸ h₁ hx hy @@ -648,6 +655,7 @@ theorem injOn_of_injective (h : Injective f) (s : Set α) : InjOn f s := fun _ _ #align set.inj_on_of_injective Set.injOn_of_injective alias injOn_of_injective ← _root_.Function.Injective.injOn +#align function.injective.inj_on Function.Injective.injOn lemma injOn_id (s : Set α) : InjOn id s := injective_id.injOn _ #align set.inj_on_id Set.injOn_id @@ -677,6 +685,7 @@ theorem injOn_iff_injective : InjOn f s ↔ Injective (s.restrict f) := #align set.inj_on_iff_injective Set.injOn_iff_injective alias Set.injOn_iff_injective ↔ InjOn.injective _ +#align set.inj_on.injective Set.InjOn.injective theorem MapsTo.restrict_inj (h : MapsTo f s t) : Injective (h.restrict f s t) ↔ InjOn f s := by rw [h.restrict_eq_codRestrict, injective_codRestrict, injOn_iff_injective] @@ -999,6 +1008,7 @@ theorem bijective_iff_bijOn_univ : Bijective f ↔ BijOn f univ univ := #align set.bijective_iff_bij_on_univ Set.bijective_iff_bijOn_univ alias bijective_iff_bijOn_univ ↔ _root_.Function.Bijective.bijOn_univ _ +#align function.bijective.bij_on_univ Function.Bijective.bijOn_univ theorem BijOn.compl (hst : BijOn f s t) (hf : Bijective f) : BijOn f (sᶜ) (tᶜ) := ⟨hst.surjOn.mapsTo_compl hf.1, hf.1.injOn _, hst.mapsTo.surjOn_compl hf.2⟩ @@ -1569,6 +1579,8 @@ theorem strictMono_restrict [Preorder α] [Preorder β] {f : α → β} {s : Set #align strict_mono_restrict strictMono_restrict alias strictMono_restrict ↔ _root_.StrictMono.of_restrict _root_.StrictMonoOn.restrict +#align strict_mono.of_restrict StrictMono.of_restrict +#align strict_mono_on.restrict StrictMonoOn.restrict theorem StrictMono.codRestrict [Preorder α] [Preorder β] {f : α → β} (hf : StrictMono f) {s : Set β} (hs : ∀ x, f x ∈ s) : StrictMono (Set.codRestrict f s hs) := diff --git a/Mathlib/Data/Set/Image.lean b/Mathlib/Data/Set/Image.lean index 7cb33e9167c1c..a93fca881901c 100644 --- a/Mathlib/Data/Set/Image.lean +++ b/Mathlib/Data/Set/Image.lean @@ -667,6 +667,7 @@ theorem range_iff_surjective : range f = univ ↔ Surjective f := -- Porting note: Lean4 unfolds `Surjective` here, ruining dot notation alias range_iff_surjective ↔ _ _root_.Function.Surjective.range_eq +#align function.surjective.range_eq Function.Surjective.range_eq @[simp] theorem image_univ {f : α → β} : f '' univ = range f := by diff --git a/Mathlib/Data/Set/Opposite.lean b/Mathlib/Data/Set/Opposite.lean index 56abd170348bf..99d3d13081af8 100644 --- a/Mathlib/Data/Set/Opposite.lean +++ b/Mathlib/Data/Set/Opposite.lean @@ -27,52 +27,69 @@ namespace Set /-- The opposite of a set `s` is the set obtained by taking the opposite of each member of `s`. -/ protected def op (s : Set α) : Set αᵒᵖ := unop ⁻¹' s +#align set.op Set.op /-- The unop of a set `s` is the set obtained by taking the unop of each member of `s`. -/ protected def unop (s : Set αᵒᵖ) : Set α := op ⁻¹' s +#align set.unop Set.unop @[simp] theorem mem_op {s : Set α} {a : αᵒᵖ} : a ∈ s.op ↔ unop a ∈ s := Iff.rfl +#align set.mem_op Set.mem_op @[simp 1100] theorem op_mem_op {s : Set α} {a : α} : op a ∈ s.op ↔ a ∈ s := by rfl +#align set.op_mem_op Set.op_mem_op @[simp] theorem mem_unop {s : Set αᵒᵖ} {a : α} : a ∈ s.unop ↔ op a ∈ s := Iff.rfl +#align set.mem_unop Set.mem_unop @[simp 1100] theorem unop_mem_unop {s : Set αᵒᵖ} {a : αᵒᵖ} : unop a ∈ s.unop ↔ a ∈ s := by rfl +#align set.unop_mem_unop Set.unop_mem_unop @[simp] theorem op_unop (s : Set α) : s.op.unop = s := rfl +#align set.op_unop Set.op_unop @[simp] theorem unop_op (s : Set αᵒᵖ) : s.unop.op = s := rfl +#align set.unop_op Set.unop_op /-- The members of the opposite of a set are in bijection with the members of the set itself. -/ @[simps] def opEquiv_self (s : Set α) : s.op ≃ s := ⟨fun x ↦ ⟨unop x, x.2⟩, fun x ↦ ⟨op x, x.2⟩, fun _ ↦ rfl, fun _ ↦ rfl⟩ #align set.op_equiv_self Set.opEquiv_self +#align set.op_equiv_self_apply_coe Set.opEquiv_self_apply_coe +#align set.op_equiv_self_symm_apply_coe Set.opEquiv_self_symm_apply_coe /-- Taking opposites as an equivalence of powersets. -/ @[simps] def opEquiv : Set α ≃ Set αᵒᵖ := ⟨Set.op, Set.unop, op_unop, unop_op⟩ +#align set.op_equiv Set.opEquiv +#align set.op_equiv_symm_apply Set.opEquiv_symm_apply +#align set.op_equiv_apply Set.opEquiv_apply @[simp] theorem singleton_op (x : α) : ({x} : Set α).op = {op x} := rfl +#align set.singleton_op Set.singleton_op @[simp] theorem singleton_unop (x : αᵒᵖ) : ({x} : Set αᵒᵖ).unop = {unop x} := rfl +#align set.singleton_unop Set.singleton_unop @[simp] theorem singleton_op_unop (x : α) : ({op x} : Set αᵒᵖ).unop = {x} := rfl +#align set.singleton_op_unop Set.singleton_op_unop @[simp] theorem singleton_unop_op (x : αᵒᵖ) : ({unop x} : Set α).op = {x} := rfl +#align set.singleton_unop_op Set.singleton_unop_op end Set diff --git a/Mathlib/Data/Set/Pointwise/Basic.lean b/Mathlib/Data/Set/Pointwise/Basic.lean index fd28f417b10f2..7c9776b9ddaf2 100644 --- a/Mathlib/Data/Set/Pointwise/Basic.lean +++ b/Mathlib/Data/Set/Pointwise/Basic.lean @@ -1140,8 +1140,10 @@ theorem not_one_mem_div_iff : (1 : α) ∉ s / t ↔ Disjoint s t := #align set.not_zero_mem_sub_iff Set.not_zero_mem_sub_iff alias not_one_mem_div_iff ↔ _ _root_.Disjoint.one_not_mem_div_set +#align disjoint.one_not_mem_div_set Disjoint.one_not_mem_div_set attribute [to_additive] Disjoint.one_not_mem_div_set +#align disjoint.zero_not_mem_sub_set Disjoint.zero_not_mem_sub_set @[to_additive] theorem Nonempty.one_mem_div (h : s.Nonempty) : (1 : α) ∈ s / s := diff --git a/Mathlib/Data/Set/Prod.lean b/Mathlib/Data/Set/Prod.lean index 2f7a4f7a464f5..faa50878424b0 100644 --- a/Mathlib/Data/Set/Prod.lean +++ b/Mathlib/Data/Set/Prod.lean @@ -555,8 +555,10 @@ theorem offDiag_eq_empty : s.offDiag = ∅ ↔ s.Subsingleton := by #align set.off_diag_eq_empty Set.offDiag_eq_empty alias offDiag_nonempty ↔ _ Nontrivial.offDiag_nonempty +#align set.nontrivial.off_diag_nonempty Set.Nontrivial.offDiag_nonempty alias offDiag_nonempty ↔ _ Subsingleton.offDiag_eq_empty +#align set.subsingleton.off_diag_eq_empty Set.Subsingleton.offDiag_eq_empty variable (s t) diff --git a/Mathlib/Data/Setoid/Basic.lean b/Mathlib/Data/Setoid/Basic.lean index a970bcc448132..283e37cd88b80 100644 --- a/Mathlib/Data/Setoid/Basic.lean +++ b/Mathlib/Data/Setoid/Basic.lean @@ -346,6 +346,8 @@ def quotientKerEquivOfRightInverse (g : β → α) (hf : Function.RightInverse g left_inv a := Quotient.inductionOn' a fun a => Quotient.sound' <| hf (f a) right_inv := hf #align setoid.quotient_ker_equiv_of_right_inverse Setoid.quotientKerEquivOfRightInverse +#align setoid.quotient_ker_equiv_of_right_inverse_symm_apply Setoid.quotientKerEquivOfRightInverse_symm_apply +#align setoid.quotient_ker_equiv_of_right_inverse_apply Setoid.quotientKerEquivOfRightInverse_apply /-- The quotient of α by the kernel of a surjective function f bijects with f's codomain. diff --git a/Mathlib/Data/Sigma/Basic.lean b/Mathlib/Data/Sigma/Basic.lean index 294fd00576acd..ac64042c71ebc 100644 --- a/Mathlib/Data/Sigma/Basic.lean +++ b/Mathlib/Data/Sigma/Basic.lean @@ -59,44 +59,54 @@ theorem mk.inj_iff {a₁ a₂ : α} {b₁ : β a₁} {b₂ : β a₂} : Sigma.mk a₁ b₁ = ⟨a₂, b₂⟩ ↔ a₁ = a₂ ∧ HEq b₁ b₂ := ⟨λ h => by cases h; exact ⟨rfl, heq_of_eq rfl⟩, -- in Lean 3 `simp` solved this λ ⟨h₁, h₂⟩ => by subst h₁; rw [eq_of_heq h₂]⟩ +#align sigma.mk.inj_iff Sigma.mk.inj_iff @[simp] theorem eta : ∀ x : Σa, β a, Sigma.mk x.1 x.2 = x | ⟨_, _⟩ => rfl +#align sigma.eta Sigma.eta @[ext] theorem ext {x₀ x₁ : Sigma β} (h₀ : x₀.1 = x₁.1) (h₁ : HEq x₀.2 x₁.2) : x₀ = x₁ := by cases x₀; cases x₁; cases h₀; cases h₁; rfl +#align sigma.ext Sigma.ext theorem ext_iff {x₀ x₁ : Sigma β} : x₀ = x₁ ↔ x₀.1 = x₁.1 ∧ HEq x₀.2 x₁.2 := by cases x₀; cases x₁; exact Sigma.mk.inj_iff +#align sigma.ext_iff Sigma.ext_iff /-- A specialized ext lemma for equality of sigma types over an indexed subtype. -/ @[ext] theorem subtype_ext {β : Type _} {p : α → β → Prop} : ∀ {x₀ x₁ : Σa, Subtype (p a)}, x₀.fst = x₁.fst → (x₀.snd : β) = x₁.snd → x₀ = x₁ | ⟨_, _, _⟩, ⟨_, _, _⟩, rfl, rfl => rfl +#align sigma.subtype_ext Sigma.subtype_ext theorem subtype_ext_iff {β : Type _} {p : α → β → Prop} {x₀ x₁ : Σa, Subtype (p a)} : x₀ = x₁ ↔ x₀.fst = x₁.fst ∧ (x₀.snd : β) = x₁.snd := ⟨fun h ↦ h ▸ ⟨rfl, rfl⟩, fun ⟨h₁, h₂⟩ ↦ subtype_ext h₁ h₂⟩ +#align sigma.subtype_ext_iff Sigma.subtype_ext_iff @[simp] theorem «forall» {p : (Σa, β a) → Prop} : (∀ x, p x) ↔ ∀ a b, p ⟨a, b⟩ := ⟨fun h a b ↦ h ⟨a, b⟩, fun h ⟨a, b⟩ ↦ h a b⟩ +#align sigma.forall Sigma.forall @[simp] theorem «exists» {p : (Σa, β a) → Prop} : (∃ x, p x) ↔ ∃ a b, p ⟨a, b⟩ := ⟨fun ⟨⟨a, b⟩, h⟩ ↦ ⟨a, b, h⟩, fun ⟨a, b, h⟩ ↦ ⟨⟨a, b⟩, h⟩⟩ +#align sigma.exists Sigma.exists /-- Map the left and right components of a sigma -/ def map (f₁ : α₁ → α₂) (f₂ : ∀ a, β₁ a → β₂ (f₁ a)) (x : Sigma β₁) : Sigma β₂ := ⟨f₁ x.1, f₂ x.1 x.2⟩ +#align sigma.map Sigma.map end Sigma theorem sigma_mk_injective {i : α} : Function.Injective (@Sigma.mk α β i) | _, _, rfl => rfl +#align sigma_mk_injective sigma_mk_injective theorem Function.Injective.sigma_map {f₁ : α₁ → α₂} {f₂ : ∀ a, β₁ a → β₂ (f₁ a)} (h₁ : Function.Injective f₁) (h₂ : ∀ a, Function.Injective (f₂ a)) : @@ -105,48 +115,57 @@ theorem Function.Injective.sigma_map {f₁ : α₁ → α₂} {f₂ : ∀ a, β obtain rfl : i = j := h₁ (Sigma.mk.inj_iff.mp h).1 obtain rfl : x = y := h₂ i (sigma_mk_injective h) rfl +#align function.injective.sigma_map Function.Injective.sigma_map theorem Function.Injective.of_sigma_map {f₁ : α₁ → α₂} {f₂ : ∀ a, β₁ a → β₂ (f₁ a)} (h : Function.Injective (Sigma.map f₁ f₂)) (a : α₁) : Function.Injective (f₂ a) := fun x y hxy ↦ sigma_mk_injective <| @h ⟨a, x⟩ ⟨a, y⟩ (Sigma.ext rfl (heq_of_eq hxy)) +#align function.injective.of_sigma_map Function.Injective.of_sigma_map theorem Function.Injective.sigma_map_iff {f₁ : α₁ → α₂} {f₂ : ∀ a, β₁ a → β₂ (f₁ a)} (h₁ : Function.Injective f₁) : Function.Injective (Sigma.map f₁ f₂) ↔ ∀ a, Function.Injective (f₂ a) := ⟨fun h ↦ h.of_sigma_map, h₁.sigma_map⟩ +#align function.injective.sigma_map_iff Function.Injective.sigma_map_iff theorem Function.Surjective.sigma_map {f₁ : α₁ → α₂} {f₂ : ∀ a, β₁ a → β₂ (f₁ a)} (h₁ : Function.Surjective f₁) (h₂ : ∀ a, Function.Surjective (f₂ a)) : Function.Surjective (Sigma.map f₁ f₂) := by simp only [Function.Surjective, Sigma.forall, h₁.forall] exact fun i ↦ (h₂ _).forall.2 fun x ↦ ⟨⟨i, x⟩, rfl⟩ +#align function.surjective.sigma_map Function.Surjective.sigma_map /-- Interpret a function on `Σ x : α, β x` as a dependent function with two arguments. This also exists as an `Equiv` as `Equiv.piCurry γ`. -/ def Sigma.curry {γ : ∀ a, β a → Type _} (f : ∀ x : Sigma β, γ x.1 x.2) (x : α) (y : β x) : γ x y := f ⟨x, y⟩ +#align sigma.curry Sigma.curry /-- Interpret a dependent function with two arguments as a function on `Σ x : α, β x`. This also exists as an `Equiv` as `(Equiv.piCurry γ).symm`. -/ def Sigma.uncurry {γ : ∀ a, β a → Type _} (f : ∀ (x) (y : β x), γ x y) (x : Sigma β) : γ x.1 x.2 := f x.1 x.2 +#align sigma.uncurry Sigma.uncurry @[simp] theorem Sigma.uncurry_curry {γ : ∀ a, β a → Type _} (f : ∀ x : Sigma β, γ x.1 x.2) : Sigma.uncurry (Sigma.curry f) = f := funext fun ⟨_, _⟩ ↦ rfl +#align sigma.uncurry_curry Sigma.uncurry_curry @[simp] theorem Sigma.curry_uncurry {γ : ∀ a, β a → Type _} (f : ∀ (x) (y : β x), γ x y) : Sigma.curry (Sigma.uncurry f) = f := rfl +#align sigma.curry_uncurry Sigma.curry_uncurry /-- Convert a product type to a Σ-type. -/ def Prod.toSigma {α β} (p : α × β) : Σ_ : α, β := ⟨p.1, p.2⟩ +#align prod.to_sigma Prod.toSigma @[simp] theorem Prod.fst_comp_toSigma {α β} : Sigma.fst ∘ @Prod.toSigma α β = Prod.fst := @@ -179,10 +198,12 @@ variable {α : Sort _} {β : α → Sort _} /-- Nondependent eliminator for `PSigma`. -/ def elim {γ} (f : ∀ a, β a → γ) (a : PSigma β) : γ := PSigma.casesOn a f +#align psigma.elim PSigma.elim @[simp] theorem elim_val {γ} (f : ∀ a, β a → γ) (a b) : PSigma.elim f ⟨a, b⟩ = f a b := rfl +#align psigma.elim_val PSigma.elim_val instance [Inhabited α] [Inhabited (β default)] : Inhabited (PSigma β) := ⟨⟨default, default⟩⟩ @@ -206,36 +227,44 @@ theorem mk.inj_iff {a₁ a₂ : α} {b₁ : β a₁} {b₂ : β a₂} : (Iff.intro PSigma.mk.inj) fun ⟨h₁, h₂⟩ ↦ match a₁, a₂, b₁, b₂, h₁, h₂ with | _, _, _, _, Eq.refl _, HEq.refl _ => rfl +#align psigma.mk.inj_iff PSigma.mk.inj_iff @[ext] theorem ext {x₀ x₁ : PSigma β} (h₀ : x₀.1 = x₁.1) (h₁ : HEq x₀.2 x₁.2) : x₀ = x₁ := by cases x₀; cases x₁; cases h₀; cases h₁; rfl +#align psigma.ext PSigma.ext theorem ext_iff {x₀ x₁ : PSigma β} : x₀ = x₁ ↔ x₀.1 = x₁.1 ∧ HEq x₀.2 x₁.2 := by cases x₀; cases x₁; exact PSigma.mk.inj_iff +#align psigma.ext_iff PSigma.ext_iff @[simp] theorem «forall» {p : (Σ'a, β a) → Prop} : (∀ x, p x) ↔ ∀ a b, p ⟨a, b⟩ := ⟨fun h a b ↦ h ⟨a, b⟩, fun h ⟨a, b⟩ ↦ h a b⟩ +#align psigma.forall PSigma.forall @[simp] theorem «exists» {p : (Σ'a, β a) → Prop} : (∃ x, p x) ↔ ∃ a b, p ⟨a, b⟩ := ⟨fun ⟨⟨a, b⟩, h⟩ ↦ ⟨a, b, h⟩, fun ⟨a, b, h⟩ ↦ ⟨⟨a, b⟩, h⟩⟩ +#align psigma.exists PSigma.exists /-- A specialized ext lemma for equality of `PSigma` types over an indexed subtype. -/ @[ext] theorem subtype_ext {β : Sort _} {p : α → β → Prop} : ∀ {x₀ x₁ : Σ'a, Subtype (p a)}, x₀.fst = x₁.fst → (x₀.snd : β) = x₁.snd → x₀ = x₁ | ⟨_, _, _⟩, ⟨_, _, _⟩, rfl, rfl => rfl +#align psigma.subtype_ext PSigma.subtype_ext theorem subtype_ext_iff {β : Sort _} {p : α → β → Prop} {x₀ x₁ : Σ'a, Subtype (p a)} : x₀ = x₁ ↔ x₀.fst = x₁.fst ∧ (x₀.snd : β) = x₁.snd := ⟨fun h ↦ h ▸ ⟨rfl, rfl⟩, fun ⟨h₁, h₂⟩ ↦ subtype_ext h₁ h₂⟩ +#align psigma.subtype_ext_iff PSigma.subtype_ext_iff variable {α₁ : Sort _} {α₂ : Sort _} {β₁ : α₁ → Sort _} {β₂ : α₂ → Sort _} /-- Map the left and right components of a sigma -/ def map (f₁ : α₁ → α₂) (f₂ : ∀ a, β₁ a → β₂ (f₁ a)) : PSigma β₁ → PSigma β₂ | ⟨a, b⟩ => ⟨f₁ a, f₂ a b⟩ +#align psigma.map PSigma.map end PSigma diff --git a/Mathlib/Data/Subtype.lean b/Mathlib/Data/Subtype.lean index c284c71fd8efd..e7b1ce30e2175 100644 --- a/Mathlib/Data/Subtype.lean +++ b/Mathlib/Data/Subtype.lean @@ -41,55 +41,68 @@ initialize_simps_projections Subtype (val → coe) instead of `x.1`. A similar result is `Subtype.mem` in `Data.Set.Basic`. -/ theorem prop (x : Subtype p) : p x := x.2 +#align subtype.prop Subtype.prop @[simp] protected theorem «forall» {q : { a // p a } → Prop} : (∀ x, q x) ↔ ∀ a b, q ⟨a, b⟩ := ⟨fun h a b ↦ h ⟨a, b⟩, fun h ⟨a, b⟩ ↦ h a b⟩ +#align subtype.forall Subtype.forall /-- An alternative version of `Subtype.forall`. This one is useful if Lean cannot figure out `q` when using `Subtype.forall` from right to left. -/ protected theorem forall' {q : ∀ x, p x → Prop} : (∀ x h, q x h) ↔ ∀ x : { a // p a }, q x x.2 := (@Subtype.forall _ _ fun x ↦ q x.1 x.2).symm +#align subtype.forall' Subtype.forall' @[simp] protected theorem «exists» {q : { a // p a } → Prop} : (∃ x, q x) ↔ ∃ a b, q ⟨a, b⟩ := ⟨fun ⟨⟨a, b⟩, h⟩ ↦ ⟨a, b, h⟩, fun ⟨a, b, h⟩ ↦ ⟨⟨a, b⟩, h⟩⟩ +#align subtype.exists Subtype.exists /-- An alternative version of `subtype.exists`. This one is useful if Lean cannot figure out `q` when using `subtype.exists` from right to left. -/ protected theorem exists' {q : ∀ x, p x → Prop} : (∃ x h, q x h) ↔ ∃ x : { a // p a }, q x x.2 := (@Subtype.exists _ _ fun x ↦ q x.1 x.2).symm +#align subtype.exists' Subtype.exists' @[ext] protected theorem ext : ∀ {a1 a2 : { x // p x }}, (a1 : α) = (a2 : α) → a1 = a2 | ⟨_, _⟩, ⟨_, _⟩, rfl => rfl +#align subtype.ext Subtype.ext theorem ext_iff {a1 a2 : { x // p x }} : a1 = a2 ↔ (a1 : α) = (a2 : α) := ⟨congr_arg _, Subtype.ext⟩ +#align subtype.ext_iff Subtype.ext_iff theorem heq_iff_coe_eq (h : ∀ x, p x ↔ q x) {a1 : { x // p x }} {a2 : { x // q x }} : HEq a1 a2 ↔ (a1 : α) = (a2 : α) := Eq.rec (motive := λ (pp: (α → Prop)) _ => ∀ a2' : {x // pp x}, HEq a1 a2' ↔ (a1 : α) = (a2' : α)) (λ _ => heq_iff_eq.trans ext_iff) (funext $ λ x => propext (h x)) a2 +#align subtype.heq_iff_coe_eq Subtype.heq_iff_coe_eq lemma heq_iff_coe_heq {α β : Sort _} {p : α → Prop} {q : β → Prop} {a : {x // p x}} {b : {y // q y}} (h : α = β) (h' : HEq p q) : HEq a b ↔ HEq (a : α) (b : β) := by subst h subst h' rw [heq_iff_eq, heq_iff_eq, ext_iff] +#align subtype.heq_iff_coe_heq Subtype.heq_iff_coe_heq theorem ext_val {a1 a2 : { x // p x }} : a1.1 = a2.1 → a1 = a2 := Subtype.ext +#align subtype.ext_val Subtype.ext_val theorem ext_iff_val {a1 a2 : { x // p x }} : a1 = a2 ↔ a1.1 = a2.1 := ext_iff +#align subtype.ext_iff_val Subtype.ext_iff_val @[simp] theorem coe_eta (a : { a // p a }) (h : p a) : mk (↑a) h = a := Subtype.ext rfl +#align subtype.coe_eta Subtype.coe_eta theorem coe_mk (a h) : (@mk α p a h : α) = a := rfl +#align subtype.coe_mk Subtype.coe_mk -- Porting note: comment out `@[simp, nolint simp_nf]` -- Porting note: not clear if "build-in reduction doesn't always work" is still relevant @@ -97,23 +110,30 @@ theorem coe_mk (a h) : (@mk α p a h : α) = a := -- @[simp, nolint simp_nf] theorem mk_eq_mk {a h a' h'} : @mk α p a h = @mk α p a' h' ↔ a = a' := ext_iff +#align subtype.mk_eq_mk Subtype.mk_eq_mk theorem coe_eq_of_eq_mk {a : { a // p a }} {b : α} (h : ↑a = b) : a = ⟨b, h ▸ a.2⟩ := Subtype.ext h +#align subtype.coe_eq_of_eq_mk Subtype.coe_eq_of_eq_mk theorem coe_eq_iff {a : { a // p a }} {b : α} : ↑a = b ↔ ∃ h, a = ⟨b, h⟩ := ⟨fun h ↦ h ▸ ⟨a.2, (coe_eta _ _).symm⟩, fun ⟨_, ha⟩ ↦ ha.symm ▸ rfl⟩ +#align subtype.coe_eq_iff Subtype.coe_eq_iff theorem coe_injective : Injective (fun (a : Subtype p) ↦ (a : α)) := fun _ _ ↦ Subtype.ext +#align subtype.coe_injective Subtype.coe_injective theorem val_injective : Injective (@val _ p) := coe_injective +#align subtype.val_injective Subtype.val_injective theorem coe_inj {a b : Subtype p} : (a : α) = b ↔ a = b := coe_injective.eq_iff +#align subtype.coe_inj Subtype.coe_inj theorem val_inj {a b : Subtype p} : a.val = b.val ↔ a = b := coe_inj +#align subtype.val_inj Subtype.val_inj -- Porting note: it is unclear why the linter doesn't like this. -- If you understand why, please replace this comment with an explanation, or resolve. @@ -121,6 +141,7 @@ theorem val_inj {a b : Subtype p} : a.val = b.val ↔ a = b := theorem _root_.exists_eq_subtype_mk_iff {a : Subtype p} {b : α} : (∃ h : p b, a = Subtype.mk b h) ↔ ↑a = b := coe_eq_iff.symm +#align exists_eq_subtype_mk_iff exists_eq_subtype_mk_iff -- Porting note: it is unclear why the linter doesn't like this. -- If you understand why, please replace this comment with an explanation, or resolve. @@ -128,21 +149,26 @@ theorem _root_.exists_eq_subtype_mk_iff {a : Subtype p} {b : α} : theorem _root_.exists_subtype_mk_eq_iff {a : Subtype p} {b : α} : (∃ h : p b, Subtype.mk b h = a) ↔ b = a := by simp only [@eq_comm _ b, exists_eq_subtype_mk_iff, @eq_comm _ _ a] +#align exists_subtype_mk_eq_iff exists_subtype_mk_eq_iff /-- Restrict a (dependent) function to a subtype -/ def restrict {α} {β : α → Type _} (p : α → Prop) (f : ∀ x, β x) (x : Subtype p) : β x.1 := f x +#align subtype.restrict Subtype.restrict theorem restrict_apply {α} {β : α → Type _} (f : ∀ x, β x) (p : α → Prop) (x : Subtype p) : restrict p f x = f x.1 := by rfl +#align subtype.restrict_apply Subtype.restrict_apply theorem restrict_def {α β} (f : α → β) (p : α → Prop) : restrict p f = f ∘ (fun (a : Subtype p) ↦ a) := rfl +#align subtype.restrict_def Subtype.restrict_def theorem restrict_injective {α β} {f : α → β} (p : α → Prop) (h : Injective f) : Injective (restrict p f) := h.comp coe_injective +#align subtype.restrict_injective Subtype.restrict_injective theorem surjective_restrict {α} {β : α → Type _} [ne : ∀ a, Nonempty (β a)] (p : α → Prop) : Surjective fun f : ∀ x, β x ↦ restrict p f := by @@ -150,64 +176,81 @@ theorem surjective_restrict {α} {β : α → Type _} [ne : ∀ a, Nonempty (β refine' fun f ↦ ⟨fun x ↦ if h : p x then f ⟨x, h⟩ else Nonempty.some (ne x), funext <| _⟩ rintro ⟨x, hx⟩ exact dif_pos hx +#align subtype.surjective_restrict Subtype.surjective_restrict /-- Defining a map into a subtype, this can be seen as an "coinduction principle" of `Subtype`-/ @[simps] def coind {α β} (f : α → β) {p : β → Prop} (h : ∀ a, p (f a)) : α → Subtype p := fun a ↦ ⟨f a, h a⟩ +#align subtype.coind Subtype.coind +#align subtype.coind_coe Subtype.coind_coe theorem coind_injective {α β} {f : α → β} {p : β → Prop} (h : ∀ a, p (f a)) (hf : Injective f) : Injective (coind f h) := fun x y hxy ↦ hf <| by apply congr_arg Subtype.val hxy +#align subtype.coind_injective Subtype.coind_injective theorem coind_surjective {α β} {f : α → β} {p : β → Prop} (h : ∀ a, p (f a)) (hf : Surjective f) : Surjective (coind f h) := fun x ↦ let ⟨a, ha⟩ := hf x ⟨a, coe_injective ha⟩ +#align subtype.coind_surjective Subtype.coind_surjective theorem coind_bijective {α β} {f : α → β} {p : β → Prop} (h : ∀ a, p (f a)) (hf : Bijective f) : Bijective (coind f h) := ⟨coind_injective h hf.1, coind_surjective h hf.2⟩ +#align subtype.coind_bijective Subtype.coind_bijective /-- Restriction of a function to a function on subtypes. -/ @[simps] def map {p : α → Prop} {q : β → Prop} (f : α → β) (h : ∀ a, p a → q (f a)) : Subtype p → Subtype q := fun x ↦ ⟨f x, h x x.prop⟩ +#align subtype.map Subtype.map +#align subtype.map_coe Subtype.map_coe theorem map_comp {p : α → Prop} {q : β → Prop} {r : γ → Prop} {x : Subtype p} (f : α → β) (h : ∀ a, p a → q (f a)) (g : β → γ) (l : ∀ a, q a → r (g a)) : map g l (map f h x) = map (g ∘ f) (fun a ha ↦ l (f a) <| h a ha) x := rfl +#align subtype.map_comp Subtype.map_comp theorem map_id {p : α → Prop} {h : ∀ a, p a → p (id a)} : map (@id α) h = id := funext fun _ ↦ rfl +#align subtype.map_id Subtype.map_id theorem map_injective {p : α → Prop} {q : β → Prop} {f : α → β} (h : ∀ a, p a → q (f a)) (hf : Injective f) : Injective (map f h) := coind_injective _ <| hf.comp coe_injective +#align subtype.map_injective Subtype.map_injective theorem map_involutive {p : α → Prop} {f : α → α} (h : ∀ a, p a → p (f a)) (hf : Involutive f) : Involutive (map f h) := fun x ↦ Subtype.ext (hf x) +#align subtype.map_involutive Subtype.map_involutive instance [HasEquiv α] (p : α → Prop) : HasEquiv (Subtype p) := ⟨fun s t ↦ (s : α) ≈ (t : α)⟩ theorem equiv_iff [HasEquiv α] {p : α → Prop} {s t : Subtype p} : s ≈ t ↔ (s : α) ≈ (t : α) := Iff.rfl +#align subtype.equiv_iff Subtype.equiv_iff variable [Setoid α] protected theorem refl (s : Subtype p) : s ≈ s := Setoid.refl _ +#align subtype.refl Subtype.refl protected theorem symm {s t : Subtype p} (h : s ≈ t) : t ≈ s := Setoid.symm h +#align subtype.symm Subtype.symm protected theorem trans {s t u : Subtype p} (h₁ : s ≈ t) (h₂ : t ≈ u) : s ≈ u := Setoid.trans h₁ h₂ +#align subtype.trans Subtype.trans theorem equivalence (p : α → Prop) : Equivalence (@HasEquiv.Equiv (Subtype p) _) := .mk (Subtype.refl) (@Subtype.symm _ p _) (@Subtype.trans _ p _) +#align subtype.equivalence Subtype.equivalence instance (p : α → Prop) : Setoid (Subtype p) := Setoid.mk (· ≈ ·) (equivalence p) @@ -222,8 +265,10 @@ variable {α β γ : Type _} {p : α → Prop} @[simp] theorem coe_prop {S : Set α} (a : { a // a ∈ S }) : ↑a ∈ S := a.prop +#align subtype.coe_prop Subtype.coe_prop theorem val_prop {S : Set α} (a : { a // a ∈ S }) : a.val ∈ S := a.property +#align subtype.val_prop Subtype.val_prop end Subtype diff --git a/Mathlib/Data/Sum/Basic.lean b/Mathlib/Data/Sum/Basic.lean index b26d9ae321947..d0f272bf3cf66 100644 --- a/Mathlib/Data/Sum/Basic.lean +++ b/Mathlib/Data/Sum/Basic.lean @@ -51,6 +51,7 @@ deriving instance BEq for Sum @[simp] theorem «forall» {p : Sum α β → Prop} : (∀ x, p x) ↔ (∀ a, p (inl a)) ∧ ∀ b, p (inr b) := ⟨fun h ↦ ⟨fun _ ↦ h _, fun _ ↦ h _⟩, fun ⟨h₁, h₂⟩ ↦ Sum.rec h₁ h₂⟩ +#align sum.forall Sum.forall @[simp] theorem «exists» {p : Sum α β → Prop} : (∃ x, p x) ↔ (∃ a, p (inl a)) ∨ ∃ b, p (inr b) := @@ -60,10 +61,13 @@ theorem «exists» {p : Sum α β → Prop} : (∃ x, p x) ↔ (∃ a, p (inl a) fun | Or.inl ⟨a, h⟩ => ⟨inl a, h⟩ | Or.inr ⟨b, h⟩ => ⟨inr b, h⟩⟩ +#align sum.exists Sum.exists theorem inl_injective : Function.Injective (inl : α → Sum α β) := fun _ _ ↦ inl.inj +#align sum.inl_injective Sum.inl_injective theorem inr_injective : Function.Injective (inr : β → Sum α β) := fun _ _ ↦ inr.inj +#align sum.inr_injective Sum.inr_injective section get @@ -71,21 +75,25 @@ section get def getLeft : Sum α β → Option α | inl a => some a | inr _ => none +#align sum.get_left Sum.getLeft /-- Check if a sum is `inr` and if so, retrieve its contents. -/ def getRight : Sum α β → Option β | inr b => some b | inl _ => none +#align sum.get_right Sum.getRight /-- Check if a sum is `inl`. -/ def isLeft : Sum α β → Bool | inl _ => true | inr _ => false +#align sum.is_left Sum.isLeft /-- Check if a sum is `inr`. -/ def isRight : Sum α β → Bool | inl _ => false | inr _ => true +#align sum.is_right Sum.isRight variable {x y : Sum α β} @@ -157,68 +165,84 @@ theorem inr.inj_iff {a b} : (inr a : Sum α β) = inr b ↔ a = b := theorem inl_ne_inr {a : α} {b : β} : inl a ≠ inr b := fun. +#align sum.inl_ne_inr Sum.inl_ne_inr theorem inr_ne_inl {a : α} {b : β} : inr b ≠ inl a := fun. +#align sum.inr_ne_inl Sum.inr_ne_inl /-- Define a function on `α ⊕ β` by giving separate definitions on `α` and `β`. -/ protected def elim {α β γ : Sort _} (f : α → γ) (g : β → γ) : Sum α β → γ := fun x ↦ Sum.casesOn x f g +#align sum.elim Sum.elim @[simp] theorem elim_inl {α β γ : Sort _} (f : α → γ) (g : β → γ) (x : α) : Sum.elim f g (inl x) = f x := rfl +#align sum.elim_inl Sum.elim_inl @[simp] theorem elim_inr {α β γ : Sort _} (f : α → γ) (g : β → γ) (x : β) : Sum.elim f g (inr x) = g x := rfl +#align sum.elim_inr Sum.elim_inr @[simp] theorem elim_comp_inl {α β γ : Sort _} (f : α → γ) (g : β → γ) : Sum.elim f g ∘ inl = f := rfl +#align sum.elim_comp_inl Sum.elim_comp_inl @[simp] theorem elim_comp_inr {α β γ : Sort _} (f : α → γ) (g : β → γ) : Sum.elim f g ∘ inr = g := rfl +#align sum.elim_comp_inr Sum.elim_comp_inr @[simp] theorem elim_inl_inr {α β : Sort _} : @Sum.elim α β _ inl inr = id := funext fun x ↦ Sum.casesOn x (fun _ ↦ rfl) fun _ ↦ rfl +#align sum.elim_inl_inr Sum.elim_inl_inr theorem comp_elim {α β γ δ : Sort _} (f : γ → δ) (g : α → γ) (h : β → γ) : f ∘ Sum.elim g h = Sum.elim (f ∘ g) (f ∘ h) := funext fun x ↦ Sum.casesOn x (fun _ ↦ rfl) fun _ ↦ rfl +#align sum.comp_elim Sum.comp_elim @[simp] theorem elim_comp_inl_inr {α β γ : Sort _} (f : Sum α β → γ) : Sum.elim (f ∘ inl) (f ∘ inr) = f := funext fun x ↦ Sum.casesOn x (fun _ ↦ rfl) fun _ ↦ rfl +#align sum.elim_comp_inl_inr Sum.elim_comp_inl_inr /-- Map `α ⊕ β` to `α' ⊕ β'` sending `α` to `α'` and `β` to `β'`. -/ protected def map (f : α → α') (g : β → β') : Sum α β → Sum α' β' := Sum.elim (inl ∘ f) (inr ∘ g) +#align sum.map Sum.map @[simp] theorem map_inl (f : α → α') (g : β → β') (x : α) : (inl x).map f g = inl (f x) := rfl +#align sum.map_inl Sum.map_inl @[simp] theorem map_inr (f : α → α') (g : β → β') (x : β) : (inr x).map f g = inr (g x) := rfl +#align sum.map_inr Sum.map_inr @[simp] theorem map_map {α'' β''} (f' : α' → α'') (g' : β' → β'') (f : α → α') (g : β → β') : ∀ x : Sum α β, (x.map f g).map f' g' = x.map (f' ∘ f) (g' ∘ g) | inl _ => rfl | inr _ => rfl +#align sum.map_map Sum.map_map @[simp] theorem map_comp_map {α'' β''} (f' : α' → α'') (g' : β' → β'') (f : α → α') (g : β → β') : Sum.map f' g' ∘ Sum.map f g = Sum.map (f' ∘ f) (g' ∘ g) := funext <| map_map f' g' f g +#align sum.map_comp_map Sum.map_comp_map @[simp] theorem map_id_id (α β) : Sum.map (@id α) (@id β) = id := funext fun x ↦ Sum.recOn x (fun _ ↦ rfl) fun _ ↦ rfl +#align sum.map_id_id Sum.map_id_id theorem elim_map {α β γ δ ε : Sort _} {f₁ : α → β} {f₂ : β → ε} {g₁ : γ → δ} {g₂ : δ → ε} {x} : Sum.elim f₂ g₂ (Sum.map f₁ g₁ x) = Sum.elim (f₂ ∘ f₁) (g₂ ∘ g₁) x := by @@ -228,6 +252,7 @@ theorem elim_map {α β γ δ ε : Sort _} {f₁ : α → β} {f₂ : β → ε} theorem elim_comp_map {α β γ δ ε : Sort _} {f₁ : α → β} {f₂ : β → ε} {g₁ : γ → δ} {g₂ : δ → ε} : Sum.elim f₂ g₂ ∘ Sum.map f₁ g₁ = Sum.elim (f₂ ∘ f₁) (g₂ ∘ g₁) := funext $ fun _ => elim_map +#align sum.elim_comp_map Sum.elim_comp_map @[simp] theorem isLeft_map (f : α → β) (g : γ → δ) (x : Sum α γ) : isLeft (x.map f g) = isLeft x := @@ -255,76 +280,93 @@ open Function (update update_eq_iff update_comp_eq_of_injective update_comp_eq_o theorem update_elim_inl [DecidableEq α] [DecidableEq (Sum α β)] {f : α → γ} {g : β → γ} {i : α} {x : γ} : update (Sum.elim f g) (inl i) x = Sum.elim (update f i x) g := update_eq_iff.2 ⟨by simp, by simp (config := { contextual := true })⟩ +#align sum.update_elim_inl Sum.update_elim_inl @[simp] theorem update_elim_inr [DecidableEq β] [DecidableEq (Sum α β)] {f : α → γ} {g : β → γ} {i : β} {x : γ} : update (Sum.elim f g) (inr i) x = Sum.elim f (update g i x) := update_eq_iff.2 ⟨by simp, by simp (config := { contextual := true })⟩ +#align sum.update_elim_inr Sum.update_elim_inr @[simp] theorem update_inl_comp_inl [DecidableEq α] [DecidableEq (Sum α β)] {f : Sum α β → γ} {i : α} {x : γ} : update f (inl i) x ∘ inl = update (f ∘ inl) i x := update_comp_eq_of_injective _ inl_injective _ _ +#align sum.update_inl_comp_inl Sum.update_inl_comp_inl @[simp] theorem update_inl_apply_inl [DecidableEq α] [DecidableEq (Sum α β)] {f : Sum α β → γ} {i j : α} {x : γ} : update f (inl i) x (inl j) = update (f ∘ inl) i x j := by rw [← update_inl_comp_inl, Function.comp_apply] +#align sum.update_inl_apply_inl Sum.update_inl_apply_inl @[simp] theorem update_inl_comp_inr [DecidableEq (Sum α β)] {f : Sum α β → γ} {i : α} {x : γ} : update f (inl i) x ∘ inr = f ∘ inr := (update_comp_eq_of_forall_ne _ _) fun _ ↦ inr_ne_inl +#align sum.update_inl_comp_inr Sum.update_inl_comp_inr theorem update_inl_apply_inr [DecidableEq (Sum α β)] {f : Sum α β → γ} {i : α} {j : β} {x : γ} : update f (inl i) x (inr j) = f (inr j) := Function.update_noteq inr_ne_inl _ _ +#align sum.update_inl_apply_inr Sum.update_inl_apply_inr @[simp] theorem update_inr_comp_inl [DecidableEq (Sum α β)] {f : Sum α β → γ} {i : β} {x : γ} : update f (inr i) x ∘ inl = f ∘ inl := (update_comp_eq_of_forall_ne _ _) fun _ ↦ inl_ne_inr +#align sum.update_inr_comp_inl Sum.update_inr_comp_inl theorem update_inr_apply_inl [DecidableEq (Sum α β)] {f : Sum α β → γ} {i : α} {j : β} {x : γ} : update f (inr j) x (inl i) = f (inl i) := Function.update_noteq inl_ne_inr _ _ +#align sum.update_inr_apply_inl Sum.update_inr_apply_inl @[simp] theorem update_inr_comp_inr [DecidableEq β] [DecidableEq (Sum α β)] {f : Sum α β → γ} {i : β} {x : γ} : update f (inr i) x ∘ inr = update (f ∘ inr) i x := update_comp_eq_of_injective _ inr_injective _ _ +#align sum.update_inr_comp_inr Sum.update_inr_comp_inr @[simp] theorem update_inr_apply_inr [DecidableEq β] [DecidableEq (Sum α β)] {f : Sum α β → γ} {i j : β} {x : γ} : update f (inr i) x (inr j) = update (f ∘ inr) i x j := by rw [← update_inr_comp_inr, Function.comp_apply] +#align sum.update_inr_apply_inr Sum.update_inr_apply_inr /-- Swap the factors of a sum type -/ def swap : Sum α β → Sum β α := Sum.elim inr inl +#align sum.swap Sum.swap @[simp] theorem swap_inl (x : α) : swap (inl x : Sum α β) = inr x := rfl +#align sum.swap_inl Sum.swap_inl @[simp] theorem swap_inr (x : β) : swap (inr x : Sum α β) = inl x := rfl +#align sum.swap_inr Sum.swap_inr @[simp] theorem swap_swap (x : Sum α β) : swap (swap x) = x := by cases x <;> rfl +#align sum.swap_swap Sum.swap_swap @[simp] theorem swap_swap_eq : swap ∘ swap = @id (Sum α β) := funext <| swap_swap +#align sum.swap_swap_eq Sum.swap_swap_eq @[simp] theorem swap_leftInverse : Function.LeftInverse (@swap α β) swap := swap_swap +#align sum.swap_left_inverse Sum.swap_leftInverse @[simp] theorem swap_rightInverse : Function.RightInverse (@swap α β) swap := swap_swap +#align sum.swap_right_inverse Sum.swap_rightInverse @[simp] theorem isLeft_swap (x : Sum α β) : x.swap.isLeft = x.isRight := by cases x <;> rfl @@ -349,6 +391,7 @@ section LiftRel inductive LiftRel (r : α → γ → Prop) (s : β → δ → Prop) : Sum α β → Sum γ δ → Prop | protected inl {a c} : r a c → LiftRel r s (inl a) (inl c) | protected inr {b d} : s b d → LiftRel r s (inr b) (inr d) +#align sum.lift_rel Sum.LiftRel variable {r r₁ r₂ : α → γ → Prop} {s s₁ s₂ : β → δ → Prop} {a : α} {b : β} {c : γ} {d : δ} {x : Sum α β} {y : Sum γ δ} @@ -393,19 +436,23 @@ theorem LiftRel.mono (hr : ∀ a b, r₁ a b → r₂ a b) (hs : ∀ a b, s₁ a cases h · exact LiftRel.inl (hr _ _ ‹_›) · exact LiftRel.inr (hs _ _ ‹_›) +#align sum.lift_rel.mono Sum.LiftRel.mono theorem LiftRel.mono_left (hr : ∀ a b, r₁ a b → r₂ a b) (h : LiftRel r₁ s x y) : LiftRel r₂ s x y := (h.mono hr) fun _ _ ↦ id +#align sum.lift_rel.mono_left Sum.LiftRel.mono_left theorem LiftRel.mono_right (hs : ∀ a b, s₁ a b → s₂ a b) (h : LiftRel r s₁ x y) : LiftRel r s₂ x y := h.mono (fun _ _ ↦ id) hs +#align sum.lift_rel.mono_right Sum.LiftRel.mono_right protected theorem LiftRel.swap (h : LiftRel r s x y) : LiftRel s r x.swap y.swap := by cases h · exact LiftRel.inr ‹_› · exact LiftRel.inl ‹_› +#align sum.lift_rel.swap Sum.LiftRel.swap @[simp] theorem liftRel_swap_iff : LiftRel s r x.swap y.swap ↔ LiftRel r s x y := @@ -428,6 +475,7 @@ inductive Lex (r : α → α → Prop) (s : β → β → Prop) : Sum α β → #align sum.lex.inl Sum.Lex.inl #align sum.lex.inr Sum.Lex.inr #align sum.lex.sep Sum.Lex.sep +#align sum.lex Sum.Lex attribute [simp] Lex.sep @@ -439,16 +487,19 @@ theorem lex_inl_inl : Lex r s (inl a₁) (inl a₂) ↔ r a₁ a₂ := ⟨fun h ↦ by cases h assumption, Lex.inl⟩ +#align sum.lex_inl_inl Sum.lex_inl_inl @[simp] theorem lex_inr_inr : Lex r s (inr b₁) (inr b₂) ↔ s b₁ b₂ := ⟨fun h ↦ by cases h assumption, Lex.inr⟩ +#align sum.lex_inr_inr Sum.lex_inr_inr @[simp] theorem lex_inr_inl : ¬Lex r s (inr b) (inl a) := fun. +#align sum.lex_inr_inl Sum.lex_inr_inl instance [DecidableRel r] [DecidableRel s] : DecidableRel (Lex r s) | inl _, inl _ => decidable_of_iff' _ lex_inl_inl @@ -460,6 +511,7 @@ protected theorem LiftRel.lex {a b : Sum α β} (h : LiftRel r s a b) : Lex r s cases h · exact Lex.inl ‹_› · exact Lex.inr ‹_› +#align sum.lift_rel.lex Sum.LiftRel.lex theorem liftRel_subrelation_lex : Subrelation (LiftRel r s) (Lex r s) := LiftRel.lex @@ -471,12 +523,15 @@ theorem Lex.mono (hr : ∀ a b, r₁ a b → r₂ a b) (hs : ∀ a b, s₁ a b · exact Lex.inl (hr _ _ ‹_›) · exact Lex.inr (hs _ _ ‹_›) · exact Lex.sep _ _ +#align sum.lex.mono Sum.Lex.mono theorem Lex.mono_left (hr : ∀ a b, r₁ a b → r₂ a b) (h : Lex r₁ s x y) : Lex r₂ s x y := (h.mono hr) fun _ _ ↦ id +#align sum.lex.mono_left Sum.Lex.mono_left theorem Lex.mono_right (hs : ∀ a b, s₁ a b → s₂ a b) (h : Lex r s₁ x y) : Lex r s₂ x y := h.mono (fun _ _ ↦ id) hs +#align sum.lex.mono_right Sum.Lex.mono_right theorem lex_acc_inl {a} (aca : Acc r a) : Acc (Lex r s) (inl a) := by induction' aca with a _ IH @@ -484,6 +539,7 @@ theorem lex_acc_inl {a} (aca : Acc r a) : Acc (Lex r s) (inl a) := by intro y h cases' h with a' _ h' exact IH _ h' +#align sum.lex_acc_inl Sum.lex_acc_inl theorem lex_acc_inr (aca : ∀ a, Acc (Lex r s) (inl a)) {b} (acb : Acc s b) : Acc (Lex r s) (inr b) := by @@ -493,10 +549,12 @@ theorem lex_acc_inr (aca : ∀ a, Acc (Lex r s) (inl a)) {b} (acb : Acc s b) : cases' h with _ _ _ b' _ h' a · exact IH _ h' · exact aca _ +#align sum.lex_acc_inr Sum.lex_acc_inr theorem lex_wf (ha : WellFounded r) (hb : WellFounded s) : WellFounded (Lex r s) := have aca : ∀ a, Acc (Lex r s) (inl a) := fun a ↦ lex_acc_inl (ha.apply a) ⟨fun x ↦ Sum.recOn x aca fun b ↦ lex_acc_inr aca (hb.apply b)⟩ +#align sum.lex_wf Sum.lex_wf end Lex @@ -512,11 +570,13 @@ theorem Injective.sum_elim {f : α → γ} {g : β → γ} (hf : Injective f) (h | inl _, inr _, h => (hfg _ _ h).elim | inr _, inl _, h => (hfg _ _ h.symm).elim | inr _, inr _, h => congr_arg inr <| hg h +#align function.injective.sum_elim Function.Injective.sum_elim theorem Injective.sum_map {f : α → β} {g : α' → β'} (hf : Injective f) (hg : Injective g) : Injective (Sum.map f g) | inl _, inl _, h => congr_arg inl <| hf <| inl.inj h | inr _, inr _, h => congr_arg inr <| hg <| inr.inj h +#align function.injective.sum_map Function.Injective.sum_map theorem Surjective.sum_map {f : α → β} {g : α' → β'} (hf : Surjective f) (hg : Surjective g) : Surjective (Sum.map f g) @@ -526,6 +586,7 @@ theorem Surjective.sum_map {f : α → β} {g : α' → β'} (hf : Surjective f) | inr y => let ⟨x, hx⟩ := hg y ⟨inr x, congr_arg inr hx⟩ +#align function.surjective.sum_map Function.Surjective.sum_map end Function @@ -537,11 +598,13 @@ theorem elim_const_const (c : γ) : Sum.elim (const _ c : α → γ) (const _ c : β → γ) = const _ c := by ext x cases x <;> rfl +#align sum.elim_const_const Sum.elim_const_const @[simp] theorem elim_lam_const_lam_const (c : γ) : (Sum.elim (fun _ : α ↦ c) fun _ : β ↦ c) = fun _ ↦ c := Sum.elim_const_const c +#align sum.elim_lam_const_lam_const Sum.elim_lam_const_lam_const theorem elim_update_left [DecidableEq α] [DecidableEq β] (f : α → γ) (g : β → γ) (i : α) (c : γ) : Sum.elim (Function.update f i c) g = Function.update (Sum.elim f g) (inl i) c := by @@ -552,6 +615,7 @@ theorem elim_update_left [DecidableEq α] [DecidableEq β] (f : α → γ) (g : simp · simp [h] · simp +#align sum.elim_update_left Sum.elim_update_left theorem elim_update_right [DecidableEq α] [DecidableEq β] (f : α → γ) (g : β → γ) (i : β) (c : γ) : Sum.elim f (Function.update g i c) = Function.update (Sum.elim f g) (inr i) c := by @@ -562,6 +626,7 @@ theorem elim_update_right [DecidableEq α] [DecidableEq β] (f : α → γ) (g : · subst h simp · simp [h] +#align sum.elim_update_right Sum.elim_update_right end Sum @@ -577,15 +642,18 @@ namespace Sum3 @[match_pattern, simp, reducible] def in₀ (a : α) : Sum α (Sum β γ) := inl a +#align sum3.in₀ Sum3.in₀ /-- The map from the second summand into a ternary sum. -/ @[match_pattern, simp, reducible] def in₁ (b : β) : Sum α (Sum β γ) := inr <| inl b +#align sum3.in₁ Sum3.in₁ /-- The map from the third summand into a ternary sum. -/ @[match_pattern, simp, reducible] def in₂ (c : γ) : Sum α (Sum β γ) := inr <| inr c +#align sum3.in₂ Sum3.in₂ end Sum3 diff --git a/Mathlib/Data/Sum/Order.lean b/Mathlib/Data/Sum/Order.lean index 7dfbc9f7981b8..83c1ebd9be590 100644 --- a/Mathlib/Data/Sum/Order.lean +++ b/Mathlib/Data/Sum/Order.lean @@ -553,6 +553,7 @@ variable [LE α] [LE β] [LE γ] (a : α) (b : β) (c : γ) def sumComm (α β : Type _) [LE α] [LE β] : Sum α β ≃o Sum β α := { Equiv.sumComm α β with map_rel_iff' := swap_le_swap_iff } #align order_iso.sum_comm OrderIso.sumComm +#align order_iso.sum_comm_apply OrderIso.sumComm_apply @[simp] theorem sumComm_symm (α β : Type _) [LE α] [LE β] : diff --git a/Mathlib/Data/Sym/Basic.lean b/Mathlib/Data/Sym/Basic.lean index ebf843ed05cf6..3979326b71eda 100644 --- a/Mathlib/Data/Sym/Basic.lean +++ b/Mathlib/Data/Sym/Basic.lean @@ -418,6 +418,8 @@ def equivCongr (e : α ≃ β) : Sym α n ≃ Sym β n left_inv x := by rw [map_map, Equiv.symm_comp_self, map_id] right_inv x := by rw [map_map, Equiv.self_comp_symm, map_id] #align sym.equiv_congr Sym.equivCongr +#align sym.equiv_congr_symm_apply Sym.equivCongr_symm_apply +#align sym.equiv_congr_apply Sym.equivCongr_apply /-- "Attach" a proof that `a ∈ s` to each element `a` in `s` to produce an element of the symmetric power on `{x // x ∈ s}`. -/ diff --git a/Mathlib/Data/TwoPointing.lean b/Mathlib/Data/TwoPointing.lean index a2fc7dbb9dad4..7eb7dab58d7b5 100644 --- a/Mathlib/Data/TwoPointing.lean +++ b/Mathlib/Data/TwoPointing.lean @@ -39,6 +39,8 @@ structure TwoPointing (α : Type _) extends α × α where fst_ne_snd : fst ≠ snd deriving DecidableEq #align two_pointing TwoPointing +#align two_pointing.ext TwoPointing.ext +#align two_pointing.ext_iff TwoPointing.ext_iff namespace TwoPointing @@ -53,6 +55,7 @@ theorem snd_ne_fst : p.snd ≠ p.fst := def swap : TwoPointing α := ⟨(p.snd, p.fst), p.snd_ne_fst⟩ #align two_pointing.swap TwoPointing.swap +#align two_pointing.swap_to_prod TwoPointing.swap_toProd theorem swap_fst : p.swap.fst = p.snd := rfl #align two_pointing.swap_fst TwoPointing.swap_fst diff --git a/Mathlib/Data/W/Basic.lean b/Mathlib/Data/W/Basic.lean index ff3e505ea9b3a..a347124e8d67e 100644 --- a/Mathlib/Data/W/Basic.lean +++ b/Mathlib/Data/W/Basic.lean @@ -82,6 +82,8 @@ def equivSigma : WType β ≃ Σa : α, β a → WType β left_inv := ofSigma_toSigma right_inv := toSigma_ofSigma #align W_type.equiv_sigma WType.equivSigma +#align W_type.equiv_sigma_symm_apply WType.equivSigma_symm_apply +#align W_type.equiv_sigma_apply WType.equivSigma_apply variable {β} diff --git a/Mathlib/Deprecated/Submonoid.lean b/Mathlib/Deprecated/Submonoid.lean index 5892667ef7b48..25e802bb120a8 100644 --- a/Mathlib/Deprecated/Submonoid.lean +++ b/Mathlib/Deprecated/Submonoid.lean @@ -89,6 +89,7 @@ theorem IsSubmonoid.inter {s₁ s₂ : Set M} (is₁ : IsSubmonoid s₁) (is₂ { one_mem := ⟨is₁.one_mem, is₂.one_mem⟩ mul_mem := @fun _ _ hx hy => ⟨is₁.mul_mem hx.1 hy.1, is₂.mul_mem hx.2 hy.2⟩ } #align is_submonoid.inter IsSubmonoid.inter +#align is_add_submonoid.inter IsAddSubmonoid.inter /-- The intersection of an indexed set of submonoids of a monoid `M` is a submonoid of `M`. -/ @[to_additive @@ -100,6 +101,7 @@ theorem IsSubmonoid.Inter {ι : Sort _} {s : ι → Set M} (h : ∀ y : ι, IsSu mul_mem := @fun _ _ h₁ h₂ => Set.mem_interᵢ.2 fun y => (h y).mul_mem (Set.mem_interᵢ.1 h₁ y) (Set.mem_interᵢ.1 h₂ y) } #align is_submonoid.Inter IsSubmonoid.Inter +#align is_add_submonoid.Inter IsAddSubmonoid.Inter /-- The union of an indexed, directed, nonempty set of submonoids of a monoid `M` is a submonoid of `M`. -/ @@ -118,6 +120,7 @@ theorem is_submonoid_Union_of_directed {ι : Type _} [hι : Nonempty ι] {s : ι let ⟨k, hk⟩ := Directed i j Set.mem_unionᵢ.2 ⟨k, (hs k).mul_mem (hk.1 hi) (hk.2 hj)⟩ } #align is_submonoid_Union_of_directed is_submonoid_Union_of_directed +#align is_add_submonoid_Union_of_directed is_addSubmonoid_Union_of_directed section powers @@ -127,12 +130,14 @@ section powers def powers (x : M) : Set M := { y | ∃ n : ℕ, x ^ n = y } #align powers powers +#align multiples multiples /-- 1 is in the set of natural number powers of an element of a monoid. -/ @[to_additive "0 is in the set of natural number multiples of an element of an `AddMonoid`."] theorem powers.one_mem {x : M} : (1 : M) ∈ powers x := ⟨0, pow_zero _⟩ #align powers.one_mem powers.one_mem +#align multiples.zero_mem multiples.zero_mem /-- An element of a monoid is in the set of that element's natural number powers. -/ @[to_additive @@ -140,6 +145,7 @@ theorem powers.one_mem {x : M} : (1 : M) ∈ powers x := theorem powers.self_mem {x : M} : x ∈ powers x := ⟨1, pow_one _⟩ #align powers.self_mem powers.self_mem +#align multiples.self_mem multiples.self_mem /-- The set of natural number powers of an element of a monoid is closed under multiplication. -/ @[to_additive @@ -148,6 +154,7 @@ theorem powers.self_mem {x : M} : x ∈ powers x := theorem powers.mul_mem {x y z : M} : y ∈ powers x → z ∈ powers x → y * z ∈ powers x := fun ⟨n₁, h₁⟩ ⟨n₂, h₂⟩ => ⟨n₁ + n₂, by simp only [pow_add, *]⟩ #align powers.mul_mem powers.mul_mem +#align multiples.add_mem multiples.add_mem /-- The set of natural number powers of an element of a monoid `M` is a submonoid of `M`. -/ @[to_additive @@ -157,11 +164,13 @@ theorem powers.is_submonoid (x : M) : IsSubmonoid (powers x) := { one_mem := powers.one_mem mul_mem := @fun _ _ => powers.mul_mem } #align powers.is_submonoid powers.is_submonoid +#align multiples.is_add_submonoid multiples.is_addSubmonoid /-- A monoid is a submonoid of itself. -/ @[to_additive "An `add_monoid` is an `add_submonoid` of itself."] theorem Univ.IsSubmonoid : IsSubmonoid (@Set.univ M) := by constructor <;> simp #align univ.is_submonoid Univ.IsSubmonoid +#align univ.is_add_submonoid Univ.IsAddSubmonoid /-- The preimage of a submonoid under a monoid hom is a submonoid of the domain. -/ @[to_additive @@ -173,6 +182,7 @@ theorem IsSubmonoid.preimage {N : Type _} [Monoid N] {f : M → N} (hf : IsMonoi mul_mem := @fun a b (ha : f a ∈ s) (hb : f b ∈ s) => show f (a * b) ∈ s by (rw [IsMonoidHom.map_mul' hf]; exact hs.mul_mem ha hb) } #align is_submonoid.preimage IsSubmonoid.preimage +#align is_add_submonoid.preimage IsAddSubmonoid.preimage /-- The image of a submonoid under a monoid hom is a submonoid of the codomain. -/ @[to_additive @@ -184,6 +194,7 @@ theorem IsSubmonoid.image {γ : Type _} [Monoid γ] {f : M → γ} (hf : IsMonoi mul_mem := @fun a b ⟨x, hx⟩ ⟨y, hy⟩ => ⟨x * y, hs.mul_mem hx.1 hy.1, by rw [hf.map_mul, hx.2, hy.2]⟩ } #align is_submonoid.image IsSubmonoid.image +#align is_add_submonoid.image IsAddSubmonoid.image /-- The image of a monoid hom is a submonoid of the codomain. -/ @[to_additive "The image of an `AddMonoid` hom is an `AddSubmonoid` of the codomain."] @@ -192,6 +203,7 @@ theorem Range.is_submonoid {γ : Type _} [Monoid γ] {f : M → γ} (hf : IsMono rw [← Set.image_univ] exact Univ.IsSubmonoid.image hf #align range.is_submonoid Range.is_submonoid +#align range.is_add_submonoid Range.is_addSubmonoid /-- Submonoids are closed under natural powers. -/ @[to_additive @@ -212,6 +224,7 @@ theorem IsSubmonoid.pow_mem {a : M} (hs : IsSubmonoid s) (h : a ∈ s) : ∀ {n theorem IsSubmonoid.power_subset {a : M} (hs : IsSubmonoid s) (h : a ∈ s) : powers a ⊆ s := fun _ ⟨_, hx⟩ => hx ▸ hs.pow_mem h #align is_submonoid.power_subset IsSubmonoid.power_subset +#align is_add_submonoid.multiples_subset IsAddSubmonoid.multiples_subset end powers @@ -227,6 +240,7 @@ theorem list_prod_mem (hs : IsSubmonoid s) : ∀ {l : List M}, (∀ x ∈ l, x have : a ∈ s ∧ ∀ x ∈ l, x ∈ s := by simpa using h hs.mul_mem this.1 (list_prod_mem hs this.2) #align is_submonoid.list_prod_mem IsSubmonoid.list_prod_mem +#align is_add_submonoid.list_sum_mem IsAddSubmonoid.list_sum_mem /-- The product of a multiset of elements of a submonoid of a `CommMonoid` is an element of the submonoid. -/ @@ -239,6 +253,7 @@ theorem multiset_prod_mem {M} [CommMonoid M] {s : Set M} (hs : IsSubmonoid s) (m rw [Multiset.quot_mk_to_coe, Multiset.coe_prod] exact list_prod_mem hs hl #align is_submonoid.multiset_prod_mem IsSubmonoid.multiset_prod_mem +#align is_add_submonoid.multiset_sum_mem IsAddSubmonoid.multiset_sum_mem /-- The product of elements of a submonoid of a `comm_monoid` indexed by a `finset` is an element of the submonoid. -/ @@ -249,6 +264,7 @@ theorem finset_prod_mem {M A} [CommMonoid M] {s : Set M} (hs : IsSubmonoid s) (f ∀ t : Finset A, (∀ b ∈ t, f b ∈ s) → (∏ b in t, f b) ∈ s | ⟨m, hm⟩, _ => multiset_prod_mem hs _ (by simpa) #align is_submonoid.finset_prod_mem IsSubmonoid.finset_prod_mem +#align is_add_submonoid.finset_sum_mem IsAddSubmonoid.finset_sum_mem end IsSubmonoid @@ -290,12 +306,14 @@ theorem closure.IsSubmonoid (s : Set M) : IsSubmonoid (Closure s) := { one_mem := InClosure.one mul_mem := @fun _ _ => InClosure.mul } #align monoid.closure.is_submonoid Monoid.closure.IsSubmonoid +#align add_monoid.closure.is_add_submonoid AddMonoid.closure.IsAddSubmonoid /-- A subset of a monoid is contained in the submonoid it generates. -/ @[to_additive "A subset of an `AddMonoid` is contained in the `AddSubmonoid` it generates."] theorem subset_closure {s : Set M} : s ⊆ Closure s := fun _ => InClosure.basic #align monoid.subset_closure Monoid.subset_closure +#align add_monoid.subset_closure AddMonoid.subset_closure /-- The submonoid generated by a set is contained in any submonoid that contains the set. -/ @[to_additive @@ -304,6 +322,7 @@ theorem subset_closure {s : Set M} : s ⊆ Closure s := fun _ => InClosure.basic theorem closure_subset {s t : Set M} (ht : IsSubmonoid t) (h : s ⊆ t) : Closure s ⊆ t := fun a ha => by induction ha <;> simp [h _, *, IsSubmonoid.one_mem, IsSubmonoid.mul_mem] #align monoid.closure_subset Monoid.closure_subset +#align add_monoid.closure_subset AddMonoid.closure_subset /-- Given subsets `t` and `s` of a monoid `M`, if `s ⊆ t`, the submonoid of `M` generated by `s` is contained in the submonoid generated by `t`. -/ @@ -313,6 +332,7 @@ theorem closure_subset {s t : Set M} (ht : IsSubmonoid t) (h : s ⊆ t) : Closur theorem closure_mono {s t : Set M} (h : s ⊆ t) : Closure s ⊆ Closure t := closure_subset (closure.IsSubmonoid t) <| Set.Subset.trans h subset_closure #align monoid.closure_mono Monoid.closure_mono +#align add_monoid.closure_mono AddMonoid.closure_mono /-- The submonoid generated by an element of a monoid equals the set of natural number powers of the element. -/ @@ -325,6 +345,7 @@ theorem closure_singleton {x : M} : Closure ({x} : Set M) = powers x := IsSubmonoid.power_subset (closure.IsSubmonoid _) <| Set.singleton_subset_iff.1 <| subset_closure #align monoid.closure_singleton Monoid.closure_singleton +#align add_monoid.closure_singleton AddMonoid.closure_singleton /-- The image under a monoid hom of the submonoid generated by a set equals the submonoid generated by the image of the set under the monoid hom. -/ @@ -345,6 +366,7 @@ theorem image_closure {A : Type _} [Monoid A] {f : M → A} (hf : IsMonoidHom f) (closure_subset (IsSubmonoid.image hf (closure.IsSubmonoid _)) <| Set.image_subset _ subset_closure) #align monoid.image_closure Monoid.image_closure +#align add_monoid.image_closure AddMonoid.image_closure /-- Given an element `a` of the submonoid of a monoid `M` generated by a set `s`, there exists a list of elements of `s` whose product is `a`. -/ @@ -363,6 +385,7 @@ theorem exists_list_of_mem_closure {s : Set M} {a : M} (h : a ∈ Closure s) : simp [eqa.symm, eqb.symm, or_imp] exact fun a => ⟨ha a, hb a⟩ #align monoid.exists_list_of_mem_closure Monoid.exists_list_of_mem_closure +#align add_monoid.exists_list_of_mem_closure AddMonoid.exists_list_of_mem_closure /-- Given sets `s, t` of a commutative monoid `M`, `x ∈ M` is in the submonoid of `M` generated by `s ∪ t` iff there exists an element of the submonoid generated by `s` and an element of the @@ -394,6 +417,7 @@ theorem mem_closure_union_iff {M : Type _} [CommMonoid M] {s t : Set M} {x : M} (closure.IsSubmonoid _).mul_mem (closure_mono (Set.subset_union_left _ _) hy) (closure_mono (Set.subset_union_right _ _) hz)⟩ #align monoid.mem_closure_union_iff Monoid.mem_closure_union_iff +#align add_monoid.mem_closure_union_iff AddMonoid.mem_closure_union_iff end Monoid @@ -402,8 +426,10 @@ end Monoid def Submonoid.of {s : Set M} (h : IsSubmonoid s) : Submonoid M := ⟨⟨s, @fun _ _ => h.2⟩, h.1⟩ #align submonoid.of Submonoid.of +#align add_submonoid.of AddSubmonoid.of @[to_additive] theorem Submonoid.is_submonoid (S : Submonoid M) : IsSubmonoid (S : Set M) := by refine' ⟨S.2, S.1.2⟩ #align submonoid.is_submonoid Submonoid.is_submonoid +#align add_submonoid.is_add_submonoid AddSubmonoid.is_addSubmonoid diff --git a/Mathlib/GroupTheory/Congruence.lean b/Mathlib/GroupTheory/Congruence.lean index 7af5527aa6655..6f68952635b42 100644 --- a/Mathlib/GroupTheory/Congruence.lean +++ b/Mathlib/GroupTheory/Congruence.lean @@ -547,6 +547,7 @@ theorem conGen_of_con (c : Con M) : conGen c = c := le_antisymm (by rw [conGen_eq]; exact infₛ_le fun _ _ => id) ConGen.Rel.of #align con.con_gen_of_con Con.conGen_of_con #align add_con.add_con_gen_of_con AddCon.addConGen_of_addCon +#align add_con.add_con_gen_of_add_con AddCon.addConGen_of_addCon --Porting note: removing simp, simp can prove it /-- The map sending a binary relation to the smallest congruence relation in which it is @@ -1104,6 +1105,8 @@ def quotientKerEquivOfRightInverse (f : M →* P) (g : P → M) (hf : Function.R right_inv := fun x => by conv_rhs => rw [← hf x]; rfl } #align con.quotient_ker_equiv_of_right_inverse Con.quotientKerEquivOfRightInverse #align add_con.quotient_ker_equiv_of_right_inverse AddCon.quotientKerEquivOfRightInverse +#align con.quotient_ker_equiv_of_right_inverse_symm_apply Con.quotientKerEquivOfRightInverse_symmApply +#align con.quotient_ker_equiv_of_right_inverse_apply Con.quotientKerEquivOfRightInverse_apply /-- The first isomorphism theorem for Monoids in the case of a surjective homomorphism. diff --git a/Mathlib/GroupTheory/FreeGroup.lean b/Mathlib/GroupTheory/FreeGroup.lean index fa983a4d74094..2974770d0d143 100644 --- a/Mathlib/GroupTheory/FreeGroup.lean +++ b/Mathlib/GroupTheory/FreeGroup.lean @@ -721,6 +721,7 @@ def lift : (α → β) ≃ (FreeGroup α →* β) simpa [Lift.aux] using ih) #align free_group.lift FreeGroup.lift #align free_add_group.lift FreeAddGroup.lift +#align free_group.lift_symm_apply FreeGroup.lift_symm_apply variable {f} @@ -865,6 +866,7 @@ def freeGroupCongr {α β} (e : α ≃ β) : FreeGroup α ≃* FreeGroup β map_mul' := MonoidHom.map_mul _ #align free_group.free_group_congr FreeGroup.freeGroupCongr #align free_add_group.free_add_group_congr FreeAddGroup.freeAddGroupCongr +#align free_group.free_group_congr_apply FreeGroup.freeGroupCongr_apply @[to_additive (attr:=simp)] theorem freeGroupCongr_refl : freeGroupCongr (Equiv.refl α) = MulEquiv.refl _ := diff --git a/Mathlib/GroupTheory/GroupAction/Defs.lean b/Mathlib/GroupTheory/GroupAction/Defs.lean index 00b901d7528b5..eef6062586c94 100644 --- a/Mathlib/GroupTheory/GroupAction/Defs.lean +++ b/Mathlib/GroupTheory/GroupAction/Defs.lean @@ -114,6 +114,10 @@ class MulAction (α : Type _) (β : Type _) [Monoid α] extends SMul α β where /-- Associativity of `•` and `*` -/ mul_smul : ∀ (x y : α) (b : β), (x * y) • b = x • y • b #align mul_action MulAction +#align mul_action.ext MulAction.ext +#align add_action.ext_iff AddAction.ext_iff +#align mul_action.ext_iff MulAction.ext_iff +#align add_action.ext AddAction.ext /-! ### (Pre)transitive action @@ -680,6 +684,7 @@ def smulOneHom {M N} [Monoid M] [Monoid N] [MulAction M N] [IsScalarTower M N N] map_mul' x y := by rw [smul_one_mul, smul_smul] #align smul_one_hom smulOneHom #align vadd_zero_hom vaddZeroHom +#align smul_one_hom_apply smulOneHom_apply end CompatibleScalar @@ -750,6 +755,7 @@ def SMulZeroClass.toZeroHom (x : M) : toFun := (· • ·) x map_zero' := smul_zero x #align smul_zero_class.to_zero_hom SMulZeroClass.toZeroHom +#align smul_zero_class.to_zero_hom_apply SMulZeroClass.toZeroHom_apply end smul_zero @@ -762,6 +768,8 @@ class DistribSMul (M A : Type _) [AddZeroClass A] extends SMulZeroClass M A wher /-- Scalar multiplication distributes across addition -/ smul_add : ∀ (a : M) (x y : A), a • (x + y) = a • x + a • y #align distrib_smul DistribSMul +#align distrib_smul.ext DistribSMul.ext +#align distrib_smul.ext_iff DistribSMul.ext_iff section DistribSMul @@ -825,6 +833,7 @@ def DistribSMul.compFun (f : N → M) : DistribSMul N A := def DistribSMul.toAddMonoidHom (x : M) : A →+ A := { SMulZeroClass.toZeroHom A x with toFun := (· • ·) x, map_add' := smul_add x } #align distrib_smul.to_add_monoid_hom DistribSMul.toAddMonoidHom +#align distrib_smul.to_add_monoid_hom_apply DistribSMul.toAddMonoidHom_apply end DistribSMul @@ -836,6 +845,8 @@ class DistribMulAction (M A : Type _) [Monoid M] [AddMonoid A] extends MulAction /-- Scalar multiplication distributes across addition -/ smul_add : ∀ (a : M) (x y : A), a • (x + y) = a • x + a • y #align distrib_mul_action DistribMulAction +#align distrib_mul_action.ext DistribMulAction.ext +#align distrib_mul_action.ext_iff DistribMulAction.ext_iff section @@ -898,6 +909,7 @@ def DistribMulAction.compHom [Monoid N] (f : N →* M) : DistribMulAction N A := def DistribMulAction.toAddMonoidHom (x : M) : A →+ A := DistribSMul.toAddMonoidHom A x #align distrib_mul_action.to_add_monoid_hom DistribMulAction.toAddMonoidHom +#align distrib_mul_action.to_add_monoid_hom_apply DistribMulAction.toAddMonoidHom_apply variable (M) @@ -909,6 +921,7 @@ def DistribMulAction.toAddMonoidEnd : map_one' := AddMonoidHom.ext <| one_smul M map_mul' x y := AddMonoidHom.ext <| mul_smul x y #align distrib_mul_action.to_add_monoid_End DistribMulAction.toAddMonoidEnd +#align distrib_mul_action.to_add_monoid_End_apply DistribMulAction.toAddMonoidEnd_apply instance AddMonoid.nat_smulCommClass : SMulCommClass ℕ M @@ -956,6 +969,8 @@ class MulDistribMulAction (M : Type _) (A : Type _) [Monoid M] [Monoid A] extend /-- Multiplying `1` by a scalar gives `1` -/ smul_one : ∀ r : M, r • (1 : A) = 1 #align mul_distrib_mul_action MulDistribMulAction +#align mul_distrib_mul_action.ext MulDistribMulAction.ext +#align mul_distrib_mul_action.ext_iff MulDistribMulAction.ext_iff export MulDistribMulAction (smul_one) @@ -1032,6 +1047,7 @@ def MulDistribMulAction.toMonoidEnd : map_one' := MonoidHom.ext <| one_smul M map_mul' x y := MonoidHom.ext <| mul_smul x y #align mul_distrib_mul_action.to_monoid_End MulDistribMulAction.toMonoidEnd +#align mul_distrib_mul_action.to_monoid_End_apply MulDistribMulAction.toMonoidEnd_apply end diff --git a/Mathlib/GroupTheory/GroupAction/Group.lean b/Mathlib/GroupTheory/GroupAction/Group.lean index 7d972d28b4bad..c9de0b8f6a669 100644 --- a/Mathlib/GroupTheory/GroupAction/Group.lean +++ b/Mathlib/GroupTheory/GroupAction/Group.lean @@ -52,6 +52,10 @@ def MulAction.toPerm (a : α) : Equiv.Perm β := ⟨fun x => a • x, fun x => a⁻¹ • x, inv_smul_smul a, smul_inv_smul a⟩ #align mul_action.to_perm MulAction.toPerm #align add_action.to_perm AddAction.toPerm +#align add_action.to_perm_apply AddAction.toPerm_apply +#align mul_action.to_perm_apply MulAction.toPerm_apply +#align add_action.to_perm_symm_apply AddAction.toPerm_symm_apply +#align mul_action.to_perm_symm_apply MulAction.toPerm_symm_apply /-- Given an action of an additive group `α` on `β`, each `g : α` defines a permutation of `β`. -/ add_decl_doc AddAction.toPerm @@ -73,6 +77,7 @@ def MulAction.toPermHom : α →* Equiv.Perm β where map_one' := Equiv.ext <| one_smul α map_mul' u₁ u₂ := Equiv.ext <| mul_smul (u₁ : α) u₂ #align mul_action.to_perm_hom MulAction.toPermHom +#align mul_action.to_perm_hom_apply MulAction.toPermHom_apply /-- Given an action of a additive group `α` on a set `β`, each `g : α` defines a permutation of `β`. -/ @@ -247,6 +252,8 @@ This is a stronger version of `MulAction.toPerm`. -/ def DistribMulAction.toAddEquiv (x : α) : β ≃+ β := { DistribMulAction.toAddMonoidHom β x, MulAction.toPermHom α β x with } #align distrib_mul_action.to_add_equiv DistribMulAction.toAddEquiv +#align distrib_mul_action.to_add_equiv_apply DistribMulAction.toAddEquiv_apply +#align distrib_mul_action.to_add_equiv_symm_apply DistribMulAction.toAddEquiv_symmApply variable (α) @@ -259,6 +266,7 @@ def DistribMulAction.toAddAut : α →* AddAut β where map_one' := AddEquiv.ext (one_smul _) map_mul' _ _ := AddEquiv.ext (mul_smul _ _) #align distrib_mul_action.to_add_aut DistribMulAction.toAddAut +#align distrib_mul_action.to_add_aut_apply DistribMulAction.toAddAut_apply variable {α β} @@ -301,6 +309,8 @@ This is a stronger version of `MulAction.toPerm`. -/ def MulDistribMulAction.toMulEquiv (x : α) : β ≃* β := { MulDistribMulAction.toMonoidHom β x, MulAction.toPermHom α β x with } #align mul_distrib_mul_action.to_mul_equiv MulDistribMulAction.toMulEquiv +#align mul_distrib_mul_action.to_mul_equiv_symm_apply MulDistribMulAction.toMulEquiv_symmApply +#align mul_distrib_mul_action.to_mul_equiv_apply MulDistribMulAction.toMulEquiv_apply variable (α) @@ -313,6 +323,7 @@ def MulDistribMulAction.toMulAut : α →* MulAut β where map_one' := MulEquiv.ext (one_smul _) map_mul' _ _ := MulEquiv.ext (mul_smul _ _) #align mul_distrib_mul_action.to_mul_aut MulDistribMulAction.toMulAut +#align mul_distrib_mul_action.to_mul_aut_apply MulDistribMulAction.toMulAut_apply variable {α β} @@ -354,6 +365,8 @@ attribute [local instance] arrowMulDistribMulAction def mulAutArrow {G A H} [Group G] [MulAction G A] [Monoid H] : G →* MulAut (A → H) := MulDistribMulAction.toMulAut _ _ #align mul_aut_arrow mulAutArrow +#align mul_aut_arrow_apply_apply mulAutArrow_apply_apply +#align mul_aut_arrow_apply_symm_apply mulAutArrow_apply_symmApply end Arrow diff --git a/Mathlib/GroupTheory/GroupAction/Prod.lean b/Mathlib/GroupTheory/GroupAction/Prod.lean index 124a4958f12ce..3b681342eccc5 100644 --- a/Mathlib/GroupTheory/GroupAction/Prod.lean +++ b/Mathlib/GroupTheory/GroupAction/Prod.lean @@ -196,6 +196,7 @@ def smulMulHom [Monoid α] [Mul β] [MulAction α β] [IsScalarTower α β β] [ toFun a := a.1 • a.2 map_mul' _ _ := (smul_mul_smul _ _ _ _).symm #align smul_mul_hom smulMulHom +#align smul_mul_hom_apply smulMulHom_apply /-- Scalar multiplication as a monoid homomorphism. -/ @[simps] @@ -203,5 +204,6 @@ def smulMonoidHom [Monoid α] [MulOneClass β] [MulAction α β] [IsScalarTower [SMulCommClass α β β] : α × β →* β := { smulMulHom with map_one' := one_smul _ _ } #align smul_monoid_hom smulMonoidHom +#align smul_monoid_hom_apply smulMonoidHom_apply end BundledSMul diff --git a/Mathlib/GroupTheory/IsFreeGroup.lean b/Mathlib/GroupTheory/IsFreeGroup.lean index d7eb72eb79da6..447c985bdd46c 100644 --- a/Mathlib/GroupTheory/IsFreeGroup.lean +++ b/Mathlib/GroupTheory/IsFreeGroup.lean @@ -69,6 +69,8 @@ def MulEquiv : FreeGroup (Generators G) ≃* G := IsFreeGroup.MulEquiv' def toFreeGroup : G ≃* FreeGroup (Generators G) := (MulEquiv G).symm #align is_free_group.to_free_group IsFreeGroup.toFreeGroup +#align is_free_group.to_free_group_apply IsFreeGroup.toFreeGroup_apply +#align is_free_group.to_free_group_symm_apply IsFreeGroup.toFreeGroup_symmApply variable {G} diff --git a/Mathlib/GroupTheory/Perm/Basic.lean b/Mathlib/GroupTheory/Perm/Basic.lean index 89b4eb5c676b3..9702db06520d3 100644 --- a/Mathlib/GroupTheory/Perm/Basic.lean +++ b/Mathlib/GroupTheory/Perm/Basic.lean @@ -55,6 +55,8 @@ def equivUnitsEnd : Perm α ≃* Units (Function.End α) where right_inv _ := Units.ext rfl map_mul' _ _ := rfl #align equiv.perm.equiv_units_End Equiv.Perm.equivUnitsEnd +#align equiv.perm.equiv_units_End_symm_apply_apply Equiv.Perm.equivUnitsEnd_symmApply_apply +#align equiv.perm.equiv_units_End_symm_apply_symm_apply Equiv.Perm.equivUnitsEnd_symmApply_symm_apply /-- Lift a monoid homomorphism `f : G →* function.End α` to a monoid homomorphism `f : G →* equiv.perm α`. -/ @@ -62,6 +64,8 @@ def equivUnitsEnd : Perm α ≃* Units (Function.End α) where def _root_.MonoidHom.toHomPerm {G : Type _} [Group G] (f : G →* Function.End α) : G →* Perm α := equivUnitsEnd.symm.toMonoidHom.comp f.toHomUnits #align monoid_hom.to_hom_perm MonoidHom.toHomPerm +#align monoid_hom.to_hom_perm_apply_symm_apply MonoidHom.toHomPerm_apply_symm_apply +#align monoid_hom.to_hom_perm_apply_apply MonoidHom.toHomPerm_apply_apply theorem mul_apply (f g : Perm α) (x) : (f * g) x = f (g x) := Equiv.trans_apply _ _ _ @@ -208,6 +212,7 @@ def sumCongrHom (α β : Type _) : Perm α × Perm β →* Perm (Sum α β) wher map_one' := sumCongr_one map_mul' _ _ := (sumCongr_mul _ _ _ _).symm #align equiv.perm.sum_congr_hom Equiv.Perm.sumCongrHom +#align equiv.perm.sum_congr_hom_apply Equiv.Perm.sumCongrHom_apply theorem sumCongrHom_injective {α β : Type _} : Function.Injective (sumCongrHom α β) := by rintro ⟨⟩ ⟨⟩ h @@ -260,6 +265,7 @@ def sigmaCongrRightHom {α : Type _} (β : α → Type _) : (∀ a, Perm (β a)) map_one' := sigmaCongrRight_one map_mul' _ _ := (sigmaCongrRight_mul _ _).symm #align equiv.perm.sigma_congr_right_hom Equiv.Perm.sigmaCongrRightHom +#align equiv.perm.sigma_congr_right_hom_apply Equiv.Perm.sigmaCongrRightHom_apply theorem sigmaCongrRightHom_injective {α : Type _} {β : α → Type _} : Function.Injective (sigmaCongrRightHom β) := by @@ -276,6 +282,7 @@ def subtypeCongrHom (p : α → Prop) [DecidablePred p] : map_one' := Perm.subtypeCongr.refl map_mul' _ _ := (Perm.subtypeCongr.trans _ _ _ _).symm #align equiv.perm.subtype_congr_hom Equiv.Perm.subtypeCongrHom +#align equiv.perm.subtype_congr_hom_apply Equiv.Perm.subtypeCongrHom_apply theorem subtypeCongrHom_injective (p : α → Prop) [DecidablePred p] : Function.Injective (subtypeCongrHom p) := by @@ -321,6 +328,7 @@ def extendDomainHom : Perm α →* Perm β where map_one' := extendDomain_one f map_mul' e e' := (extendDomain_mul f e e').symm #align equiv.perm.extend_domain_hom Equiv.Perm.extendDomainHom +#align equiv.perm.extend_domain_hom_apply Equiv.Perm.extendDomainHom_apply theorem extendDomainHom_injective : Function.Injective (extendDomainHom f) := (injective_iff_map_eq_one (extendDomainHom f)).mpr fun e he => @@ -470,6 +478,8 @@ protected def subtypeEquivSubtypePerm (p : α → Prop) [DecidablePred p] : right_inv f := Subtype.ext ((Equiv.Perm.ofSubtype_subtypePerm _) fun a => Not.decidable_imp_symm <| f.prop a) #align equiv.perm.subtype_equiv_subtype_perm Equiv.Perm.subtypeEquivSubtypePerm +#align equiv.perm.subtype_equiv_subtype_perm_symm_apply Equiv.Perm.subtypeEquivSubtypePerm_symm_apply +#align equiv.perm.subtype_equiv_subtype_perm_apply_coe Equiv.Perm.subtypeEquivSubtypePerm_apply_coe theorem subtypeEquivSubtypePerm_apply_of_mem (f : Perm (Subtype p)) (h : p a) : -- Porting note: was `Perm.subtypeEquivSubtypePerm p f a` @@ -584,28 +594,38 @@ variable [AddGroup α] (a b : α) @[simp] lemma addLeft_zero : Equiv.addLeft (0 : α) = 1 := ext zero_add @[simp] lemma addRight_zero : Equiv.addRight (0 : α) = 1 := ext add_zero +#align equiv.add_right_zero Equiv.addRight_zero +#align equiv.add_left_zero Equiv.addLeft_zero @[simp] lemma addLeft_add : Equiv.addLeft (a + b) = Equiv.addLeft a * Equiv.addLeft b := ext $ add_assoc _ _ +#align equiv.add_left_add Equiv.addLeft_add @[simp] lemma addRight_add : Equiv.addRight (a + b) = Equiv.addRight b * Equiv.addRight a := ext $ fun _ ↦ (add_assoc _ _ _).symm +#align equiv.add_right_add Equiv.addRight_add @[simp] lemma inv_addLeft : (Equiv.addLeft a)⁻¹ = Equiv.addLeft (-a) := Equiv.coe_inj.1 rfl @[simp] lemma inv_addRight : (Equiv.addRight a)⁻¹ = Equiv.addRight (-a) := Equiv.coe_inj.1 rfl +#align equiv.inv_add_right Equiv.inv_addRight +#align equiv.inv_add_left Equiv.inv_addLeft @[simp] lemma pow_addLeft (n : ℕ) : Equiv.addLeft a ^ n = Equiv.addLeft (n • a) := by ext; simp [Perm.coe_pow] +#align equiv.pow_add_left Equiv.pow_addLeft @[simp] lemma pow_addRight (n : ℕ) : Equiv.addRight a ^ n = Equiv.addRight (n • a) := by ext; simp [Perm.coe_pow] +#align equiv.pow_add_right Equiv.pow_addRight @[simp] lemma zpow_addLeft (n : ℤ) : Equiv.addLeft a ^ n = Equiv.addLeft (n • a) := (map_zsmul (⟨⟨Equiv.addLeft, addLeft_zero⟩, addLeft_add⟩ : α →+ Additive (Perm α)) _ _).symm +#align equiv.zpow_add_left Equiv.zpow_addLeft @[simp] lemma zpow_addRight : ∀ (n : ℤ), Equiv.addRight a ^ n = Equiv.addRight (n • a) | (Int.ofNat n) => by simp | (Int.negSucc n) => by simp +#align equiv.zpow_add_right Equiv.zpow_addRight end AddGroup @@ -614,36 +634,46 @@ variable [Group α] (a b : α) @[simp, to_additive] lemma mulLeft_one : Equiv.mulLeft (1 : α) = 1 := ext one_mul @[simp, to_additive] lemma mulRight_one : Equiv.mulRight (1 : α) = 1 := ext mul_one +#align equiv.mul_right_one Equiv.mulRight_one +#align equiv.mul_left_one Equiv.mulLeft_one @[simp, to_additive] lemma mulLeft_mul : Equiv.mulLeft (a * b) = Equiv.mulLeft a * Equiv.mulLeft b := ext $ mul_assoc _ _ +#align equiv.mul_left_mul Equiv.mulLeft_mul @[simp, to_additive] lemma mulRight_mul : Equiv.mulRight (a * b) = Equiv.mulRight b * Equiv.mulRight a := ext $ fun _ ↦ (mul_assoc _ _ _).symm +#align equiv.mul_right_mul Equiv.mulRight_mul @[simp, to_additive inv_addLeft] lemma inv_mulLeft : (Equiv.mulLeft a)⁻¹ = Equiv.mulLeft a⁻¹ := Equiv.coe_inj.1 rfl @[simp, to_additive inv_addRight] lemma inv_mulRight : (Equiv.mulRight a)⁻¹ = Equiv.mulRight a⁻¹ := Equiv.coe_inj.1 rfl +#align equiv.inv_mul_right Equiv.inv_mulRight +#align equiv.inv_mul_left Equiv.inv_mulLeft @[simp, to_additive pow_addLeft] lemma pow_mulLeft (n : ℕ) : Equiv.mulLeft a ^ n = Equiv.mulLeft (a ^ n) := by ext; simp [Perm.coe_pow] +#align equiv.pow_mul_left Equiv.pow_mulLeft @[simp, to_additive pow_addRight] lemma pow_mulRight (n : ℕ) : Equiv.mulRight a ^ n = Equiv.mulRight (a ^ n) := by ext; simp [Perm.coe_pow] +#align equiv.pow_mul_right Equiv.pow_mulRight @[simp, to_additive zpow_addLeft] lemma zpow_mulLeft (n : ℤ) : Equiv.mulLeft a ^ n = Equiv.mulLeft (a ^ n) := (map_zpow (⟨⟨Equiv.mulLeft, mulLeft_one⟩, mulLeft_mul⟩ : α →* Perm α) _ _).symm +#align equiv.zpow_mul_left Equiv.zpow_mulLeft @[simp, to_additive zpow_addRight] lemma zpow_mulRight : ∀ n : ℤ, Equiv.mulRight a ^ n = Equiv.mulRight (a ^ n) | (Int.ofNat n) => by simp | (Int.negSucc n) => by simp +#align equiv.zpow_mul_right Equiv.zpow_mulRight end Group end Equiv diff --git a/Mathlib/GroupTheory/Subgroup/Basic.lean b/Mathlib/GroupTheory/Subgroup/Basic.lean index 42a20970545dd..28dfe00bd3e04 100644 --- a/Mathlib/GroupTheory/Subgroup/Basic.lean +++ b/Mathlib/GroupTheory/Subgroup/Basic.lean @@ -378,9 +378,11 @@ attribute [to_additive] Subgroup /-- Reinterpret a `Subgroup` as a `Submonoid`. -/ add_decl_doc Subgroup.toSubmonoid +#align subgroup.to_submonoid Subgroup.toSubmonoid /-- Reinterpret an `AddSubgroup` as an `AddSubmonoid`. -/ add_decl_doc AddSubgroup.toAddSubmonoid +#align add_subgroup.to_add_submonoid AddSubgroup.toAddSubmonoid namespace Subgroup @@ -513,6 +515,8 @@ def Subgroup.toAddSubgroup : Subgroup G ≃o AddSubgroup (Additive G) where right_inv x := by cases x; rfl map_rel_iff' := Iff.rfl #align subgroup.to_add_subgroup Subgroup.toAddSubgroup +#align subgroup.to_add_subgroup_symm_apply_coe Subgroup.toAddSubgroup_symmApply_coe +#align subgroup.to_add_subgroup_apply_coe Subgroup.toAddSubgroup_apply_coe /-- Additive subgroup of an additive group `Additive G` are isomorphic to subgroup of `G`. -/ abbrev AddSubgroup.toSubgroup' : AddSubgroup (Additive G) ≃o Subgroup G := @@ -529,6 +533,8 @@ def AddSubgroup.toSubgroup : AddSubgroup A ≃o Subgroup (Multiplicative A) wher right_inv x := by cases x; rfl map_rel_iff' := Iff.rfl #align add_subgroup.to_subgroup AddSubgroup.toSubgroup +#align add_subgroup.to_subgroup_apply_coe AddSubgroup.toSubgroup_apply_coe +#align add_subgroup.to_subgroup_symm_apply_coe AddSubgroup.toSubgroup_symmApply_coe /-- Subgroups of an additive group `Multiplicative A` are isomorphic to additive subgroups of `A`. -/ @@ -859,6 +865,8 @@ def topEquiv : (⊤ : Subgroup G) ≃* G := Submonoid.topEquiv #align subgroup.top_equiv Subgroup.topEquiv #align add_subgroup.top_equiv AddSubgroup.topEquiv +#align subgroup.top_equiv_symm_apply_coe Subgroup.topEquiv_symmApply_coe +#align subgroup.top_equiv_apply Subgroup.topEquiv_apply /-- The trivial subgroup `{1}` of an group `G`. -/ @[to_additive "The trivial `AddSubgroup` `{0}` of an `AddGroup` `G`."] @@ -1611,6 +1619,8 @@ def subgroupOfEquivOfLe {G : Type _} [Group G] {H K : Subgroup G} (h : H ≤ K) map_mul' _g _h := rfl #align subgroup.subgroup_of_equiv_of_le Subgroup.subgroupOfEquivOfLe #align add_subgroup.add_subgroup_of_equiv_of_le AddSubgroup.addSubgroupOfEquivOfLe +#align subgroup.subgroup_of_equiv_of_le_symm_apply_coe_coe Subgroup.subgroupOfEquivOfLe_symmApply_coe_coe +#align subgroup.subgroup_of_equiv_of_le_apply_coe Subgroup.subgroupOfEquivOfLe_apply_coe @[to_additive (attr := simp)] theorem comap_subtype (H K : Subgroup G) : H.comap K.subtype = H.subgroupOf K := @@ -3437,6 +3447,8 @@ def subgroupComap (f : G →* G') (H' : Subgroup G') : H'.comap f →* H' := f.submonoidComap H'.toSubmonoid #align monoid_hom.subgroup_comap MonoidHom.subgroupComap #align add_monoid_hom.add_subgroup_comap AddMonoidHom.addSubgroupComap +#align add_monoid_hom.add_subgroup_comap_apply_coe AddMonoidHom.addSubgroupComap_apply_coe +#align monoid_hom.subgroup_comap_apply_coe MonoidHom.subgroupComap_apply_coe /-- The `MonoidHom` from a subgroup to its image. -/ @[to_additive (attr := simps) "the `add_monoid_hom` from an additive subgroup to its image"] @@ -3444,6 +3456,8 @@ def subgroupMap (f : G →* G') (H : Subgroup G) : H →* H.map f := f.submonoidMap H.toSubmonoid #align monoid_hom.subgroup_map MonoidHom.subgroupMap #align add_monoid_hom.add_subgroup_map AddMonoidHom.addSubgroupMap +#align add_monoid_hom.add_subgroup_map_apply_coe AddMonoidHom.addSubgroupMap_apply_coe +#align monoid_hom.subgroup_map_apply_coe MonoidHom.subgroupMap_apply_coe @[to_additive] theorem subgroupMap_surjective (f : G →* G') (H : Subgroup G) : diff --git a/Mathlib/GroupTheory/Subgroup/MulOpposite.lean b/Mathlib/GroupTheory/Subgroup/MulOpposite.lean index 6d9856de5cd23..abfc5e23725b0 100644 --- a/Mathlib/GroupTheory/Subgroup/MulOpposite.lean +++ b/Mathlib/GroupTheory/Subgroup/MulOpposite.lean @@ -49,6 +49,7 @@ def oppositeEquiv (H : Subgroup G) : H ≃ opposite H := MulOpposite.opEquiv.subtypeEquiv fun _ => Iff.rfl #align subgroup.opposite_equiv Subgroup.oppositeEquiv #align add_subgroup.opposite_equiv AddSubgroup.oppositeEquiv +#align subgroup.opposite_equiv_symm_apply_coe Subgroup.oppositeEquiv_symm_apply_coe @[to_additive] instance (H : Subgroup G) [Encodable H] : Encodable (opposite H) := diff --git a/Mathlib/GroupTheory/Subgroup/Zpowers.lean b/Mathlib/GroupTheory/Subgroup/Zpowers.lean index 5b6ae5c861f66..cb3acadf49978 100644 --- a/Mathlib/GroupTheory/Subgroup/Zpowers.lean +++ b/Mathlib/GroupTheory/Subgroup/Zpowers.lean @@ -98,28 +98,38 @@ theorem range_zmultiplesHom (a : A) : (zmultiplesHom A a).range = zmultiples a : attribute [to_additive AddSubgroup.zmultiples] Subgroup.zpowers attribute [to_additive (attr := simp) AddSubgroup.mem_zmultiples] Subgroup.mem_zpowers +#align add_subgroup.mem_zmultiples AddSubgroup.mem_zmultiples attribute [to_additive AddSubgroup.zmultiples_eq_closure] Subgroup.zpowers_eq_closure +#align add_subgroup.zmultiples_eq_closure AddSubgroup.zmultiples_eq_closure attribute [to_additive (attr := simp) AddSubgroup.range_zmultiplesHom] Subgroup.range_zpowersHom attribute [to_additive AddSubgroup.zmultiples_subset] Subgroup.zpowers_subset +#align add_subgroup.zmultiples_subset AddSubgroup.zmultiples_subset attribute [to_additive AddSubgroup.mem_zmultiples_iff] Subgroup.mem_zpowers_iff +#align add_subgroup.mem_zmultiples_iff AddSubgroup.mem_zmultiples_iff attribute [to_additive (attr := simp) AddSubgroup.zsmul_mem_zmultiples] Subgroup.zpow_mem_zpowers +#align add_subgroup.zsmul_mem_zmultiples AddSubgroup.zsmul_mem_zmultiples attribute [to_additive (attr := simp) AddSubgroup.nsmul_mem_zmultiples] Subgroup.npow_mem_zpowers +#align add_subgroup.nsmul_mem_zmultiples AddSubgroup.nsmul_mem_zmultiples --Porting note: increasing simp priority. Better lemma than `Subtype.forall` attribute [to_additive (attr := simp 1100) AddSubgroup.forall_zmultiples] Subgroup.forall_zpowers +#align add_subgroup.forall_zmultiples AddSubgroup.forall_zmultiples attribute [to_additive AddSubgroup.forall_mem_zmultiples] Subgroup.forall_mem_zpowers +#align add_subgroup.forall_mem_zmultiples AddSubgroup.forall_mem_zmultiples --Porting note: increasing simp priority. Better lemma than `Subtype.exists` attribute [to_additive (attr := simp 1100) AddSubgroup.exists_zmultiples] Subgroup.exists_zpowers +#align add_subgroup.exists_zmultiples AddSubgroup.exists_zmultiples attribute [to_additive AddSubgroup.exists_mem_zmultiples] Subgroup.exists_mem_zpowers +#align add_subgroup.exists_mem_zmultiples AddSubgroup.exists_mem_zmultiples instance (a : A) : Countable (zmultiples a) := (zmultiplesHom A a).rangeRestrict_surjective.countable diff --git a/Mathlib/GroupTheory/Submonoid/Basic.lean b/Mathlib/GroupTheory/Submonoid/Basic.lean index f7c71a075abc1..a8de9915c9460 100644 --- a/Mathlib/GroupTheory/Submonoid/Basic.lean +++ b/Mathlib/GroupTheory/Submonoid/Basic.lean @@ -101,6 +101,7 @@ end /-- A submonoid of a monoid `M` can be considered as a subsemigroup of that monoid. -/ add_decl_doc Submonoid.toSubsemigroup +#align submonoid.to_subsemigroup Submonoid.toSubsemigroup /-- `SubmonoidClass S M` says `S` is a type of subsets `s ≤ M` that contain `1` and are closed under `(*)` -/ @@ -122,6 +123,7 @@ end /-- An additive submonoid of an additive monoid `M` can be considered as an additive subsemigroup of that additive monoid. -/ add_decl_doc AddSubmonoid.toAddSubsemigroup +#align add_submonoid.to_add_subsemigroup AddSubmonoid.toAddSubsemigroup /-- `AddSubmonoidClass S M` says `S` is a type of subsets `s ≤ M` that contain `0` and are closed under `(+)` -/ diff --git a/Mathlib/GroupTheory/Submonoid/Membership.lean b/Mathlib/GroupTheory/Submonoid/Membership.lean index 2e34d74e990ce..1e81f0d50230c 100644 --- a/Mathlib/GroupTheory/Submonoid/Membership.lean +++ b/Mathlib/GroupTheory/Submonoid/Membership.lean @@ -464,6 +464,7 @@ theorem powers_one : powers (1 : M) = ⊥ := def pow (n : M) (m : ℕ) : powers n := (powersHom M n).mrangeRestrict (Multiplicative.ofAdd m) #align submonoid.pow Submonoid.pow +#align submonoid.pow_coe Submonoid.pow_coe theorem pow_apply (n : M) (m : ℕ) : Submonoid.pow n m = ⟨n ^ m, m, rfl⟩ := rfl @@ -502,6 +503,8 @@ def powLogEquiv [DecidableEq M] {n : M} (h : Function.Injective fun m : ℕ => n right_inv := pow_log_eq_self map_mul' _ _ := by simp only [pow, map_mul, ofAdd_add, toAdd_mul] #align submonoid.pow_log_equiv Submonoid.powLogEquiv +#align submonoid.pow_log_equiv_symm_apply Submonoid.powLogEquiv_symmApply +#align submonoid.pow_log_equiv_apply Submonoid.powLogEquiv_apply theorem log_mul [DecidableEq M] {n : M} (h : Function.Injective fun m : ℕ => n ^ m) (x y : powers (n : M)) : log (x * y) = log x + log y := diff --git a/Mathlib/GroupTheory/Submonoid/Operations.lean b/Mathlib/GroupTheory/Submonoid/Operations.lean index 0f3812cc5abe6..e08f86d07a0c4 100644 --- a/Mathlib/GroupTheory/Submonoid/Operations.lean +++ b/Mathlib/GroupTheory/Submonoid/Operations.lean @@ -92,6 +92,8 @@ def Submonoid.toAddSubmonoid : Submonoid M ≃o AddSubmonoid (Additive M) where right_inv x := by cases x; rfl map_rel_iff' := Iff.rfl #align submonoid.to_add_submonoid Submonoid.toAddSubmonoid +#align submonoid.to_add_submonoid_symm_apply_coe Submonoid.toAddSubmonoid_symmApply_coe +#align submonoid.to_add_submonoid_apply_coe Submonoid.toAddSubmonoid_apply_coe /-- Additive submonoids of an additive monoid `Additive M` are isomorphic to submonoids of `M`. -/ abbrev AddSubmonoid.toSubmonoid' : AddSubmonoid (Additive M) ≃o Submonoid M := @@ -138,6 +140,8 @@ def AddSubmonoid.toSubmonoid : AddSubmonoid A ≃o Submonoid (Multiplicative A) right_inv x := by cases x; rfl map_rel_iff' := Iff.rfl #align add_submonoid.to_submonoid AddSubmonoid.toSubmonoid +#align add_submonoid.to_submonoid_symm_apply_coe AddSubmonoid.toSubmonoid_symmApply_coe +#align add_submonoid.to_submonoid_apply_coe AddSubmonoid.toSubmonoid_apply_coe /-- Submonoids of a monoid `Multiplicative A` are isomorphic to additive submonoids of `A`. -/ abbrev Submonoid.toAddSubmonoid' : Submonoid (Multiplicative A) ≃o AddSubmonoid A := @@ -797,6 +801,8 @@ def topEquiv : (⊤ : Submonoid M) ≃* M where map_mul' _ _ := rfl #align submonoid.top_equiv Submonoid.topEquiv #align add_submonoid.top_equiv AddSubmonoid.topEquiv +#align submonoid.top_equiv_apply Submonoid.topEquiv_apply +#align submonoid.top_equiv_symm_apply_coe Submonoid.topEquiv_symmApply_coe @[to_additive (attr := simp)] theorem top_equiv_toMonoidHom : (topEquiv : _ ≃* M).toMonoidHom = (⊤ : Submonoid M).subtype := @@ -1122,6 +1128,7 @@ def codRestrict {S} [SetLike S N] [SubmonoidClass S N] (f : M →* N) (s : S) (h map_mul' x y := Subtype.eq (f.map_mul x y) #align monoid_hom.cod_restrict MonoidHom.codRestrict #align add_monoid_hom.cod_restrict AddMonoidHom.codRestrict +#align monoid_hom.cod_restrict_apply MonoidHom.codRestrict_apply /-- Restriction of a monoid hom to its range interpreted as a submonoid. -/ @[to_additive @@ -1246,6 +1253,7 @@ def submonoidComap (f : M →* N) (N' : Submonoid N) : map_mul' x y := Subtype.eq (f.map_mul x y) #align monoid_hom.submonoid_comap MonoidHom.submonoidComap #align add_monoid_hom.add_submonoid_comap AddMonoidHom.addSubmonoidComap +#align monoid_hom.submonoid_comap_apply_coe MonoidHom.submonoidComap_apply_coe /-- The `MonoidHom` from a submonoid to its image. See `MulEquiv.SubmonoidMap` for a variant for `MulEquiv`s. -/ @@ -1260,6 +1268,7 @@ def submonoidMap (f : M →* N) (M' : Submonoid M) : map_mul' x y := Subtype.eq <| f.map_mul x y #align monoid_hom.submonoid_map MonoidHom.submonoidMap #align add_monoid_hom.add_submonoid_map AddMonoidHom.addSubmonoidMap +#align monoid_hom.submonoid_map_apply_coe MonoidHom.submonoidMap_apply_coe @[to_additive] theorem submonoidMap_surjective (f : M →* N) (M' : Submonoid M) : @@ -1420,6 +1429,8 @@ def ofLeftInverse' (f : M →* N) {g : N → M} (h : Function.LeftInverse g f) : show f (g x) = x by rw [← hx', h x'] } #align mul_equiv.of_left_inverse' MulEquiv.ofLeftInverse' #align add_equiv.of_left_inverse' AddEquiv.ofLeftInverse' +#align mul_equiv.of_left_inverse'_apply MulEquiv.ofLeftInverse'_apply +#align mul_equiv.of_left_inverse'_symm_apply MulEquiv.ofLeftInverse'_symmApply /-- A `MulEquiv` `φ` between two monoids `M` and `N` induces a `MulEquiv` between a submonoid `S ≤ M` and the submonoid `φ(S) ≤ N`. diff --git a/Mathlib/GroupTheory/Subsemigroup/Operations.lean b/Mathlib/GroupTheory/Subsemigroup/Operations.lean index a5a8d5799d149..dd0360f4ce74c 100644 --- a/Mathlib/GroupTheory/Subsemigroup/Operations.lean +++ b/Mathlib/GroupTheory/Subsemigroup/Operations.lean @@ -92,6 +92,8 @@ def Subsemigroup.toAddSubsemigroup : Subsemigroup M ≃o AddSubsemigroup (Additi right_inv _ := rfl map_rel_iff':= Iff.rfl #align subsemigroup.to_add_subsemigroup Subsemigroup.toAddSubsemigroup +#align subsemigroup.to_add_subsemigroup_symm_apply_coe Subsemigroup.toAddSubsemigroup_symmApply_coe +#align subsemigroup.to_add_subsemigroup_apply_coe Subsemigroup.toAddSubsemigroup_apply_coe /-- Additive subsemigroups of an additive semigroup `Additive M` are isomorphic to subsemigroups of `M`. -/ @@ -137,6 +139,8 @@ def AddSubsemigroup.toSubsemigroup : AddSubsemigroup A ≃o Subsemigroup (Multip right_inv _ := rfl map_rel_iff' := Iff.rfl #align add_subsemigroup.to_subsemigroup AddSubsemigroup.toSubsemigroup +#align add_subsemigroup.to_subsemigroup_apply_coe AddSubsemigroup.toSubsemigroup_apply_coe +#align add_subsemigroup.to_subsemigroup_symm_apply_coe AddSubsemigroup.toSubsemigroup_symmApply_coe /-- Subsemigroups of a semigroup `Multiplicative A` are isomorphic to additive subsemigroups of `A`. -/ @@ -591,6 +595,8 @@ def topEquiv : (⊤ : Subsemigroup M) ≃* M where map_mul' _ _ := rfl #align subsemigroup.top_equiv Subsemigroup.topEquiv #align add_subsemigroup.top_equiv AddSubsemigroup.topEquiv +#align subsemigroup.top_equiv_symm_apply_coe Subsemigroup.topEquiv_symmApply_coe +#align subsemigroup.top_equiv_apply Subsemigroup.topEquiv_apply @[to_additive (attr := simp)] theorem topEquiv_toMulHom : @@ -827,6 +833,7 @@ def codRestrict [SetLike σ N] [MulMemClass σ N] (f : M →ₙ* N) (S : σ) (h map_mul' x y := Subtype.eq (map_mul f x y) #align mul_hom.cod_restrict MulHom.codRestrict #align add_hom.cod_restrict AddHom.codRestrict +#align mul_hom.cod_restrict_apply_coe MulHom.codRestrict_apply_coe /-- Restriction of a semigroup hom to its range interpreted as a subsemigroup. -/ @[to_additive "Restriction of an `AddSemigroup` hom to its range interpreted as a subsemigroup."] @@ -863,6 +870,7 @@ def subsemigroupComap (f : M →ₙ* N) (N' : Subsemigroup N) : map_mul' x y := Subtype.eq (@map_mul M N _ _ _ _ f x y) #align mul_hom.subsemigroup_comap MulHom.subsemigroupComap #align add_hom.subsemigroup_comap AddHom.subsemigroupComap +#align mul_hom.subsemigroup_comap_apply_coe MulHom.subsemigroupComap_apply_coe /-- The `MulHom` from a subsemigroup to its image. See `MulEquiv.subsemigroupMap` for a variant for `MulEquiv`s. -/ @@ -876,6 +884,7 @@ def subsemigroupMap (f : M →ₙ* N) (M' : Subsemigroup M) : map_mul' x y := Subtype.eq <| @map_mul M N _ _ _ _ f x y #align mul_hom.subsemigroup_map MulHom.subsemigroupMap #align add_hom.subsemigroup_map AddHom.subsemigroupMap +#align mul_hom.subsemigroup_map_apply_coe MulHom.subsemigroupMap_apply_coe @[to_additive] theorem subsemigroupMap_surjective (f : M →ₙ* N) (M' : Subsemigroup M) : @@ -969,6 +978,8 @@ def ofLeftInverse (f : M →ₙ* N) {g : N → M} (h : Function.LeftInverse g f) show f (g x) = x by rw [← hx', h x'] } #align mul_equiv.of_left_inverse MulEquiv.ofLeftInverse #align add_equiv.of_left_inverse AddEquiv.ofLeftInverse +#align mul_equiv.of_left_inverse_apply MulEquiv.ofLeftInverse_apply +#align mul_equiv.of_left_inverse_symm_apply MulEquiv.ofLeftInverse_symmApply /-- A `MulEquiv` `φ` between two semigroups `M` and `N` induces a `MulEquiv` between a subsemigroup `S ≤ M` and the subsemigroup `φ(S) ≤ N`. @@ -986,5 +997,7 @@ def subsemigroupMap (e : M ≃* N) (S : Subsemigroup M) : S ≃* S.map e.toMulHo invFun := fun x => ⟨e.symm x, _⟩ } #align mul_equiv.subsemigroup_map MulEquiv.subsemigroupMap #align add_equiv.subsemigroup_map AddEquiv.subsemigroupMap +#align mul_equiv.subsemigroup_map_apply_coe MulEquiv.subsemigroupMap_apply_coe +#align mul_equiv.subsemigroup_map_symm_apply_coe MulEquiv.subsemigroupMap_symmApply_coe end MulEquiv diff --git a/Mathlib/Init/CcLemmas.lean b/Mathlib/Init/CcLemmas.lean index 797b696eaaa46..b0fc41d68ff7e 100644 --- a/Mathlib/Init/CcLemmas.lean +++ b/Mathlib/Init/CcLemmas.lean @@ -119,8 +119,10 @@ theorem ne_of_eq_of_ne {α : Sort u} {a b c : α} (h₁ : a = b) (h₂ : b ≠ c h₁.symm ▸ h₂ alias ne_of_eq_of_ne ← Eq.trans_ne +#align eq.trans_ne Eq.trans_ne theorem ne_of_ne_of_eq {α : Sort u} {a b c : α} (h₁ : a ≠ b) (h₂ : b = c) : a ≠ c := h₂ ▸ h₁ alias ne_of_ne_of_eq ← Ne.trans_eq +#align ne.trans_eq Ne.trans_eq diff --git a/Mathlib/Init/Classical.lean b/Mathlib/Init/Classical.lean index 53f836d694011..ad99233b5f905 100644 --- a/Mathlib/Init/Classical.lean +++ b/Mathlib/Init/Classical.lean @@ -29,6 +29,7 @@ theorem cases_on (a : Prop) {p : Prop → Prop} (h1 : p True) (h2 : p False) : p @cases_true_false p h1 h2 a theorem cases {p : Prop → Prop} (h1 : p True) (h2 : p False) (a) : p a := cases_on a h1 h2 +#align classical.cases Classical.cases alias byCases ← by_cases alias byContradiction ← by_contradiction diff --git a/Mathlib/Init/Logic.lean b/Mathlib/Init/Logic.lean index ed13358878a2a..3b19e44501b24 100644 --- a/Mathlib/Init/Logic.lean +++ b/Mathlib/Init/Logic.lean @@ -243,11 +243,13 @@ theorem exists_unique_of_exists_of_unique {α : Sort u} {p : α → Prop} theorem ExistsUnique.exists {p : α → Prop} : (∃! x, p x) → ∃ x, p x | ⟨x, h, _⟩ => ⟨x, h⟩ #align exists_of_exists_unique ExistsUnique.exists +#align exists_unique.exists ExistsUnique.exists theorem ExistsUnique.unique {α : Sort u} {p : α → Prop} (h : ∃! x, p x) {y₁ y₂ : α} (py₁ : p y₁) (py₂ : p y₂) : y₁ = y₂ := let ⟨_, _, hy⟩ := h; (hy _ py₁).trans (hy _ py₂).symm #align unique_of_exists_unique ExistsUnique.unique +#align exists_unique.unique ExistsUnique.unique /- exists, forall, exists unique congruences -/ diff --git a/Mathlib/Logic/Basic.lean b/Mathlib/Logic/Basic.lean index 73e5206a8acbe..4ae5ad6810257 100644 --- a/Mathlib/Logic/Basic.lean +++ b/Mathlib/Logic/Basic.lean @@ -47,6 +47,7 @@ attribute [simp] cast_eq cast_heq if it is applied to a large term, so it can be used for elision, as done in the `elide` and `unelide` tactics. -/ @[reducible] def hidden {α : Sort _} {a : α} := a +#align hidden hidden instance (priority := 10) decidableEq_of_subsingleton [Subsingleton α] : DecidableEq α := fun a b ↦ isTrue (Subsingleton.elim a b) @@ -60,31 +61,40 @@ instance (α : Sort _) [Subsingleton α] (p : α → Prop) : Subsingleton (Subty theorem congr_heq {α β γ : Sort _} {f : α → γ} {g : β → γ} {x : α} {y : β} (h₁ : HEq f g) (h₂ : HEq x y) : f x = g y := by cases h₂; cases h₁; rfl +#align congr_heq congr_heq theorem congr_arg_heq {α} {β : α → Sort _} (f : ∀ a, β a) : ∀ {a₁ a₂ : α}, a₁ = a₂ → HEq (f a₁) (f a₂) | _, _, rfl => HEq.rfl +#align congr_arg_heq congr_arg_heq theorem ULift.down_injective {α : Sort _} : Function.Injective (@ULift.down α) | ⟨a⟩, ⟨b⟩, _ => by congr +#align ulift.down_injective ULift.down_injective @[simp] theorem ULift.down_inj {α : Sort _} {a b : ULift α} : a.down = b.down ↔ a = b := ⟨fun h ↦ ULift.down_injective h, fun h ↦ by rw [h]⟩ +#align ulift.down_inj ULift.down_inj theorem PLift.down_injective {α : Sort _} : Function.Injective (@PLift.down α) | ⟨a⟩, ⟨b⟩, _ => by congr +#align plift.down_injective PLift.down_injective @[simp] theorem PLift.down_inj {α : Sort _} {a b : PLift α} : a.down = b.down ↔ a = b := ⟨fun h ↦ PLift.down_injective h, fun h ↦ by rw [h]⟩ +#align plift.down_inj PLift.down_inj @[simp] theorem eq_iff_eq_cancel_left {b c : α} : (∀ {a}, a = b ↔ a = c) ↔ b = c := ⟨fun h ↦ by rw [← h], fun h a ↦ by rw [h]⟩ +#align eq_iff_eq_cancel_left eq_iff_eq_cancel_left @[simp] theorem eq_iff_eq_cancel_right {a b : α} : (∀ {c}, a = c ↔ b = c) ↔ a = b := ⟨fun h ↦ by rw [h], fun h a ↦ by rw [h]⟩ +#align eq_iff_eq_cancel_right eq_iff_eq_cancel_right lemma ne_and_eq_iff_right {α : Sort _} {a b c : α} (h : b ≠ c) : a ≠ b ∧ a = c ↔ a = c := and_iff_right_of_imp (fun h2 => h2.symm ▸ h.symm) +#align ne_and_eq_iff_right ne_and_eq_iff_right /-- Wrapper for adding elementary propositions to the type class systems. Warning: this can easily be abused. See the rest of this docstring for details. @@ -107,6 +117,7 @@ class Fact (p : Prop) : Prop where /-- `Fact.out` contains the unwrapped witness for the fact represented by the instance of `Fact p`. -/ out : p +#align fact Fact library_note "fact non-instances"/-- In most cases, we should not have global instances of `Fact`; typeclass search only reads the head @@ -116,11 +127,14 @@ everywhere. We instead make them as lemmata and make them local instances as req theorem Fact.elim {p : Prop} (h : Fact p) : p := h.1 theorem fact_iff {p : Prop} : Fact p ↔ p := ⟨fun h ↦ h.1, fun h ↦ ⟨h⟩⟩ +#align fact_iff fact_iff +#align fact.elim Fact.elim /-- Swaps two pairs of arguments to a function. -/ @[reducible] def Function.swap₂ {κ₁ : ι₁ → Sort _} {κ₂ : ι₂ → Sort _} {φ : ∀ i₁, κ₁ i₁ → ∀ i₂, κ₂ i₂ → Sort _} (f : ∀ i₁ j₁ i₂ j₂, φ i₁ j₁ i₂ j₂) (i₂ j₂ i₁ j₁) : φ i₁ j₁ i₂ j₂ := f i₁ j₁ i₂ j₂ +#align function.swap₂ Function.swap₂ -- Porting note: these don't work as intended any more -- /-- If `x : α . tac_name` then `x.out : α`. These are definitionally equal, but this can @@ -150,49 +164,65 @@ instance : IsRefl Prop Iff := ⟨Iff.refl⟩ instance : IsTrans Prop Iff := ⟨fun _ _ _ ↦ Iff.trans⟩ alias imp_congr ← Iff.imp +#align iff.imp Iff.imp @[simp] theorem eq_true_eq_id : Eq True = id := by funext _; simp only [true_iff, id.def, eq_iff_iff] +#align eq_true_eq_id eq_true_eq_id #align imp_and_distrib imp_and #align imp_iff_right imp_iff_rightₓ -- reorder implicits #align imp_iff_not imp_iff_notₓ -- reorder implicits @[simp] theorem imp_iff_right_iff : (a → b ↔ b) ↔ a ∨ b := Decidable.imp_iff_right_iff +#align imp_iff_right_iff imp_iff_right_iff @[simp] theorem and_or_imp : a ∧ b ∨ (a → c) ↔ a → b ∨ c := Decidable.and_or_imp +#align and_or_imp and_or_imp /-- Provide modus tollens (`mt`) as dot notation for implications. -/ protected theorem Function.mt : (a → b) → ¬b → ¬a := mt +#align function.mt Function.mt /-! ### Declarations about `not` -/ alias Decidable.em ← dec_em +#align dec_em dec_em theorem dec_em' (p : Prop) [Decidable p] : ¬p ∨ p := (dec_em p).symm +#align dec_em' dec_em' alias Classical.em ← em +#align em em theorem em' (p : Prop) : ¬p ∨ p := (em p).symm +#align em' em' theorem or_not {p : Prop} : p ∨ ¬p := em _ +#align or_not or_not theorem Decidable.eq_or_ne (x y : α) [Decidable (x = y)] : x = y ∨ x ≠ y := dec_em <| x = y +#align decidable.eq_or_ne Decidable.eq_or_ne theorem Decidable.ne_or_eq (x y : α) [Decidable (x = y)] : x ≠ y ∨ x = y := dec_em' <| x = y +#align decidable.ne_or_eq Decidable.ne_or_eq theorem eq_or_ne (x y : α) : x = y ∨ x ≠ y := em <| x = y +#align eq_or_ne eq_or_ne theorem ne_or_eq (x y : α) : x ≠ y ∨ x = y := em' <| x = y +#align ne_or_eq ne_or_eq theorem by_contradiction : (¬p → False) → p := Decidable.by_contradiction #align classical.by_contradiction by_contradiction +#align by_contradiction by_contradiction theorem by_cases {q : Prop} (hpq : p → q) (hnpq : ¬p → q) : q := if hp : p then hpq hp else hnpq hp #align classical.by_cases by_cases alias by_contradiction ← by_contra +#align by_contra by_contra library_note "decidable namespace"/-- In most of mathlib, we use the law of excluded middle (LEM) and the axiom of choice (AC) freely. @@ -218,34 +248,49 @@ classical ones, as these may cause instance mismatch errors later. The left-to-right direction, double negation elimination (DNE), is classically true but not constructively. -/ @[simp] theorem not_not : ¬¬a ↔ a := Decidable.not_not +#align not_not not_not theorem of_not_not : ¬¬a → a := by_contra +#align of_not_not of_not_not theorem not_ne_iff : ¬a ≠ b ↔ a = b := not_not +#align not_ne_iff not_ne_iff theorem of_not_imp {a b : Prop} : ¬(a → b) → a := Decidable.of_not_imp +#align of_not_imp of_not_imp alias Decidable.not_imp_symm ← Not.decidable_imp_symm +#align not.decidable_imp_symm Not.decidable_imp_symm theorem Not.imp_symm : (¬a → b) → ¬b → a := Not.decidable_imp_symm +#align not.imp_symm Not.imp_symm theorem not_imp_comm : ¬a → b ↔ ¬b → a := Decidable.not_imp_comm +#align not_imp_comm not_imp_comm @[simp] theorem not_imp_self : ¬a → a ↔ a := Decidable.not_imp_self +#align not_imp_self not_imp_self theorem Imp.swap : a → b → c ↔ b → a → c := ⟨Function.swap, Function.swap⟩ +#align imp.swap Imp.swap alias not_congr ← Iff.not theorem Iff.not_left (h : a ↔ ¬b) : ¬a ↔ b := h.not.trans not_not theorem Iff.not_right (h : ¬a ↔ b) : a ↔ ¬b := not_not.symm.trans h.not +#align iff.not_right Iff.not_right +#align iff.not_left Iff.not_left +#align iff.not Iff.not /-! ### Declarations about `xor` -/ @[simp] theorem xor_true : Xor' True = Not := by simp [Xor'] +#align xor_true xor_true @[simp] theorem xor_false : Xor' False = id := by ext; simp [Xor'] +#align xor_false xor_false theorem xor_comm (a b) : Xor' a b = Xor' b a := by simp [Xor', and_comm, or_comm] +#align xor_comm xor_comm instance : IsCommutative Prop Xor' := ⟨xor_comm⟩ @@ -254,6 +299,11 @@ instance : IsCommutative Prop Xor' := ⟨xor_comm⟩ @[simp] theorem xor_not_right : Xor' a (¬b) ↔ (a ↔ b) := by by_cases a <;> simp [*] theorem xor_not_not : Xor' (¬a) (¬b) ↔ Xor' a b := by simp [Xor', or_comm, and_comm] protected theorem xor.or (h : Xor' a b) : a ∨ b := h.imp And.left And.left +#align xor.or xor.or +#align xor_not_not xor_not_not +#align xor_not_right xor_not_right +#align xor_not_left xor_not_left +#align xor_self xor_self /-! ### Declarations about `and` -/ @@ -266,6 +316,8 @@ alias and_congr ← Iff.and alias and_rotate ↔ And.rotate _ #align and.congr_right_iff and_congr_right_iff #align and.congr_left_iff and_congr_left_iffₓ -- reorder implicits +#align and.rotate And.rotate +#align iff.and Iff.and theorem and_symm_right (a b : α) (p : Prop) : p ∧ a = b ↔ p ∧ b = a := by simp [eq_comm] theorem and_symm_left (a b : α) (p : Prop) : a = b ∧ p ↔ b = a ∧ p := by simp [eq_comm] @@ -277,44 +329,60 @@ alias or_congr ← Iff.or #align or_congr_right' or_congr_rightₓ -- reorder implicits #align or.right_comm or_right_comm alias or_rotate ↔ Or.rotate _ +#align or.rotate Or.rotate +#align iff.or Iff.or @[deprecated Or.imp] theorem or_of_or_of_imp_of_imp (h₁ : a ∨ b) (h₂ : a → c) (h₃ : b → d) : c ∨ d := Or.imp h₂ h₃ h₁ +#align or_of_or_of_imp_of_imp or_of_or_of_imp_of_imp @[deprecated Or.imp_left] theorem or_of_or_of_imp_left (h₁ : a ∨ c) (h : a → b) : b ∨ c := Or.imp_left h h₁ +#align or_of_or_of_imp_left or_of_or_of_imp_left @[deprecated Or.imp_right] theorem or_of_or_of_imp_right (h₁ : c ∨ a) (h : a → b) : c ∨ b := Or.imp_right h h₁ +#align or_of_or_of_imp_right or_of_or_of_imp_right theorem Or.elim3 {d : Prop} (h : a ∨ b ∨ c) (ha : a → d) (hb : b → d) (hc : c → d) : d := Or.elim h ha fun h₂ ↦ Or.elim h₂ hb hc +#align or.elim3 Or.elim3 theorem Or.imp3 (had : a → d) (hbe : b → e) (hcf : c → f) : a ∨ b ∨ c → d ∨ e ∨ f := Or.imp had <| Or.imp hbe hcf +#align or.imp3 Or.imp3 #align or_imp_distrib or_imp theorem or_iff_not_imp_left : a ∨ b ↔ ¬a → b := Decidable.or_iff_not_imp_left +#align or_iff_not_imp_left or_iff_not_imp_left theorem or_iff_not_imp_right : a ∨ b ↔ ¬b → a := Decidable.or_iff_not_imp_right +#align or_iff_not_imp_right or_iff_not_imp_right theorem not_or_of_imp : (a → b) → ¬a ∨ b := Decidable.not_or_of_imp +#align not_or_of_imp not_or_of_imp -- See Note [decidable namespace] protected theorem Decidable.or_not_of_imp [Decidable a] (h : a → b) : b ∨ ¬a := dite _ (Or.inl ∘ h) Or.inr +#align decidable.or_not_of_imp Decidable.or_not_of_imp theorem or_not_of_imp : (a → b) → b ∨ ¬a := Decidable.or_not_of_imp +#align or_not_of_imp or_not_of_imp theorem imp_iff_not_or : a → b ↔ ¬a ∨ b := Decidable.imp_iff_not_or +#align imp_iff_not_or imp_iff_not_or theorem imp_iff_or_not : b → a ↔ a ∨ ¬b := Decidable.imp_iff_or_not +#align imp_iff_or_not imp_iff_or_not theorem not_imp_not : ¬a → ¬b ↔ b → a := Decidable.not_imp_not +#align not_imp_not not_imp_not /-- Provide the reverse of modus tollens (`mt`) as dot notation for implications. -/ protected theorem Function.mtr : (¬a → ¬b) → b → a := not_imp_not.mp +#align function.mtr Function.mtr #align decidable.or_congr_left Decidable.or_congr_left' #align decidable.or_congr_right Decidable.or_congr_right' @@ -339,9 +407,11 @@ theorem or_congr_right' (h : ¬a → (b ↔ c)) : a ∨ b ↔ a ∨ c := Decidab /-! Declarations about `iff` -/ alias iff_congr ← Iff.iff +#align iff.iff Iff.iff -- @[simp] -- FIXME simp ignores proof rewrites theorem iff_mpr_iff_true_intro (h : P) : Iff.mpr (iff_true_intro h) True.intro = h := rfl +#align iff_mpr_iff_true_intro iff_mpr_iff_true_intro #align decidable.imp_or_distrib Decidable.imp_or @@ -354,24 +424,33 @@ theorem imp_or' : a → b ∨ c ↔ (a → b) ∨ (a → c) := Decidable.imp_or' #align imp_or_distrib' imp_or'ₓ -- universes theorem not_imp : ¬(a → b) ↔ a ∧ ¬b := Decidable.not_imp +#align not_imp not_imp theorem peirce (a b : Prop) : ((a → b) → a) → a := Decidable.peirce _ _ +#align peirce peirce theorem not_iff_not : (¬a ↔ ¬b) ↔ (a ↔ b) := Decidable.not_iff_not +#align not_iff_not not_iff_not theorem not_iff_comm : (¬a ↔ b) ↔ (¬b ↔ a) := Decidable.not_iff_comm +#align not_iff_comm not_iff_comm theorem not_iff : ¬(a ↔ b) ↔ (¬a ↔ b) := Decidable.not_iff +#align not_iff not_iff theorem iff_not_comm : (a ↔ ¬b) ↔ (b ↔ ¬a) := Decidable.iff_not_comm +#align iff_not_comm iff_not_comm theorem iff_iff_and_or_not_and_not : (a ↔ b) ↔ a ∧ b ∨ ¬a ∧ ¬b := Decidable.iff_iff_and_or_not_and_not +#align iff_iff_and_or_not_and_not iff_iff_and_or_not_and_not theorem iff_iff_not_or_and_or_not : (a ↔ b) ↔ (¬a ∨ b) ∧ (a ∨ ¬b) := Decidable.iff_iff_not_or_and_or_not +#align iff_iff_not_or_and_or_not iff_iff_not_or_and_or_not theorem not_and_not_right : ¬(a ∧ ¬b) ↔ a → b := Decidable.not_and_not_right +#align not_and_not_right not_and_not_right #align decidable_of_iff decidable_of_iff #align decidable_of_iff' decidable_of_iff' @@ -390,15 +469,21 @@ theorem not_and_or : ¬(a ∧ b) ↔ ¬a ∨ ¬b := Decidable.not_and #align not_or_distrib not_or theorem or_iff_not_and_not : a ∨ b ↔ ¬(¬a ∧ ¬b) := Decidable.or_iff_not_and_not +#align or_iff_not_and_not or_iff_not_and_not theorem and_iff_not_or_not : a ∧ b ↔ ¬(¬a ∨ ¬b) := Decidable.and_iff_not_or_not +#align and_iff_not_or_not and_iff_not_or_not @[simp] theorem not_xor (P Q : Prop) : ¬Xor' P Q ↔ (P ↔ Q) := by simp only [not_and, Xor', not_or, not_not, ← iff_iff_implies_and_implies] +#align not_xor not_xor theorem xor_iff_not_iff (P Q : Prop) : Xor' P Q ↔ ¬ (P ↔ Q) := (not_xor P Q).not_right theorem xor_iff_iff_not : Xor' a b ↔ (a ↔ ¬b) := by simp only [← @xor_not_right a, not_not] theorem xor_iff_not_iff' : Xor' a b ↔ (¬a ↔ b) := by simp only [← @xor_not_left _ b, not_not] +#align xor_iff_not_iff' xor_iff_not_iff' +#align xor_iff_iff_not xor_iff_iff_not +#align xor_iff_not_iff xor_iff_not_iff end Propositional @@ -413,48 +498,62 @@ section Equality theorem ball_cond_comm {α} {s : α → Prop} {p : α → α → Prop} : (∀ a, s a → ∀ b, s b → p a b) ↔ ∀ a b, s a → s b → p a b := ⟨fun h a b ha hb ↦ h a ha b hb, fun h a ha b hb ↦ h a b ha hb⟩ +#align ball_cond_comm ball_cond_comm theorem ball_mem_comm {α β} [Membership α β] {s : β} {p : α → α → Prop} : (∀ a (_ : a ∈ s) b (_ : b ∈ s), p a b) ↔ ∀ a b, a ∈ s → b ∈ s → p a b := ball_cond_comm +#align ball_mem_comm ball_mem_comm theorem ne_of_apply_ne {α β : Sort _} (f : α → β) {x y : α} (h : f x ≠ f y) : x ≠ y := fun w : x = y ↦ h (congr_arg f w) +#align ne_of_apply_ne ne_of_apply_ne theorem eq_equivalence : Equivalence (@Eq α) := ⟨Eq.refl, @Eq.symm _, @Eq.trans _⟩ +#align eq_equivalence eq_equivalence @[simp] theorem eq_mp_eq_cast (h : α = β) : Eq.mp h = cast h := rfl +#align eq_mp_eq_cast eq_mp_eq_cast @[simp] theorem eq_mpr_eq_cast (h : α = β) : Eq.mpr h = cast h.symm := rfl +#align eq_mpr_eq_cast eq_mpr_eq_cast @[simp] theorem cast_cast : ∀ (ha : α = β) (hb : β = γ) (a : α), cast hb (cast ha a) = cast (ha.trans hb) a | rfl, rfl, _ => rfl +#align cast_cast cast_cast -- @[simp] -- FIXME simp ignores proof rewrites theorem congr_refl_left (f : α → β) {a b : α} (h : a = b) : congr (Eq.refl f) h = congr_arg f h := rfl +#align congr_refl_left congr_refl_left -- @[simp] -- FIXME simp ignores proof rewrites theorem congr_refl_right {f g : α → β} (h : f = g) (a : α) : congr h (Eq.refl a) = congr_fun h a := rfl +#align congr_refl_right congr_refl_right -- @[simp] -- FIXME simp ignores proof rewrites theorem congr_arg_refl (f : α → β) (a : α) : congr_arg f (Eq.refl a) = Eq.refl (f a) := rfl +#align congr_arg_refl congr_arg_refl -- @[simp] -- FIXME simp ignores proof rewrites theorem congr_fun_rfl (f : α → β) (a : α) : congr_fun (Eq.refl f) a = Eq.refl (f a) := rfl +#align congr_fun_rfl congr_fun_rfl -- @[simp] -- FIXME simp ignores proof rewrites theorem congr_fun_congr_arg (f : α → β → γ) {a a' : α} (p : a = a') (b : β) : congr_fun (congr_arg f p) b = congr_arg (fun a ↦ f a b) p := rfl +#align congr_fun_congr_arg congr_fun_congr_arg theorem heq_of_cast_eq : ∀ (e : α = β) (_ : cast e a = a'), HEq a a' | rfl, h => Eq.recOn h (HEq.refl _) +#align heq_of_cast_eq heq_of_cast_eq theorem cast_eq_iff_heq : cast e a = a' ↔ HEq a a' := ⟨heq_of_cast_eq _, fun h ↦ by cases h; rfl⟩ +#align cast_eq_iff_heq cast_eq_iff_heq --Porting note: new theorem. More general version of `eqRec_heq` theorem eqRec_heq' {α : Sort u_1} {a' : α} {motive : (a : α) → a' = a → Sort u} @@ -464,13 +563,17 @@ theorem eqRec_heq' {α : Sort u_1} {a' : α} {motive : (a : α) → a' = a → S theorem rec_heq_of_heq {C : α → Sort _} {x : C a} {y : β} (e : a = b) (h : HEq x y) : HEq (@Eq.ndrec α a C x b e) y := by subst e; exact h +#align rec_heq_of_heq rec_heq_of_heq protected theorem Eq.congr (h₁ : x₁ = y₁) (h₂ : x₂ = y₂) : x₁ = x₂ ↔ y₁ = y₂ := by subst h₁; subst h₂; rfl +#align eq.congr Eq.congr theorem Eq.congr_left {x y z : α} (h : x = y) : x = z ↔ y = z := by rw [h] +#align eq.congr_left Eq.congr_left theorem Eq.congr_right {x y z : α} (h : x = y) : z = x ↔ z = y := by rw [h] +#align eq.congr_right Eq.congr_right alias congrArg₂ ← congr_arg₂ #align congr_arg2 congr_arg₂ @@ -479,16 +582,20 @@ variable {β : α → Sort _} {γ : ∀ a, β a → Sort _} {δ : ∀ a b, γ a theorem congr_fun₂ {f g : ∀ a b, γ a b} (h : f = g) (a : α) (b : β a) : f a b = g a b := congr_fun (congr_fun h _) _ +#align congr_fun₂ congr_fun₂ theorem congr_fun₃ {f g : ∀ a b c, δ a b c} (h : f = g) (a : α) (b : β a) (c : γ a b) : f a b c = g a b c := congr_fun₂ (congr_fun h _) _ _ +#align congr_fun₃ congr_fun₃ theorem funext₂ {f g : ∀ a b, γ a b} (h : ∀ a b, f a b = g a b) : f = g := funext fun _ ↦ funext <| h _ +#align funext₂ funext₂ theorem funext₃ {f g : ∀ a b c, δ a b c} (h : ∀ a b c, f a b c = g a b c) : f = g := funext fun _ ↦ funext₂ <| h _ +#align funext₃ funext₃ end Equality @@ -503,6 +610,7 @@ variable {β : α → Sort _} {γ : ∀ a, β a → Sort _} {δ : ∀ a b, γ a theorem pi_congr {β' : α → Sort _} (h : ∀ a, β a = β' a) : (∀ a, β a) = ∀ a, β' a := (funext h : β = β') ▸ rfl +#align pi_congr pi_congr -- Porting note: some higher order lemmas such as `forall₂_congr` and `exists₂_congr` -- were moved to `Std4` @@ -510,18 +618,22 @@ theorem pi_congr {β' : α → Sort _} (h : ∀ a, β a = β' a) : (∀ a, β a) theorem forall₂_imp {p q : ∀ a, β a → Prop} (h : ∀ a b, p a b → q a b) : (∀ a b, p a b) → ∀ a b, q a b := forall_imp fun i ↦ forall_imp <| h i +#align forall₂_imp forall₂_imp theorem forall₃_imp {p q : ∀ a b, γ a b → Prop} (h : ∀ a b c, p a b c → q a b c) : (∀ a b c, p a b c) → ∀ a b c, q a b c := forall_imp fun a ↦ forall₂_imp <| h a +#align forall₃_imp forall₃_imp theorem Exists₂.imp {p q : ∀ a, β a → Prop} (h : ∀ a b, p a b → q a b) : (∃ a b, p a b) → ∃ a b, q a b := Exists.imp fun a ↦ Exists.imp <| h a +#align Exists₂.imp Exists₂.imp theorem Exists₃.imp {p q : ∀ a b, γ a b → Prop} (h : ∀ a b c, p a b c → q a b c) : (∃ a b c, p a b c) → ∃ a b c, q a b c := Exists.imp fun a ↦ Exists₂.imp <| h a +#align Exists₃.imp Exists₃.imp end Dependent @@ -530,24 +642,30 @@ variable {κ : ι → Sort _} {p q : α → Prop} #align exists_imp_exists' Exists.imp' theorem forall_swap {p : α → β → Prop} : (∀ x y, p x y) ↔ ∀ y x, p x y := ⟨swap, swap⟩ +#align forall_swap forall_swap theorem forall₂_swap {κ₁ : ι₁ → Sort _} {κ₂ : ι₂ → Sort _} {p : ∀ i₁, κ₁ i₁ → ∀ i₂, κ₂ i₂ → Prop} : (∀ i₁ j₁ i₂ j₂, p i₁ j₁ i₂ j₂) ↔ ∀ i₂ j₂ i₁ j₁, p i₁ j₁ i₂ j₂ := ⟨swap₂, swap₂⟩ +#align forall₂_swap forall₂_swap /-- We intentionally restrict the type of `α` in this lemma so that this is a safer to use in simp than `forall_swap`. -/ theorem imp_forall_iff {α : Type _} {p : Prop} {q : α → Prop} : (p → ∀ x, q x) ↔ ∀ x, p → q x := forall_swap +#align imp_forall_iff imp_forall_iff theorem exists_swap {p : α → β → Prop} : (∃ x y, p x y) ↔ ∃ y x, p x y := ⟨fun ⟨x, y, h⟩ ↦ ⟨y, x, h⟩, fun ⟨y, x, h⟩ ↦ ⟨x, y, h⟩⟩ +#align exists_swap exists_swap @[simp] theorem forall_exists_index {q : (∃ x, p x) → Prop} : (∀ h, q h) ↔ ∀ x (h : p x), q ⟨x, h⟩ := ⟨fun h x hpx ↦ h ⟨x, hpx⟩, fun h ⟨x, hpx⟩ ↦ h x hpx⟩ +#align forall_exists_index forall_exists_index #align exists_imp_distrib exists_imp alias exists_imp ↔ _ not_exists_of_forall_not +#align not_exists_of_forall_not not_exists_of_forall_not #align Exists.some Exists.choose #align Exists.some_spec Exists.choose_spec @@ -557,34 +675,43 @@ protected theorem Decidable.not_forall {p : α → Prop} [Decidable (∃ x, ¬p [∀ x, Decidable (p x)] : (¬∀ x, p x) ↔ ∃ x, ¬p x := ⟨Not.decidable_imp_symm fun nx x ↦ nx.decidable_imp_symm fun h ↦ ⟨x, h⟩, not_forall_of_exists_not⟩ +#align decidable.not_forall Decidable.not_forall @[simp] theorem not_forall {p : α → Prop} : (¬∀ x, p x) ↔ ∃ x, ¬p x := Decidable.not_forall +#align not_forall not_forall -- See Note [decidable namespace] protected theorem Decidable.not_forall_not [Decidable (∃ x, p x)] : (¬∀ x, ¬p x) ↔ ∃ x, p x := (@Decidable.not_iff_comm _ _ _ (decidable_of_iff (¬∃ x, p x) not_exists)).1 not_exists +#align decidable.not_forall_not Decidable.not_forall_not theorem not_forall_not : (¬∀ x, ¬p x) ↔ ∃ x, p x := Decidable.not_forall_not +#align not_forall_not not_forall_not -- See Note [decidable namespace] protected theorem Decidable.not_exists_not [∀ x, Decidable (p x)] : (¬∃ x, ¬p x) ↔ ∀ x, p x := by simp only [not_exists, Decidable.not_not] +#align decidable.not_exists_not Decidable.not_exists_not theorem not_exists_not : (¬∃ x, ¬p x) ↔ ∀ x, p x := Decidable.not_exists_not +#align not_exists_not not_exists_not theorem forall_imp_iff_exists_imp [ha : Nonempty α] : (∀ x, p x) → b ↔ ∃ x, p x → b := by let ⟨a⟩ := ha refine ⟨fun h ↦ not_forall_not.1 fun h' ↦ ?_, fun ⟨x, hx⟩ h ↦ hx (h x)⟩ exact if hb : b then h' a fun _ ↦ hb else hb <| h fun x ↦ (not_imp.1 (h' x)).1 +#align forall_imp_iff_exists_imp forall_imp_iff_exists_imp theorem forall_true_iff : (α → True) ↔ True := imp_true_iff _ +#align forall_true_iff forall_true_iff -- Unfortunately this causes simp to loop sometimes, so we -- add the 2 and 3 cases as simp lemmas instead theorem forall_true_iff' (h : ∀ a, p a ↔ True) : (∀ a, p a) ↔ True := iff_true_intro fun _ ↦ of_iff_true (h _) +#align forall_true_iff' forall_true_iff' -- This is not marked `@[simp]` because `implies_true : (α → True) = True` works theorem forall₂_true_iff {β : α → Sort _} : (∀ a, β a → True) ↔ True := by simp @@ -598,14 +725,17 @@ theorem forall₃_true_iff {β : α → Sort _} {γ : ∀ a, β a → Sort _} : @[simp] theorem exists_unique_iff_exists [Subsingleton α] {p : α → Prop} : (∃! x, p x) ↔ ∃ x, p x := ⟨fun h ↦ h.exists, Exists.imp fun x hx ↦ ⟨hx, fun y _ ↦ Subsingleton.elim y x⟩⟩ +#align exists_unique_iff_exists exists_unique_iff_exists -- forall_forall_const is no longer needed @[simp] theorem exists_const (α) [i : Nonempty α] : (∃ _ : α, b) ↔ b := ⟨fun ⟨_, h⟩ ↦ h, i.elim Exists.intro⟩ +#align exists_const exists_const theorem exists_unique_const (α) [i : Nonempty α] [Subsingleton α] : (∃! _ : α, b) ↔ b := by simp +#align exists_unique_const exists_unique_const #align forall_and_distrib forall_and #align exists_or_distrib exists_or @@ -624,15 +754,19 @@ theorem and_forall_ne (a : α) : (p a ∧ ∀ b, b ≠ a → p b) ↔ ∀ b, p b theorem Ne.ne_or_ne {x y : α} (z : α) (h : x ≠ y) : x ≠ z ∨ y ≠ z := not_and_or.1 <| mt (and_imp.2 (· ▸ ·)) h.symm +#align ne.ne_or_ne Ne.ne_or_ne @[simp] theorem exists_unique_eq {a' : α} : ∃! a, a = a' := by simp only [eq_comm, ExistsUnique, and_self, forall_eq', exists_eq'] +#align exists_unique_eq exists_unique_eq @[simp] theorem exists_unique_eq' {a' : α} : ∃! a, a' = a := by simp only [ExistsUnique, and_self, forall_eq', exists_eq'] +#align exists_unique_eq' exists_unique_eq' -- @[simp] -- FIXME simp does not apply this lemma for some reason theorem exists_apply_eq_apply' (f : α → β) (a' : α) : ∃ a, f a' = f a := ⟨a', rfl⟩ +#align exists_apply_eq_apply' exists_apply_eq_apply' -- porting note: an alternative workaround theorem: theorem exists_apply_eq (a : α) (b : β) : ∃ f : α → β, f a = b := ⟨fun _ ↦ b, rfl⟩ @@ -640,48 +774,64 @@ theorem exists_apply_eq (a : α) (b : β) : ∃ f : α → β, f a = b := ⟨fun @[simp] theorem exists_exists_and_eq_and {f : α → β} {p : α → Prop} {q : β → Prop} : (∃ b, (∃ a, p a ∧ f a = b) ∧ q b) ↔ ∃ a, p a ∧ q (f a) := ⟨fun ⟨_, ⟨a, ha, hab⟩, hb⟩ ↦ ⟨a, ha, hab.symm ▸ hb⟩, fun ⟨a, hp, hq⟩ ↦ ⟨f a, ⟨a, hp, rfl⟩, hq⟩⟩ +#align exists_exists_and_eq_and exists_exists_and_eq_and @[simp] theorem exists_exists_eq_and {f : α → β} {p : β → Prop} : (∃ b, (∃ a, f a = b) ∧ p b) ↔ ∃ a, p (f a) := ⟨fun ⟨_, ⟨a, ha⟩, hb⟩ ↦ ⟨a, ha.symm ▸ hb⟩, fun ⟨a, ha⟩ ↦ ⟨f a, ⟨a, rfl⟩, ha⟩⟩ +#align exists_exists_eq_and exists_exists_eq_and @[simp] theorem exists_or_eq_left (y : α) (p : α → Prop) : ∃ x : α, x = y ∨ p x := ⟨y, .inl rfl⟩ +#align exists_or_eq_left exists_or_eq_left @[simp] theorem exists_or_eq_right (y : α) (p : α → Prop) : ∃ x : α, p x ∨ x = y := ⟨y, .inr rfl⟩ +#align exists_or_eq_right exists_or_eq_right @[simp] theorem exists_or_eq_left' (y : α) (p : α → Prop) : ∃ x : α, y = x ∨ p x := ⟨y, .inl rfl⟩ +#align exists_or_eq_left' exists_or_eq_left' @[simp] theorem exists_or_eq_right' (y : α) (p : α → Prop) : ∃ x : α, p x ∨ y = x := ⟨y, .inr rfl⟩ +#align exists_or_eq_right' exists_or_eq_right' theorem forall_apply_eq_imp_iff {f : α → β} {p : β → Prop} : (∀ a b, f a = b → p b) ↔ ∀ a, p (f a) := by simp +#align forall_apply_eq_imp_iff forall_apply_eq_imp_iff @[simp] theorem forall_apply_eq_imp_iff' {f : α → β} {p : β → Prop} : (∀ b a, f a = b → p b) ↔ ∀ a, p (f a) := by simp [forall_swap] +#align forall_apply_eq_imp_iff' forall_apply_eq_imp_iff' theorem forall_eq_apply_imp_iff {f : α → β} {p : β → Prop} : (∀ a b, b = f a → p b) ↔ ∀ a, p (f a) := by simp +#align forall_eq_apply_imp_iff forall_eq_apply_imp_iff @[simp] theorem forall_eq_apply_imp_iff' {f : α → β} {p : β → Prop} : (∀ b a, b = f a → p b) ↔ ∀ a, p (f a) := by simp [forall_swap] +#align forall_eq_apply_imp_iff' forall_eq_apply_imp_iff' @[simp] theorem forall_apply_eq_imp_iff₂ {f : α → β} {p : α → Prop} {q : β → Prop} : (∀ b a, p a → f a = b → q b) ↔ ∀ a, p a → q (f a) := ⟨fun h a ha ↦ h (f a) a ha rfl, fun h _ a ha hb ↦ hb ▸ h a ha⟩ +#align forall_apply_eq_imp_iff₂ forall_apply_eq_imp_iff₂ @[simp] theorem exists_eq_right' {a' : α} : (∃ a, p a ∧ a' = a) ↔ p a' := by simp [@eq_comm _ a'] +#align exists_eq_right' exists_eq_right' theorem exists_comm {p : α → β → Prop} : (∃ a b, p a b) ↔ ∃ b a, p a b := ⟨fun ⟨a, b, h⟩ ↦ ⟨b, a, h⟩, fun ⟨b, a, h⟩ ↦ ⟨a, b, h⟩⟩ +#align exists_comm exists_comm theorem exists₂_comm {κ₁ : ι₁ → Sort _} {κ₂ : ι₂ → Sort _} {p : ∀ i₁, κ₁ i₁ → ∀ i₂, κ₂ i₂ → Prop} : (∃ i₁ j₁ i₂ j₂, p i₁ j₁ i₂ j₂) ↔ ∃ i₂ j₂ i₁ j₁, p i₁ j₁ i₂ j₂ := by simp only [@exists_comm (κ₁ _), @exists_comm ι₁] +#align exists₂_comm exists₂_comm theorem And.exists {p q : Prop} {f : p ∧ q → Prop} : (∃ h, f h) ↔ ∃ hp hq, f ⟨hp, hq⟩ := ⟨fun ⟨h, H⟩ ↦ ⟨h.1, h.2, H⟩, fun ⟨hp, hq, H⟩ ↦ ⟨⟨hp, hq⟩, H⟩⟩ +#align and.exists And.exists theorem forall_or_of_or_forall (h : b ∨ ∀ x, p x) (x) : b ∨ p x := h.imp_right fun h₂ ↦ h₂ x +#align forall_or_of_or_forall forall_or_of_or_forall -- See Note [decidable namespace] protected theorem Decidable.forall_or_left {q : Prop} {p : α → Prop} [Decidable q] : @@ -704,59 +854,74 @@ theorem forall_or_right {q} {p : α → Prop} : (∀ x, p x ∨ q) ↔ (∀ x, p #align forall_or_distrib_right forall_or_right theorem exists_unique_prop {p q : Prop} : (∃! _ : p, q) ↔ p ∧ q := by simp +#align exists_unique_prop exists_unique_prop @[simp] theorem exists_unique_false : ¬∃! _ : α, False := fun ⟨_, h, _⟩ ↦ h +#align exists_unique_false exists_unique_false theorem Exists.fst {b : Prop} {p : b → Prop} : Exists p → b | ⟨h, _⟩ => h +#align Exists.fst Exists.fst theorem Exists.snd {b : Prop} {p : b → Prop} : ∀ h : Exists p, p h.fst | ⟨_, h⟩ => h +#align Exists.snd Exists.snd theorem exists_prop_of_true {p : Prop} {q : p → Prop} (h : p) : (∃ h' : p, q h') ↔ q h := @exists_const (q h) p ⟨h⟩ +#align exists_prop_of_true exists_prop_of_true theorem exists_iff_of_forall {p : Prop} {q : p → Prop} (h : ∀ h, q h) : (∃ h, q h) ↔ p := ⟨Exists.fst, fun H ↦ ⟨H, h H⟩⟩ +#align exists_iff_of_forall exists_iff_of_forall theorem exists_unique_prop_of_true {p : Prop} {q : p → Prop} (h : p) : (∃! h' : p, q h') ↔ q h := @exists_unique_const (q h) p ⟨h⟩ _ +#align exists_unique_prop_of_true exists_unique_prop_of_true theorem forall_prop_of_false {p : Prop} {q : p → Prop} (hn : ¬p) : (∀ h' : p, q h') ↔ True := iff_true_intro fun h ↦ hn.elim h +#align forall_prop_of_false forall_prop_of_false theorem exists_prop_of_false {p : Prop} {q : p → Prop} : ¬p → ¬∃ h' : p, q h' := mt Exists.fst +#align exists_prop_of_false exists_prop_of_false @[congr] theorem exists_prop_congr {p p' : Prop} {q q' : p → Prop} (hq : ∀ h, q h ↔ q' h) (hp : p ↔ p') : Exists q ↔ ∃ h : p', q' (hp.2 h) := ⟨fun ⟨_, _⟩ ↦ ⟨hp.1 ‹_›, (hq _).1 ‹_›⟩, fun ⟨_, _⟩ ↦ ⟨_, (hq _).2 ‹_›⟩⟩ +#align exists_prop_congr exists_prop_congr @[congr] theorem exists_prop_congr' {p p' : Prop} {q q' : p → Prop} (hq : ∀ h, q h ↔ q' h) (hp : p ↔ p') : Exists q = ∃ h : p', q' (hp.2 h) := propext (exists_prop_congr hq hp) +#align exists_prop_congr' exists_prop_congr' /-- See `IsEmpty.exists_iff` for the `false` version. -/ @[simp] theorem exists_true_left (p : True → Prop) : (∃ x, p x) ↔ p True.intro := exists_prop_of_true _ +#align exists_true_left exists_true_left -- Porting note: `@[congr]` commented out for now. -- @[congr] theorem forall_prop_congr {p p' : Prop} {q q' : p → Prop} (hq : ∀ h, q h ↔ q' h) (hp : p ↔ p') : (∀ h, q h) ↔ ∀ h : p', q' (hp.2 h) := ⟨fun h1 h2 ↦ (hq _).1 (h1 (hp.2 h2)), fun h1 h2 ↦ (hq _).2 (h1 (hp.1 h2))⟩ +#align forall_prop_congr forall_prop_congr -- Porting note: `@[congr]` commented out for now. -- @[congr] theorem forall_prop_congr' {p p' : Prop} {q q' : p → Prop} (hq : ∀ h, q h ↔ q' h) (hp : p ↔ p') : (∀ h, q h) = ∀ h : p', q' (hp.2 h) := propext (forall_prop_congr hq hp) +#align forall_prop_congr' forall_prop_congr' /-- See `IsEmpty.forall_iff` for the `false` version. -/ @[simp] theorem forall_true_left (p : True → Prop) : (∀ x, p x) ↔ p True.intro := forall_prop_of_true _ +#align forall_true_left forall_true_left theorem ExistsUnique.elim₂ {α : Sort _} {p : α → Sort _} [∀ x, Subsingleton (p x)] {q : ∀ (x) (_ : p x), Prop} {b : Prop} (h₂ : ∃! (x : _) (h : p x), q x h) @@ -795,21 +960,26 @@ variable {p : α → Prop} -- use shortened names to avoid conflict when classical namespace is open. /-- Any prop `p` is decidable classically. A shorthand for `classical.prop_decidable`. -/ noncomputable def dec (p : Prop) : Decidable p := by infer_instance +#align classical.dec Classical.dec /-- Any predicate `p` is decidable classically. -/ noncomputable def decPred (p : α → Prop) : DecidablePred p := by infer_instance +#align classical.dec_pred Classical.decPred /-- Any relation `p` is decidable classically. -/ noncomputable def decRel (p : α → α → Prop) : DecidableRel p := by infer_instance +#align classical.dec_rel Classical.decRel /-- Any type `α` has decidable equality classically. -/ noncomputable def decEq (α : Sort u) : DecidableEq α := by infer_instance +#align classical.dec_eq Classical.decEq /-- Construct a function from a default value `H0`, and a function to use if there exists a value satisfying the predicate. -/ -- @[elab_as_elim] -- FIXME noncomputable def existsCases (H0 : C) (H : ∀ a, p a → C) : C := if h : ∃ a, p a then H (Classical.choose h) (Classical.choose_spec h) else H0 +#align classical.exists_cases Classical.existsCases theorem some_spec₂ {α : Sort _} {p : α → Prop} {h : ∃ a, p a} (q : α → Prop) (hpq : ∀ a, p a → q a) : q (choose h) := hpq _ <| choose_spec _ @@ -823,6 +993,7 @@ noncomputable def subtype_of_exists {α : Type _} {P : α → Prop} (h : ∃ x, /-- A version of `byContradiction` that uses types instead of propositions. -/ protected noncomputable def byContradiction' {α : Sort _} (H : ¬(α → False)) : α := Classical.choice <| (peirce _ False) fun h ↦ (H fun a ↦ h ⟨a⟩).elim +#align classical.by_contradiction' Classical.byContradiction' /-- `classical.byContradiction'` is equivalent to lean's axiom `classical.choice`. -/ def choice_of_byContradiction' {α : Sort _} (contra : ¬(α → False) → α) : Nonempty α → α := @@ -837,6 +1008,7 @@ using the axiom of choice. -/ -- @[elab_as_elim] -- FIXME noncomputable def Exists.classicalRecOn {p : α → Prop} (h : ∃ a, p a) {C} (H : ∀ a, p a → C) : C := H (Classical.choose h) (Classical.choose_spec h) +#align exists.classical_rec_on Exists.classicalRecOn /-! ### Declarations about bounded quantifiers -/ @@ -845,6 +1017,7 @@ variable {r p q : α → Prop} {P Q : ∀ x, p x → Prop} {b : Prop} theorem bex_def : (∃ (x : _) (_ : p x), q x) ↔ ∃ x, p x ∧ q x := ⟨fun ⟨x, px, qx⟩ ↦ ⟨x, px, qx⟩, fun ⟨x, px, qx⟩ ↦ ⟨x, px, qx⟩⟩ +#align bex_def bex_def theorem BEx.elim {b : Prop} : (∃ x h, P x h) → (∀ a h, P a h → b) → b | ⟨a, h₁, h₂⟩, h' => h' a h₁ h₂ @@ -856,12 +1029,15 @@ theorem BEx.intro (a : α) (h₁ : p a) (h₂ : P a h₁) : ∃ (x : _) (h : p x theorem ball_congr (H : ∀ x h, P x h ↔ Q x h) : (∀ x h, P x h) ↔ ∀ x h, Q x h := forall_congr' fun x ↦ forall_congr' (H x) +#align ball_congr ball_congr theorem bex_congr (H : ∀ x h, P x h ↔ Q x h) : (∃ x h, P x h) ↔ ∃ x h, Q x h := exists_congr fun x ↦ exists_congr (H x) +#align bex_congr bex_congr theorem bex_eq_left {a : α} : (∃ (x : _) (_ : x = a), p x) ↔ p a := by simp only [exists_prop, exists_eq_left] +#align bex_eq_left bex_eq_left theorem BAll.imp_right (H : ∀ x h, P x h → Q x h) (h₁ : ∀ x h, P x h) (x h) : Q x h := H _ _ <| h₁ _ _ @@ -880,33 +1056,42 @@ theorem BEx.imp_left (H : ∀ x, p x → q x) : (∃ (x : _) (_ : p x), r x) → #align bex.imp_left BEx.imp_left theorem ball_of_forall (h : ∀ x, p x) (x) : p x := h x +#align ball_of_forall ball_of_forall theorem forall_of_ball (H : ∀ x, p x) (h : ∀ x, p x → q x) (x) : q x := h x <| H x +#align forall_of_ball forall_of_ball theorem bex_of_exists (H : ∀ x, p x) : (∃ x, q x) → ∃ (x : _) (_ : p x), q x | ⟨x, hq⟩ => ⟨x, H x, hq⟩ +#align bex_of_exists bex_of_exists theorem exists_of_bex : (∃ (x : _) (_ : p x), q x) → ∃ x, q x | ⟨x, _, hq⟩ => ⟨x, hq⟩ +#align exists_of_bex exists_of_bex theorem bex_imp : (∃ x h, P x h) → b ↔ ∀ x h, P x h → b := by simp #align bex_imp_distrib bex_imp theorem not_bex : (¬∃ x h, P x h) ↔ ∀ x h, ¬P x h := bex_imp +#align not_bex not_bex theorem not_ball_of_bex_not : (∃ x h, ¬P x h) → ¬∀ x h, P x h | ⟨x, h, hp⟩, al => hp <| al x h +#align not_ball_of_bex_not not_ball_of_bex_not -- See Note [decidable namespace] protected theorem Decidable.not_ball [Decidable (∃ x h, ¬P x h)] [∀ x h, Decidable (P x h)] : (¬∀ x h, P x h) ↔ ∃ x h, ¬P x h := ⟨Not.decidable_imp_symm fun nx x h ↦ nx.decidable_imp_symm fun h' ↦ ⟨x, h, h'⟩, not_ball_of_bex_not⟩ +#align decidable.not_ball Decidable.not_ball theorem not_ball : (¬∀ x h, P x h) ↔ ∃ x h, ¬P x h := Decidable.not_ball +#align not_ball not_ball theorem ball_true_iff (p : α → Prop) : (∀ x, p x → True) ↔ True := iff_true_intro fun _ _ ↦ trivial +#align ball_true_iff ball_true_iff theorem ball_and : (∀ x h, P x h ∧ Q x h) ↔ (∀ x h, P x h) ∧ ∀ x h, Q x h := Iff.trans (forall_congr' fun _ ↦ forall_and) forall_and @@ -937,9 +1122,11 @@ variable {σ : α → Sort _} (f : α → β) {P Q : Prop} [Decidable P] [Decida theorem dite_eq_iff : dite P A B = c ↔ (∃ h, A h = c) ∨ ∃ h, B h = c := by by_cases P <;> simp [*, exists_prop_of_true, exists_prop_of_false] +#align dite_eq_iff dite_eq_iff theorem ite_eq_iff : ite P a b = c ↔ P ∧ a = c ∨ ¬P ∧ b = c := dite_eq_iff.trans <| by simp only; rw [exists_prop, exists_prop] +#align ite_eq_iff ite_eq_iff theorem eq_ite_iff : a = ite P b c ↔ P ∧ a = b ∨ ¬P ∧ a = c := eq_comm.trans <| ite_eq_iff.trans <| (Iff.rfl.and eq_comm).or (Iff.rfl.and eq_comm) @@ -947,65 +1134,86 @@ eq_comm.trans <| ite_eq_iff.trans <| (Iff.rfl.and eq_comm).or (Iff.rfl.and eq_co theorem dite_eq_iff' : dite P A B = c ↔ (∀ h, A h = c) ∧ ∀ h, B h = c := ⟨fun he ↦ ⟨fun h ↦ (dif_pos h).symm.trans he, fun h ↦ (dif_neg h).symm.trans he⟩, fun he ↦ (em P).elim (fun h ↦ (dif_pos h).trans <| he.1 h) fun h ↦ (dif_neg h).trans <| he.2 h⟩ +#align dite_eq_iff' dite_eq_iff' theorem ite_eq_iff' : ite P a b = c ↔ (P → a = c) ∧ (¬P → b = c) := dite_eq_iff' +#align ite_eq_iff' ite_eq_iff' @[simp] theorem dite_eq_left_iff : dite P (fun _ ↦ a) B = a ↔ ∀ h, B h = a := by by_cases P <;> simp [*, forall_prop_of_true, forall_prop_of_false] +#align dite_eq_left_iff dite_eq_left_iff @[simp] theorem dite_eq_right_iff : (dite P A fun _ ↦ b) = b ↔ ∀ h, A h = b := by by_cases P <;> simp [*, forall_prop_of_true, forall_prop_of_false] +#align dite_eq_right_iff dite_eq_right_iff @[simp] theorem ite_eq_left_iff : ite P a b = a ↔ ¬P → b = a := dite_eq_left_iff @[simp] theorem ite_eq_right_iff : ite P a b = b ↔ P → a = b := dite_eq_right_iff +#align ite_eq_right_iff ite_eq_right_iff +#align ite_eq_left_iff ite_eq_left_iff theorem dite_ne_left_iff : dite P (fun _ ↦ a) B ≠ a ↔ ∃ h, a ≠ B h := by rw [Ne.def, dite_eq_left_iff, not_forall] exact exists_congr fun h ↦ by rw [ne_comm] +#align dite_ne_left_iff dite_ne_left_iff theorem dite_ne_right_iff : (dite P A fun _ ↦ b) ≠ b ↔ ∃ h, A h ≠ b := by simp only [Ne.def, dite_eq_right_iff, not_forall] +#align dite_ne_right_iff dite_ne_right_iff theorem ite_ne_left_iff : ite P a b ≠ a ↔ ¬P ∧ a ≠ b := dite_ne_left_iff.trans <| by simp only; rw [exists_prop] +#align ite_ne_left_iff ite_ne_left_iff theorem ite_ne_right_iff : ite P a b ≠ b ↔ P ∧ a ≠ b := dite_ne_right_iff.trans <| by simp only; rw [exists_prop] +#align ite_ne_right_iff ite_ne_right_iff protected theorem Ne.dite_eq_left_iff (h : ∀ h, a ≠ B h) : dite P (fun _ ↦ a) B = a ↔ P := dite_eq_left_iff.trans ⟨fun H ↦ of_not_not fun h' ↦ h h' (H h').symm, fun h H ↦ (H h).elim⟩ +#align ne.dite_eq_left_iff Ne.dite_eq_left_iff protected theorem Ne.dite_eq_right_iff (h : ∀ h, A h ≠ b) : (dite P A fun _ ↦ b) = b ↔ ¬P := dite_eq_right_iff.trans ⟨fun H h' ↦ h h' (H h'), fun h' H ↦ (h' H).elim⟩ +#align ne.dite_eq_right_iff Ne.dite_eq_right_iff protected theorem Ne.ite_eq_left_iff (h : a ≠ b) : ite P a b = a ↔ P := Ne.dite_eq_left_iff fun _ ↦ h +#align ne.ite_eq_left_iff Ne.ite_eq_left_iff protected theorem Ne.ite_eq_right_iff (h : a ≠ b) : ite P a b = b ↔ ¬P := Ne.dite_eq_right_iff fun _ ↦ h +#align ne.ite_eq_right_iff Ne.ite_eq_right_iff protected theorem Ne.dite_ne_left_iff (h : ∀ h, a ≠ B h) : dite P (fun _ ↦ a) B ≠ a ↔ ¬P := dite_ne_left_iff.trans <| exists_iff_of_forall h +#align ne.dite_ne_left_iff Ne.dite_ne_left_iff protected theorem Ne.dite_ne_right_iff (h : ∀ h, A h ≠ b) : (dite P A fun _ ↦ b) ≠ b ↔ P := dite_ne_right_iff.trans <| exists_iff_of_forall h +#align ne.dite_ne_right_iff Ne.dite_ne_right_iff protected theorem Ne.ite_ne_left_iff (h : a ≠ b) : ite P a b ≠ a ↔ ¬P := Ne.dite_ne_left_iff fun _ ↦ h +#align ne.ite_ne_left_iff Ne.ite_ne_left_iff protected theorem Ne.ite_ne_right_iff (h : a ≠ b) : ite P a b ≠ b ↔ P := Ne.dite_ne_right_iff fun _ ↦ h +#align ne.ite_ne_right_iff Ne.ite_ne_right_iff variable (P Q a b) /-- A `dite` whose results do not actually depend on the condition may be reduced to an `ite`. -/ @[simp] theorem dite_eq_ite : (dite P (fun _ ↦ a) fun _ ↦ b) = ite P a b := rfl +#align dite_eq_ite dite_eq_ite theorem dite_eq_or_eq : (∃ h, dite P A B = A h) ∨ ∃ h, dite P A B = B h := if h : _ then .inl ⟨h, dif_pos h⟩ else .inr ⟨h, dif_neg h⟩ +#align dite_eq_or_eq dite_eq_or_eq theorem ite_eq_or_eq : ite P a b = a ∨ ite P a b = b := if h : _ then .inl (if_pos h) else .inr (if_neg h) +#align ite_eq_or_eq ite_eq_or_eq /-- A two-argument function applied to two `dite`s is a `dite` of that two-argument function applied to each of the branches. -/ @@ -1026,14 +1234,17 @@ theorem apply_ite₂ (f : α → β → γ) (P : Prop) [Decidable P] (a b : α) either branch to `a`. -/ theorem dite_apply (f : P → ∀ a, σ a) (g : ¬P → ∀ a, σ a) (a : α) : (dite P f g) a = dite P (fun h ↦ f h a) fun h ↦ g h a := by by_cases h:P <;> simp [h] +#align dite_apply dite_apply /-- A 'ite' producing a `Pi` type `Π a, σ a`, applied to a value `a : α` is a `ite` that applies either branch to `a`. -/ theorem ite_apply (f g : ∀ a, σ a) (a : α) : (ite P f g) a = ite P (f a) (g a) := dite_apply P (fun _ ↦ f) (fun _ ↦ g) a +#align ite_apply ite_apply theorem ite_and : ite (P ∧ Q) a b = ite P (ite Q a b) b := by by_cases hp : P <;> by_cases hq : Q <;> simp [hp, hq] +#align ite_and ite_and theorem dite_dite_comm {B : Q → α} {C : ¬P → ¬Q → α} (h : P → ¬Q) : (if p : P then A p else if q : Q then B q else C p q) = @@ -1041,10 +1252,12 @@ theorem dite_dite_comm {B : Q → α} {C : ¬P → ¬Q → α} (h : P → ¬Q) : dite_eq_iff'.2 ⟨ fun p ↦ by rw [dif_neg (h p), dif_pos p], fun np ↦ by congr; funext _; rw [dif_neg np]⟩ +#align dite_dite_comm dite_dite_comm theorem ite_ite_comm (h : P → ¬Q) : (if P then a else if Q then b else c) = if Q then b else if P then a else c := dite_dite_comm P Q h +#align ite_ite_comm ite_ite_comm end ite diff --git a/Mathlib/Logic/Embedding/Basic.lean b/Mathlib/Logic/Embedding/Basic.lean index 1852f68e3869c..4cd2b0746a524 100644 --- a/Mathlib/Logic/Embedding/Basic.lean +++ b/Mathlib/Logic/Embedding/Basic.lean @@ -138,12 +138,14 @@ theorem apply_eq_iff_eq {α β} (f : α ↪ β) (x y : α) : f x = f y ↔ x = y protected def refl (α : Sort _) : α ↪ α := ⟨id, injective_id⟩ #align function.embedding.refl Function.Embedding.refl +#align function.embedding.refl_apply Function.Embedding.refl_apply /-- Composition of `f : α ↪ β` and `g : β ↪ γ`. -/ @[trans, simps (config := { simpRhs := true })] protected def trans {α β γ} (f : α ↪ β) (g : β ↪ γ) : α ↪ γ := ⟨g ∘ f, g.injective.comp f.injective⟩ #align function.embedding.trans Function.Embedding.trans +#align function.embedding.trans_apply Function.Embedding.trans_apply instance : Trans Embedding Embedding Embedding := ⟨Embedding.trans⟩ @@ -167,6 +169,7 @@ protected def congr {α : Sort u} {β : Sort v} {γ : Sort w} {δ : Sort x} (e (f : α ↪ γ) : β ↪ δ := (Equiv.toEmbedding e₁.symm).trans (f.trans e₂.toEmbedding) #align function.embedding.congr Function.Embedding.congr +#align function.embedding.congr_apply Function.Embedding.congr_apply /-- A right inverse `surjInv` of a surjective function as an `Embedding`. -/ protected noncomputable def ofSurjective {α β} (f : β → α) (hf : Surjective f) : α ↪ β := @@ -213,6 +216,7 @@ theorem setValue_eq {α β} (f : α ↪ β) (a : α) (b : β) [∀ a', Decidable protected def some {α} : α ↪ Option α := ⟨some, Option.some_injective α⟩ #align function.embedding.some Function.Embedding.some +#align function.embedding.some_apply Function.Embedding.some_apply -- porting note: Lean 4 unfolds coercion `α → Option α` to `some`, so there is no separate -- `Function.Embedding.coeOption`. @@ -223,6 +227,7 @@ protected def some {α} : α ↪ Option α := def optionMap {α β} (f : α ↪ β) : Option α ↪ Option β := ⟨Option.map f, Option.map_injective f.injective⟩ #align function.embedding.option_map Function.Embedding.optionMap +#align function.embedding.option_map_apply Function.Embedding.optionMap_apply /-- Embedding of a `Subtype`. -/ def subtype {α} (p : α → Prop) : Subtype p ↪ α := @@ -256,12 +261,14 @@ def punit {β : Sort _} (b : β) : PUnit ↪ β := def sectl (α : Sort _) {β : Sort _} (b : β) : α ↪ α × β := ⟨fun a => (a, b), fun _ _ h => congr_arg Prod.fst h⟩ #align function.embedding.sectl Function.Embedding.sectl +#align function.embedding.sectl_apply Function.Embedding.sectl_apply /-- Fixing an element `a : α` gives an embedding `β ↪ α × β`. -/ @[simps] def sectr {α : Sort _} (a : α) (β : Sort _) : β ↪ α × β := ⟨fun b => (a, b), fun _ _ h => congr_arg Prod.snd h⟩ #align function.embedding.sectr Function.Embedding.sectr +#align function.embedding.sectr_apply Function.Embedding.sectr_apply /-- If `e₁` and `e₂` are embeddings, then so is `prod.map e₁ e₂ : (a, b) ↦ (e₁ a, e₂ b)`. -/ def prodMap {α β γ δ : Type _} (e₁ : α ↪ β) (e₂ : γ ↪ δ) : α × γ ↪ β × δ := @@ -298,12 +305,14 @@ theorem coe_sumMap {α β γ δ} (e₁ : α ↪ β) (e₂ : γ ↪ δ) : sumMap def inl {α β : Type _} : α ↪ Sum α β := ⟨Sum.inl, fun _ _ => Sum.inl.inj⟩ #align function.embedding.inl Function.Embedding.inl +#align function.embedding.inl_apply Function.Embedding.inl_apply /-- The embedding of `β` into the sum `α ⊕ β`. -/ @[simps] def inr {α β : Type _} : β ↪ Sum α β := ⟨Sum.inr, fun _ _ => Sum.inr.inj⟩ #align function.embedding.inr Function.Embedding.inr +#align function.embedding.inr_apply Function.Embedding.inr_apply end Sum @@ -324,6 +333,7 @@ of embeddings, then `Sigma.map f g` is an embedding. -/ def sigmaMap (f : α ↪ α') (g : ∀ a, β a ↪ β' (f a)) : (Σa, β a) ↪ Σa', β' a' := ⟨Sigma.map f fun a => g a, f.injective.sigma_map fun a => (g a).injective⟩ #align function.embedding.sigma_map Function.Embedding.sigmaMap +#align function.embedding.sigma_map_apply Function.Embedding.sigmaMap_apply end Sigma @@ -333,6 +343,7 @@ end Sigma def piCongrRight {α : Sort _} {β γ : α → Sort _} (e : ∀ a, β a ↪ γ a) : (∀ a, β a) ↪ ∀ a, γ a := ⟨fun f a => e a (f a), fun _ _ h => funext fun a => (e a).injective (congr_fun h a)⟩ #align function.embedding.Pi_congr_right Function.Embedding.piCongrRight +#align function.embedding.Pi_congr_right_apply Function.Embedding.piCongrRight_apply /-- An embedding `e : α ↪ β` defines an embedding `(γ → α) ↪ (γ → β)` that sends each `f` to `e ∘ f`. -/ @@ -388,6 +399,7 @@ set. -/ def asEmbedding {p : β → Prop} (e : α ≃ Subtype p) : α ↪ β := e.toEmbedding.trans (subtype p) #align equiv.as_embedding Equiv.asEmbedding +#align equiv.as_embedding_apply Equiv.asEmbedding_apply /-- The type of embeddings `α ↪ β` is equivalent to the subtype of all injective functions `α → β`. -/ @@ -413,6 +425,7 @@ def embeddingCongr {α β γ δ : Sort _} (h : α ≃ β) (h' : γ ≃ δ) : (α ext simp #align equiv.embedding_congr Equiv.embeddingCongr +#align equiv.embedding_congr_apply Equiv.embeddingCongr_apply @[simp] theorem embeddingCongr_refl {α β : Sort _} : @@ -487,5 +500,6 @@ if `p x → q x` for all `x : α`. -/ def Subtype.impEmbedding (p q : α → Prop) (h : ∀ x, p x → q x) : { x // p x } ↪ { x // q x } := ⟨fun x => ⟨x, h x x.prop⟩, fun x y => by simp [Subtype.ext_iff]⟩ #align subtype.imp_embedding Subtype.impEmbedding +#align subtype.imp_embedding_apply_coe Subtype.impEmbedding_apply_coe end Subtype diff --git a/Mathlib/Logic/Embedding/Set.lean b/Mathlib/Logic/Embedding/Set.lean index c3cb4d270a36b..d447f91b0cdca 100644 --- a/Mathlib/Logic/Embedding/Set.lean +++ b/Mathlib/Logic/Embedding/Set.lean @@ -40,6 +40,7 @@ namespace Embedding def coeWithTop {α} : α ↪ WithTop α := { Embedding.some with toFun := WithTop.some } #align function.embedding.coe_with_top Function.Embedding.coeWithTop +#align function.embedding.coe_with_top_apply Function.Embedding.coeWithTop_apply /-- Given an embedding `f : α ↪ β` and a point outside of `Set.range f`, construct an embedding `Option α ↪ β`. -/ @@ -47,6 +48,7 @@ def coeWithTop {α} : α ↪ WithTop α := def optionElim {α β} (f : α ↪ β) (x : β) (h : x ∉ Set.range f) : Option α ↪ β := ⟨Option.elim' x f, Option.injective_iff.2 ⟨f.2, h⟩⟩ #align function.embedding.option_elim Function.Embedding.optionElim +#align function.embedding.option_elim_apply Function.Embedding.optionElim_apply /-- Equivalence between embeddings of `Option α` and a sigma type over the embeddings of `α`. -/ @[simps] @@ -56,6 +58,9 @@ def optionEmbeddingEquiv (α β) : (Option α ↪ β) ≃ Σ f : α ↪ β, ↥( left_inv f := ext <| by rintro (_ | _) <;> simp [Option.coe_def] ; rfl right_inv := fun ⟨f, y, hy⟩ ↦ by ext <;> simp [Option.coe_def] ; rfl #align function.embedding.option_embedding_equiv Function.Embedding.optionEmbeddingEquiv +#align function.embedding.option_embedding_equiv_apply_snd_coe Function.Embedding.optionEmbeddingEquiv_apply_snd_coe +#align function.embedding.option_embedding_equiv_symm_apply Function.Embedding.optionEmbeddingEquiv_symm_apply +#align function.embedding.option_embedding_equiv_apply_fst Function.Embedding.optionEmbeddingEquiv_apply_fst /-- Restrict the codomain of an embedding. -/ def codRestrict {α β} (p : Set β) (f : α ↪ β) (H : ∀ a, f a ∈ p) : α ↪ p := @@ -74,6 +79,7 @@ open Set protected def image {α β} (f : α ↪ β) : Set α ↪ Set β := ⟨image f, f.2.image_injective⟩ #align function.embedding.image Function.Embedding.image +#align function.embedding.image_apply Function.Embedding.image_apply end Embedding @@ -88,6 +94,7 @@ def embeddingOfSubset {α} (s t : Set α) (h : s ⊆ t) : s ↪ t := congr injection h⟩ #align set.embedding_of_subset Set.embeddingOfSubset +#align set.embedding_of_subset_apply Set.embeddingOfSubset_apply end Set @@ -128,6 +135,7 @@ def subtypeOrEquiv (p q : α → Prop) [DecidablePred p] (h : Disjoint p q) : intro hp simpa using h.le_bot x ⟨hp, x.prop⟩ #align subtype_or_equiv subtypeOrEquiv +#align subtype_or_equiv_apply subtypeOrEquiv_apply @[simp] theorem subtypeOrEquiv_symm_inl (p q : α → Prop) [DecidablePred p] (h : Disjoint p q) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 097fad42e60ab..b429640a5ad68 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -60,6 +60,8 @@ def pprodEquivProd : PProd α β ≃ α × β where left_inv := fun _ => rfl right_inv := fun _ => rfl #align equiv.pprod_equiv_prod Equiv.pprodEquivProd +#align equiv.pprod_equiv_prod_apply Equiv.pprodEquivProd_apply +#align equiv.pprod_equiv_prod_symm_apply Equiv.pprodEquivProd_symm_apply /-- Product of two equivalences, in terms of `PProd`. If `α ≃ β` and `γ ≃ δ`, then `PProd α γ ≃ PProd β δ`. -/ @@ -71,6 +73,7 @@ def pprodCongr (e₁ : α ≃ β) (e₂ : γ ≃ δ) : PProd α γ ≃ PProd β left_inv := fun ⟨x, y⟩ => by simp right_inv := fun ⟨x, y⟩ => by simp #align equiv.pprod_congr Equiv.pprodCongr +#align equiv.pprod_congr_apply Equiv.pprodCongr_apply /-- Combine two equivalences using `PProd` in the domain and `Prod` in the codomain. -/ @[simps apply symm_apply] @@ -78,6 +81,8 @@ def pprodProd (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) : PProd α₁ β₁ ≃ α₂ × β₂ := (ea.pprodCongr eb).trans pprodEquivProd #align equiv.pprod_prod Equiv.pprodProd +#align equiv.pprod_prod_apply Equiv.pprodProd_apply +#align equiv.pprod_prod_symm_apply Equiv.pprodProd_symm_apply /-- Combine two equivalences using `PProd` in the codomain and `Prod` in the domain. -/ @[simps apply symm_apply] @@ -85,12 +90,16 @@ def prodPProd (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) : α₁ × β₁ ≃ PProd α₂ β₂ := (ea.symm.pprodProd eb.symm).symm #align equiv.prod_pprod Equiv.prodPProd +#align equiv.prod_pprod_symm_apply Equiv.prodPProd_symm_apply +#align equiv.prod_pprod_apply Equiv.prodPProd_apply /-- `PProd α β` is equivalent to `PLift α × PLift β` -/ @[simps apply symm_apply] def pprodEquivProdPLift : PProd α β ≃ PLift α × PLift β := Equiv.plift.symm.pprodProd Equiv.plift.symm #align equiv.pprod_equiv_prod_plift Equiv.pprodEquivProdPLift +#align equiv.pprod_equiv_prod_plift_symm_apply Equiv.pprodEquivProdPLift_symm_apply +#align equiv.pprod_equiv_prod_plift_apply Equiv.pprodEquivProdPLift_apply /-- Product of two equivalences. If `α₁ ≃ α₂` and `β₁ ≃ β₂`, then `α₁ × β₁ ≃ α₂ × β₂`. This is `Prod.map` as an equivalence. -/ @@ -99,6 +108,7 @@ def pprodEquivProdPLift : PProd α β ≃ PLift α × PLift β := def prodCongr (e₁ : α₁ ≃ α₂) (e₂ : β₁ ≃ β₂) : α₁ × β₁ ≃ α₂ × β₂ := ⟨Prod.map e₁ e₂, Prod.map e₁.symm e₂.symm, fun ⟨a, b⟩ => by simp, fun ⟨a, b⟩ => by simp⟩ #align equiv.prod_congr Equiv.prodCongr +#align equiv.prod_congr_apply Equiv.prodCongr_apply @[simp] theorem prodCongr_symm (e₁ : α₁ ≃ α₂) (e₂ : β₁ ≃ β₂) : @@ -133,6 +143,8 @@ def prodAssoc (α β γ) : (α × β) × γ ≃ α × β × γ := ⟨fun p => (p.1.1, p.1.2, p.2), fun p => ((p.1, p.2.1), p.2.2), fun ⟨⟨_, _⟩, _⟩ => rfl, fun ⟨_, ⟨_, _⟩⟩ => rfl⟩ #align equiv.prod_assoc Equiv.prodAssoc +#align equiv.prod_assoc_symm_apply Equiv.prodAssoc_symm_apply +#align equiv.prod_assoc_apply Equiv.prodAssoc_apply /-- `γ`-valued functions on `α × β` are equivalent to functions `α → β → γ`. -/ @[simps (config := { fullyApplied := false })] @@ -142,6 +154,8 @@ def curry (α β γ) : (α × β → γ) ≃ (α → β → γ) where left_inv := uncurry_curry right_inv := curry_uncurry #align equiv.curry Equiv.curry +#align equiv.curry_symm_apply Equiv.curry_symm_apply +#align equiv.curry_apply Equiv.curry_apply section @@ -150,6 +164,8 @@ section def prodPUnit (α) : α × PUnit ≃ α := ⟨fun p => p.1, fun a => (a, PUnit.unit), fun ⟨_, PUnit.unit⟩ => rfl, fun _ => rfl⟩ #align equiv.prod_punit Equiv.prodPUnit +#align equiv.prod_punit_apply Equiv.prodPUnit_apply +#align equiv.prod_punit_symm_apply Equiv.prodPUnit_symm_apply /-- `PUnit` is a left identity for type product up to an equivalence. -/ @[simps] @@ -158,6 +174,8 @@ def punitProd (α) : PUnit × α ≃ α := PUnit × α ≃ α × PUnit := prodComm _ _ _ ≃ α := prodPUnit _ #align equiv.punit_prod Equiv.punitProd +#align equiv.punit_prod_symm_apply Equiv.punitProd_symm_apply +#align equiv.punit_prod_apply Equiv.punitProd_apply /-- Any `Unique` type is a right identity for type product up to equivalence. -/ def prodUnique (α β) [Unique β] : α × β ≃ α := @@ -238,6 +256,7 @@ def psumEquivSum (α β) : PSum α β ≃ Sum α β where def sumCongr (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) : Sum α₁ β₁ ≃ Sum α₂ β₂ := ⟨Sum.map ea eb, Sum.map ea.symm eb.symm, fun x => by simp, fun x => by simp⟩ #align equiv.sum_congr Equiv.sumCongr +#align equiv.sum_congr_apply Equiv.sumCongr_apply /-- If `α ≃ α'` and `β ≃ β'`, then `PSum α β ≃ PSum α' β'`. -/ def psumCongr (e₁ : α ≃ β) (e₂ : γ ≃ δ) : PSum α γ ≃ PSum β δ where @@ -321,6 +340,7 @@ def boolEquivPUnitSumPUnit : Bool ≃ Sum PUnit.{u + 1} PUnit.{v + 1} := def sumComm (α β) : Sum α β ≃ Sum β α := ⟨Sum.swap, Sum.swap, Sum.swap_swap, Sum.swap_swap⟩ #align equiv.sum_comm Equiv.sumComm +#align equiv.sum_comm_apply Equiv.sumComm_apply @[simp] theorem sumComm_symm (α β) : (sumComm α β).symm = sumComm β α := @@ -377,6 +397,7 @@ def sumEmpty (α β) [IsEmpty β] : Sum α β ≃ α where · exact isEmptyElim x right_inv _ := rfl #align equiv.sum_empty Equiv.sumEmpty +#align equiv.sum_empty_symm_apply Equiv.sumEmpty_symm_apply @[simp] theorem sumEmpty_apply_inl [IsEmpty β] (a : α) : sumEmpty α β (Sum.inl a) = a := @@ -388,6 +409,7 @@ theorem sumEmpty_apply_inl [IsEmpty β] (a : α) : sumEmpty α β (Sum.inl a) = def emptySum (α β) [IsEmpty α] : Sum α β ≃ β := (sumComm _ _).trans <| sumEmpty _ _ #align equiv.empty_sum Equiv.emptySum +#align equiv.empty_sum_symm_apply Equiv.emptySum_symm_apply @[simp] theorem emptySum_apply_inr [IsEmpty α] (b : β) : emptySum α β (Sum.inr b) = b := @@ -434,6 +456,8 @@ def optionIsSomeEquiv (α) : { x : Option α // x.isSome } ≃ α where left_inv _ := Subtype.eq <| Option.some_get _ right_inv _ := Option.get_some _ _ #align equiv.option_is_some_equiv Equiv.optionIsSomeEquiv +#align equiv.option_is_some_equiv_apply Equiv.optionIsSomeEquiv_apply +#align equiv.option_is_some_equiv_symm_apply_coe Equiv.optionIsSomeEquiv_symm_apply_coe /-- The product over `Option α` of `β a` is the binary product of the product over `α` of `β (some α)` and `β none` -/ @@ -445,6 +469,8 @@ def piOptionEquivProd {β : Option α → Type _} : left_inv f := funext fun a => by cases a <;> rfl right_inv x := by simp #align equiv.pi_option_equiv_prod Equiv.piOptionEquivProd +#align equiv.pi_option_equiv_prod_symm_apply Equiv.piOptionEquivProd_symm_apply +#align equiv.pi_option_equiv_prod_apply Equiv.piOptionEquivProd_apply /-- `α ⊕ β` is equivalent to a `Sigma`-type over `Bool`. Note that this definition assumes `α` and `β` to be types from the same universe, so it cannot by used directly to transfer theorems about @@ -465,6 +491,9 @@ the type of all fibres of `f` and the total space `α`. -/ def sigmaFiberEquiv {α β : Type _} (f : α → β) : (Σ y : β, { x // f x = y }) ≃ α := ⟨fun x => ↑x.2, fun x => ⟨f x, x, rfl⟩, fun ⟨_, _, rfl⟩ => rfl, fun _ => rfl⟩ #align equiv.sigma_fiber_equiv Equiv.sigmaFiberEquiv +#align equiv.sigma_fiber_equiv_apply Equiv.sigmaFiberEquiv_apply +#align equiv.sigma_fiber_equiv_symm_apply_fst Equiv.sigmaFiberEquiv_symm_apply_fst +#align equiv.sigma_fiber_equiv_symm_apply_snd_coe Equiv.sigmaFiberEquiv_symm_apply_snd_coe end @@ -614,6 +643,8 @@ def subtypePreimage : { x : α → β // x ∘ Subtype.val = x₀ } ≃ ({ a // dsimp only rw [dif_neg h] #align equiv.subtype_preimage Equiv.subtypePreimage +#align equiv.subtype_preimage_symm_apply_coe Equiv.subtypePreimage_symm_apply_coe +#align equiv.subtype_preimage_apply Equiv.subtypePreimage_apply theorem subtypePreimage_symm_apply_coe_pos (x : { a // ¬p a } → β) (a : α) (h : p a) : ((subtypePreimage p x₀).symm x : α → β) a = x₀ ⟨a, h⟩ := @@ -642,6 +673,7 @@ This is `Function.swap` as an `Equiv`. -/ def piComm (φ : α → β → Sort _) : (∀ a b, φ a b) ≃ ∀ b a, φ a b := ⟨swap, swap, fun _ => rfl, fun _ => rfl⟩ #align equiv.Pi_comm Equiv.piComm +#align equiv.Pi_comm_apply Equiv.piComm_apply @[simp] theorem piComm_symm {φ : α → β → Sort _} : (piComm φ).symm = (piComm <| swap φ) := @@ -749,6 +781,8 @@ theorem sigmaEquivProd_sigmaCongrRight : def ofFiberEquiv {f : α → γ} {g : β → γ} (e : ∀ c, { a // f a = c } ≃ { b // g b = c }) : α ≃ β := (sigmaFiberEquiv f).symm.trans <| (Equiv.sigmaCongrRight e).trans (sigmaFiberEquiv g) #align equiv.of_fiber_equiv Equiv.ofFiberEquiv +#align equiv.of_fiber_equiv_apply Equiv.ofFiberEquiv_apply +#align equiv.of_fiber_equiv_symm_apply Equiv.ofFiberEquiv_symm_apply theorem ofFiberEquiv_map {α β γ} {f : α → γ} {g : β → γ} (e : ∀ c, { a // f a = c } ≃ { b // g b = c }) (a : α) : g (ofFiberEquiv e a) = f a := @@ -769,6 +803,8 @@ def prodShear (e₁ : α₁ ≃ α₂) (e₂ : α₁ → β₁ ≃ β₂) : α rintro ⟨x₁, y₁⟩ simp only [apply_symm_apply] #align equiv.prod_shear Equiv.prodShear +#align equiv.prod_shear_apply Equiv.prodShear_apply +#align equiv.prod_shear_symm_apply Equiv.prodShear_symm_apply end prodCongr @@ -940,6 +976,8 @@ def sigmaSumDistrib (α β : ι → Type _) : Sum.elim (Sigma.map id fun _ => Sum.inl) (Sigma.map id fun _ => Sum.inr), fun p => by rcases p with ⟨i, a | b⟩ <;> rfl, fun p => by rcases p with (⟨i, a⟩ | ⟨i, b⟩) <;> rfl⟩ #align equiv.sigma_sum_distrib Equiv.sigmaSumDistrib +#align equiv.sigma_sum_distrib_apply Equiv.sigmaSumDistrib_apply +#align equiv.sigma_sum_distrib_symm_apply Equiv.sigmaSumDistrib_symm_apply /-- The product of an indexed sum of types (formally, a `Sigma`-type `Σ i, α i`) by a type `β` is equivalent to the sum of products `Σ i, (α i × β)`. -/ @@ -969,6 +1007,8 @@ def boolProdEquivSum (α) : Bool × α ≃ Sum α α where left_inv := by rintro ⟨_ | _, _⟩ <;> rfl right_inv := by rintro (_ | _) <;> rfl #align equiv.bool_prod_equiv_sum Equiv.boolProdEquivSum +#align equiv.bool_prod_equiv_sum_apply Equiv.boolProdEquivSum_apply +#align equiv.bool_prod_equiv_sum_symm_apply Equiv.boolProdEquivSum_symm_apply /-- The function type `Bool → α` is equivalent to `α × α`. -/ @[simps] @@ -978,6 +1018,8 @@ def boolArrowEquivProd (α) : (Bool → α) ≃ α × α where left_inv _ := funext <| Bool.forall_bool.2 ⟨rfl, rfl⟩ right_inv := fun _ => rfl #align equiv.bool_arrow_equiv_prod Equiv.boolArrowEquivProd +#align equiv.bool_arrow_equiv_prod_apply Equiv.boolArrowEquivProd_apply +#align equiv.bool_arrow_equiv_prod_symm_apply Equiv.boolArrowEquivProd_symm_apply end @@ -1085,6 +1127,8 @@ theorem subtypeEquiv_apply {p : α → Prop} {q : β → Prop} def subtypeEquivRight {p q : α → Prop} (e : ∀ x, p x ↔ q x) : { x // p x } ≃ { x // q x } := subtypeEquiv (Equiv.refl _) e #align equiv.subtype_equiv_right Equiv.subtypeEquivRight +#align equiv.subtype_equiv_right_apply_coe Equiv.subtypeEquivRight_apply_coe +#align equiv.subtype_equiv_right_symm_apply_coe Equiv.subtypeEquivRight_symm_apply_coe /-- If `α ≃ β`, then for any predicate `p : β → Prop` the subtype `{a // p (e a)}` is equivalent to the subtype `{b // p b}`. -/ @@ -1115,6 +1159,8 @@ def subtypeSubtypeEquivSubtypeExists (p : α → Prop) (q : Subtype p → Prop) exact haq⟩, fun a => ⟨⟨a, a.2.fst⟩, a.2.snd⟩, fun ⟨⟨a, ha⟩, h⟩ => rfl, fun ⟨a, h₁, h₂⟩ => rfl⟩ #align equiv.subtype_subtype_equiv_subtype_exists Equiv.subtypeSubtypeEquivSubtypeExists +#align equiv.subtype_subtype_equiv_subtype_exists_symm_apply_coe_coe Equiv.subtypeSubtypeEquivSubtypeExists_symm_apply_coe_coe +#align equiv.subtype_subtype_equiv_subtype_exists_apply_coe Equiv.subtypeSubtypeEquivSubtypeExists_apply_coe /-- A subtype of a subtype is equivalent to the subtype of elements satisfying both predicates. -/ @[simps] @@ -1123,6 +1169,8 @@ def subtypeSubtypeEquivSubtypeInter (p q : α → Prop) : (subtypeSubtypeEquivSubtypeExists p _).trans <| subtypeEquivRight fun x => @exists_prop (q x) (p x) #align equiv.subtype_subtype_equiv_subtype_inter Equiv.subtypeSubtypeEquivSubtypeInter +#align equiv.subtype_subtype_equiv_subtype_inter_apply_coe Equiv.subtypeSubtypeEquivSubtypeInter_apply_coe +#align equiv.subtype_subtype_equiv_subtype_inter_symm_apply_coe_coe Equiv.subtypeSubtypeEquivSubtypeInter_symm_apply_coe_coe /-- If the outer subtype has more restrictive predicate than the inner one, then we can drop the latter. -/ @@ -1131,6 +1179,8 @@ def subtypeSubtypeEquivSubtype {p q : α → Prop} (h : ∀ {x}, q x → p x) : { x : Subtype p // q x.1 } ≃ Subtype q := (subtypeSubtypeEquivSubtypeInter p _).trans <| subtypeEquivRight fun _ => and_iff_right_of_imp h #align equiv.subtype_subtype_equiv_subtype Equiv.subtypeSubtypeEquivSubtype +#align equiv.subtype_subtype_equiv_subtype_apply_coe Equiv.subtypeSubtypeEquivSubtype_apply_coe +#align equiv.subtype_subtype_equiv_subtype_symm_apply_coe_coe Equiv.subtypeSubtypeEquivSubtype_symm_apply_coe_coe /-- If a proposition holds for all elements, then the subtype is equivalent to the original type. -/ @@ -1138,6 +1188,8 @@ equivalent to the original type. -/ def subtypeUnivEquiv {p : α → Prop} (h : ∀ x, p x) : Subtype p ≃ α := ⟨fun x => x, fun x => ⟨x, h x⟩, fun _ => Subtype.eq rfl, fun _ => rfl⟩ #align equiv.subtype_univ_equiv Equiv.subtypeUnivEquiv +#align equiv.subtype_univ_equiv_apply Equiv.subtypeUnivEquiv_apply +#align equiv.subtype_univ_equiv_symm_apply Equiv.subtypeUnivEquiv_symm_apply /-- A subtype of a sigma-type is a sigma-type over a subtype. -/ def subtypeSigmaEquiv (p : α → Type v) (q : α → Prop) : { y : Sigma p // q y.1 } ≃ Σ x : @@ -1258,6 +1310,8 @@ def piEquivPiSubtypeProd {α : Type _} (p : α → Prop) (β : α → Type _) [D by_cases h:p x <;> · simp only [h, dif_neg, dif_pos, not_false_iff] #align equiv.pi_equiv_pi_subtype_prod Equiv.piEquivPiSubtypeProd +#align equiv.pi_equiv_pi_subtype_prod_symm_apply Equiv.piEquivPiSubtypeProd_symm_apply +#align equiv.pi_equiv_pi_subtype_prod_apply Equiv.piEquivPiSubtypeProd_apply /-- A product of types can be split as the binary product of one of the types and the product of all the remaining types. -/ @@ -1276,6 +1330,8 @@ def piSplitAt {α : Type _} [DecidableEq α] (i : α) (β : α → Type _) : · subst h; rfl · rfl #align equiv.pi_split_at Equiv.piSplitAt +#align equiv.pi_split_at_apply Equiv.piSplitAt_apply +#align equiv.pi_split_at_symm_apply Equiv.piSplitAt_symm_apply /-- A product of copies of a type can be split as the binary product of one copy and the product of all the remaining copies. -/ @@ -1284,6 +1340,8 @@ def funSplitAt {α : Type _} [DecidableEq α] (i : α) (β : Type _) : (α → β) ≃ β × ({ j // j ≠ i } → β) := piSplitAt i _ #align equiv.fun_split_at Equiv.funSplitAt +#align equiv.fun_split_at_symm_apply Equiv.funSplitAt_symm_apply +#align equiv.fun_split_at_apply Equiv.funSplitAt_apply end @@ -1355,6 +1413,7 @@ noncomputable def ofBijective (f : α → β) (hf : Bijective f) : α ≃ β whe left_inv := Function.leftInverse_surjInv hf right_inv := Function.rightInverse_surjInv _ #align equiv.of_bijective Equiv.ofBijective +#align equiv.of_bijective_apply Equiv.ofBijective_apply theorem ofBijective_apply_symm_apply (f : α → β) (hf : Bijective f) (x : β) : f ((ofBijective f hf).symm x) = x := @@ -1706,6 +1765,8 @@ def piCongrLeft' (P : α → Sort _) (e : α ≃ β) : (∀ a, P a) ≃ ∀ b, P (by rintro _ rfl; rfl : ∀ {y} (h : y = x), (congr_arg e.symm h).ndrec (f y) = f x) (e.apply_symm_apply x) #align equiv.Pi_congr_left' Equiv.piCongrLeft' +#align equiv.Pi_congr_left'_apply Equiv.piCongrLeft'_apply +#align equiv.Pi_congr_left'_symm_apply Equiv.piCongrLeft'_symm_apply end diff --git a/Mathlib/Logic/Equiv/Defs.lean b/Mathlib/Logic/Equiv/Defs.lean index c4c8121c9e939..a162d5fd4db3e 100644 --- a/Mathlib/Logic/Equiv/Defs.lean +++ b/Mathlib/Logic/Equiv/Defs.lean @@ -73,6 +73,7 @@ structure Equiv (α : Sort _) (β : Sort _) where invFun : β → α left_inv : LeftInverse invFun toFun right_inv : RightInverse invFun toFun +#align equiv Equiv infixl:25 " ≃ " => Equiv @@ -93,6 +94,7 @@ instance {F} [EquivLike F α β] : CoeTC F (α ≃ β) := @[reducible] def Equiv.Perm (α : Sort _) := Equiv α α +#align equiv.perm Equiv.Perm namespace Equiv @@ -105,47 +107,61 @@ instance : EquivLike (α ≃ β) α β where @[simp] theorem coe_fn_mk (f : α → β) (g l r) : (Equiv.mk f g l r : α → β) = f := rfl +#align equiv.coe_fn_mk Equiv.coe_fn_mk /-- The map `(r ≃ s) → (r → s)` is injective. -/ theorem coe_fn_injective : @Function.Injective (α ≃ β) (α → β) (fun e => e) := FunLike.coe_injective' +#align equiv.coe_fn_injective Equiv.coe_fn_injective protected theorem coe_inj {e₁ e₂ : α ≃ β} : (e₁ : α → β) = e₂ ↔ e₁ = e₂ := @FunLike.coe_fn_eq _ _ _ _ e₁ e₂ +#align equiv.coe_inj Equiv.coe_inj @[ext] theorem ext {f g : Equiv α β} (H : ∀ x, f x = g x) : f = g := FunLike.ext f g H +#align equiv.ext Equiv.ext protected theorem congr_arg {f : Equiv α β} {x x' : α} : x = x' → f x = f x' := FunLike.congr_arg f +#align equiv.congr_arg Equiv.congr_arg protected theorem congr_fun {f g : Equiv α β} (h : f = g) (x : α) : f x = g x := FunLike.congr_fun h x +#align equiv.congr_fun Equiv.congr_fun theorem ext_iff {f g : Equiv α β} : f = g ↔ ∀ x, f x = g x := FunLike.ext_iff +#align equiv.ext_iff Equiv.ext_iff @[ext] theorem Perm.ext {σ τ : Equiv.Perm α} (H : ∀ x, σ x = τ x) : σ = τ := Equiv.ext H +#align equiv.perm.ext Equiv.Perm.ext protected theorem Perm.congr_arg {f : Equiv.Perm α} {x x' : α} : x = x' → f x = f x' := Equiv.congr_arg +#align equiv.perm.congr_arg Equiv.Perm.congr_arg protected theorem Perm.congr_fun {f g : Equiv.Perm α} (h : f = g) (x : α) : f x = g x := Equiv.congr_fun h x +#align equiv.perm.congr_fun Equiv.Perm.congr_fun theorem Perm.ext_iff {σ τ : Equiv.Perm α} : σ = τ ↔ ∀ x, σ x = τ x := Equiv.ext_iff +#align equiv.perm.ext_iff Equiv.Perm.ext_iff /-- Any type is equivalent to itself. -/ @[refl] protected def refl (α : Sort _) : α ≃ α := ⟨id, id, fun _ => rfl, fun _ => rfl⟩ +#align equiv.refl Equiv.refl instance inhabited' : Inhabited (α ≃ α) := ⟨Equiv.refl α⟩ /-- Inverse of an equivalence `e : α ≃ β`. -/ @[symm] protected def symm (e : α ≃ β) : β ≃ α := ⟨e.invFun, e.toFun, e.right_inv, e.left_inv⟩ +#align equiv.symm Equiv.symm /-- See Note [custom simps projection] -/ def Simps.apply (e : α ≃ β) : α → β := e /-- See Note [custom simps projection] -/ def Simps.symm_apply (e : α ≃ β) : β → α := e.symm +#align equiv.simps.symm_apply Equiv.Simps.symm_apply initialize_simps_projections Equiv (toFun → apply, invFun → symm_apply) @@ -160,6 +176,7 @@ theorem right_inv' (e : α ≃ β) : Function.RightInverse e.symm e := e.right_i @[trans] protected def trans (e₁ : α ≃ β) (e₂ : β ≃ γ) : α ≃ γ := ⟨e₂ ∘ e₁, e₁.symm ∘ e₂.symm, e₂.left_inv.comp e₁.left_inv, e₂.right_inv.comp e₁.right_inv⟩ +#align equiv.trans Equiv.trans instance : Trans Equiv Equiv Equiv where trans := Equiv.trans @@ -193,9 +210,11 @@ protected theorem subsingleton (e : α ≃ β) [Subsingleton β] : Subsingleton protected theorem subsingleton.symm (e : α ≃ β) [Subsingleton α] : Subsingleton β := e.symm.injective.subsingleton +#align equiv.subsingleton.symm Equiv.subsingleton.symm theorem subsingleton_congr (e : α ≃ β) : Subsingleton α ↔ Subsingleton β := ⟨fun _ => e.symm.subsingleton, fun _ => e.subsingleton⟩ +#align equiv.subsingleton_congr Equiv.subsingleton_congr instance equiv_subsingleton_cod [Subsingleton β] : Subsingleton (α ≃ β) := ⟨fun _ _ => Equiv.ext fun _ => Subsingleton.elim _ _⟩ @@ -208,120 +227,163 @@ instance permUnique [Subsingleton α] : Unique (Perm α) := theorem Perm.subsingleton_eq_refl [Subsingleton α] (e : Perm α) : e = Equiv.refl α := Subsingleton.elim _ _ +#align equiv.perm.subsingleton_eq_refl Equiv.Perm.subsingleton_eq_refl /-- Transfer `DecidableEq` across an equivalence. -/ protected def decidableEq (e : α ≃ β) [DecidableEq β] : DecidableEq α := e.injective.decidableEq +#align equiv.decidable_eq Equiv.decidableEq theorem nonempty_congr (e : α ≃ β) : Nonempty α ↔ Nonempty β := Nonempty.congr e e.symm +#align equiv.nonempty_congr Equiv.nonempty_congr protected theorem nonempty (e : α ≃ β) [Nonempty β] : Nonempty α := e.nonempty_congr.mpr ‹_› +#align equiv.nonempty Equiv.nonempty /-- If `α ≃ β` and `β` is inhabited, then so is `α`. -/ protected def inhabited [Inhabited β] (e : α ≃ β) : Inhabited α := ⟨e.symm default⟩ +#align equiv.inhabited Equiv.inhabited /-- If `α ≃ β` and `β` is a singleton type, then so is `α`. -/ protected def unique [Unique β] (e : α ≃ β) : Unique α := e.symm.surjective.unique +#align equiv.unique Equiv.unique /-- Equivalence between equal types. -/ protected def cast {α β : Sort _} (h : α = β) : α ≃ β := ⟨cast h, cast h.symm, fun _ => by cases h; rfl, fun _ => by cases h; rfl⟩ +#align equiv.cast Equiv.cast @[simp] theorem coe_fn_symm_mk (f : α → β) (g l r) : ((Equiv.mk f g l r).symm : β → α) = g := rfl +#align equiv.coe_fn_symm_mk Equiv.coe_fn_symm_mk @[simp] theorem coe_refl : (Equiv.refl α : α → α) = id := rfl +#align equiv.coe_refl Equiv.coe_refl /-- This cannot be a `simp` lemmas as it incorrectly matches against `e : α ≃ synonym α`, when `synonym α` is semireducible. This makes a mess of `multiplicative.of_add` etc. -/ theorem Perm.coe_subsingleton {α : Type _} [Subsingleton α] (e : Perm α) : (e : α → α) = id := by rw [Perm.subsingleton_eq_refl e, coe_refl] +#align equiv.perm.coe_subsingleton Equiv.Perm.coe_subsingleton -- porting note: marking this as `@[simp]` because `simp` doesn't fire on `coe_refl` -- in an expression such as `Equiv.refl a x` @[simp] theorem refl_apply (x : α) : Equiv.refl α x = x := rfl +#align equiv.refl_apply Equiv.refl_apply @[simp] theorem coe_trans (f : α ≃ β) (g : β ≃ γ) : (f.trans g : α → γ) = g ∘ f := rfl +#align equiv.coe_trans Equiv.coe_trans -- porting note: marking this as `@[simp]` because `simp` doesn't fire on `coe_trans` -- in an expression such as `Equiv.trans f g x` @[simp] theorem trans_apply (f : α ≃ β) (g : β ≃ γ) (a : α) : (f.trans g) a = g (f a) := rfl +#align equiv.trans_apply Equiv.trans_apply @[simp] theorem apply_symm_apply (e : α ≃ β) (x : β) : e (e.symm x) = x := e.right_inv x +#align equiv.apply_symm_apply Equiv.apply_symm_apply @[simp] theorem symm_apply_apply (e : α ≃ β) (x : α) : e.symm (e x) = x := e.left_inv x +#align equiv.symm_apply_apply Equiv.symm_apply_apply @[simp] theorem symm_comp_self (e : α ≃ β) : e.symm ∘ e = id := funext e.symm_apply_apply +#align equiv.symm_comp_self Equiv.symm_comp_self @[simp] theorem self_comp_symm (e : α ≃ β) : e ∘ e.symm = id := funext e.apply_symm_apply +#align equiv.self_comp_symm Equiv.self_comp_symm @[simp] theorem symm_trans_apply (f : α ≃ β) (g : β ≃ γ) (a : γ) : (f.trans g).symm a = f.symm (g.symm a) := rfl +#align equiv.symm_trans_apply Equiv.symm_trans_apply -- The `simp` attribute is needed to make this a `dsimp` lemma. -- `simp` will always rewrite with `equiv.symm_symm` before this has a chance to fire. @[simp, nolint simpNF] theorem symm_symm_apply (f : α ≃ β) (b : α) : f.symm.symm b = f b := rfl +#align equiv.symm_symm_apply Equiv.symm_symm_apply theorem apply_eq_iff_eq (f : α ≃ β) {x y : α} : f x = f y ↔ x = y := EquivLike.apply_eq_iff_eq f +#align equiv.apply_eq_iff_eq Equiv.apply_eq_iff_eq theorem apply_eq_iff_eq_symm_apply (f : α ≃ β) : f x = y ↔ x = f.symm y := by conv_lhs => rw [← apply_symm_apply f y] rw [apply_eq_iff_eq] +#align equiv.apply_eq_iff_eq_symm_apply Equiv.apply_eq_iff_eq_symm_apply @[simp] theorem cast_apply {α β} (h : α = β) (x : α) : Equiv.cast h x = cast h x := rfl +#align equiv.cast_apply Equiv.cast_apply @[simp] theorem cast_symm {α β} (h : α = β) : (Equiv.cast h).symm = Equiv.cast h.symm := rfl +#align equiv.cast_symm Equiv.cast_symm @[simp] theorem cast_refl {α} (h : α = α := rfl) : Equiv.cast h = Equiv.refl α := rfl +#align equiv.cast_refl Equiv.cast_refl @[simp] theorem cast_trans {α β γ} (h : α = β) (h2 : β = γ) : (Equiv.cast h).trans (Equiv.cast h2) = Equiv.cast (h.trans h2) := ext fun x => by substs h h2; rfl +#align equiv.cast_trans Equiv.cast_trans theorem cast_eq_iff_heq {α β} (h : α = β) {a : α} {b : β} : Equiv.cast h a = b ↔ HEq a b := by subst h; simp [coe_refl] +#align equiv.cast_eq_iff_heq Equiv.cast_eq_iff_heq theorem symm_apply_eq {α β} (e : α ≃ β) {x y} : e.symm x = y ↔ x = e y := ⟨fun H => by simp [H.symm], fun H => by simp [H]⟩ +#align equiv.symm_apply_eq Equiv.symm_apply_eq theorem eq_symm_apply {α β} (e : α ≃ β) {x y} : y = e.symm x ↔ e y = x := (eq_comm.trans e.symm_apply_eq).trans eq_comm +#align equiv.eq_symm_apply Equiv.eq_symm_apply @[simp] theorem symm_symm (e : α ≃ β) : e.symm.symm = e := by cases e; rfl +#align equiv.symm_symm Equiv.symm_symm @[simp] theorem trans_refl (e : α ≃ β) : e.trans (Equiv.refl β) = e := by cases e; rfl +#align equiv.trans_refl Equiv.trans_refl @[simp] theorem refl_symm : (Equiv.refl α).symm = Equiv.refl α := rfl +#align equiv.refl_symm Equiv.refl_symm @[simp] theorem refl_trans (e : α ≃ β) : (Equiv.refl α).trans e = e := by cases e; rfl +#align equiv.refl_trans Equiv.refl_trans @[simp] theorem symm_trans_self (e : α ≃ β) : e.symm.trans e = Equiv.refl β := ext <| by simp +#align equiv.symm_trans_self Equiv.symm_trans_self @[simp] theorem self_trans_symm (e : α ≃ β) : e.trans e.symm = Equiv.refl α := ext <| by simp +#align equiv.self_trans_symm Equiv.self_trans_symm theorem trans_assoc {δ} (ab : α ≃ β) (bc : β ≃ γ) (cd : γ ≃ δ) : (ab.trans bc).trans cd = ab.trans (bc.trans cd) := Equiv.ext fun _ => rfl +#align equiv.trans_assoc Equiv.trans_assoc theorem leftInverse_symm (f : Equiv α β) : LeftInverse f.symm f := f.left_inv +#align equiv.left_inverse_symm Equiv.leftInverse_symm theorem rightInverse_symm (f : Equiv α β) : Function.RightInverse f.symm f := f.right_inv +#align equiv.right_inverse_symm Equiv.rightInverse_symm theorem injective_comp (e : α ≃ β) (f : β → γ) : Injective (f ∘ e) ↔ Injective f := EquivLike.injective_comp e f +#align equiv.injective_comp Equiv.injective_comp theorem comp_injective (f : α → β) (e : β ≃ γ) : Injective (e ∘ f) ↔ Injective f := EquivLike.comp_injective f e +#align equiv.comp_injective Equiv.comp_injective theorem surjective_comp (e : α ≃ β) (f : β → γ) : Surjective (f ∘ e) ↔ Surjective f := EquivLike.surjective_comp e f +#align equiv.surjective_comp Equiv.surjective_comp theorem comp_surjective (f : α → β) (e : β ≃ γ) : Surjective (e ∘ f) ↔ Surjective f := EquivLike.comp_surjective f e +#align equiv.comp_surjective Equiv.comp_surjective theorem bijective_comp (e : α ≃ β) (f : β → γ) : Bijective (f ∘ e) ↔ Bijective f := EquivLike.bijective_comp e f +#align equiv.bijective_comp Equiv.bijective_comp theorem comp_bijective (f : α → β) (e : β ≃ γ) : Bijective (e ∘ f) ↔ Bijective f := EquivLike.comp_bijective f e +#align equiv.comp_bijective Equiv.comp_bijective /-- If `α` is equivalent to `β` and `γ` is equivalent to `δ`, then the type of equivalences `α ≃ γ` is equivalent to the type of equivalences `β ≃ δ`. -/ @@ -330,6 +392,7 @@ def equivCongr (ab : α ≃ β) (cd : γ ≃ δ) : (α ≃ γ) ≃ (β ≃ δ) w invFun bd := ab.trans <| bd.trans <| cd.symm left_inv ac := by ext x; simp only [trans_apply, comp_apply, symm_apply_apply] right_inv ac := by ext x; simp only [trans_apply, comp_apply, apply_symm_apply] +#align equiv.equiv_congr Equiv.equivCongr @[simp] theorem equivCongr_refl {α β} : (Equiv.refl α).equivCongr (Equiv.refl β) = Equiv.refl (α ≃ β) := by ext; rfl @@ -362,6 +425,7 @@ variable {α' β' : Type _} (e : α' ≃ β') /-- If `α` is equivalent to `β`, then `Perm α` is equivalent to `Perm β`. -/ def permCongr : Perm α' ≃ Perm β' := equivCongr e e +#align equiv.perm_congr Equiv.permCongr theorem permCongr_def (p : Equiv.Perm α') : e.permCongr p = (e.symm.trans p).trans e := rfl #align equiv.perm_congr_def Equiv.permCongr_def @@ -390,9 +454,11 @@ end permCongr /-- Two empty types are equivalent. -/ def equivOfIsEmpty (α β : Sort _) [IsEmpty α] [IsEmpty β] : α ≃ β := ⟨isEmptyElim, isEmptyElim, isEmptyElim, isEmptyElim⟩ +#align equiv.equiv_of_is_empty Equiv.equivOfIsEmpty /-- If `α` is an empty type, then it is equivalent to the `Empty` type. -/ def equivEmpty (α : Sort u) [IsEmpty α] : α ≃ Empty := equivOfIsEmpty α _ +#align equiv.equiv_empty Equiv.equivEmpty /-- If `α` is an empty type, then it is equivalent to the `PEmpty` type in any universe. -/ def equivPEmpty (α : Sort v) [IsEmpty α] : α ≃ PEmpty.{u} := equivOfIsEmpty α _ @@ -413,6 +479,7 @@ def equivOfUnique (α β : Sort _) [Unique.{u} α] [Unique.{v} β] : α ≃ β w invFun := default left_inv _ := Subsingleton.elim _ _ right_inv _ := Subsingleton.elim _ _ +#align equiv.equiv_of_unique Equiv.equivOfUnique /-- If `α` has a unique element, then it is equivalent to any `PUnit`. -/ def equivPUnit (α : Sort u) [Unique α] : α ≃ PUnit.{v} := equivOfUnique α _ @@ -426,13 +493,18 @@ def propEquivPUnit {p : Prop} (h : p) : p ≃ PUnit.{0} := @equivPUnit p <| uniq @[simps (config := { fullyApplied := false }) apply] protected def ulift {α : Type v} : ULift.{u} α ≃ α := ⟨ULift.down, ULift.up, ULift.up_down, fun _ => rfl⟩ +#align equiv.ulift Equiv.ulift +#align equiv.ulift_apply Equiv.ulift_apply /-- `PLift α` is equivalent to `α`. -/ @[simps (config := { fullyApplied := false }) apply] protected def plift : PLift α ≃ α := ⟨PLift.down, PLift.up, PLift.up_down, PLift.down_up⟩ +#align equiv.plift Equiv.plift +#align equiv.plift_apply Equiv.plift_apply /-- equivalence of propositions is the same as iff -/ def ofIff {P Q : Prop} (h : P ↔ Q) : P ≃ Q := ⟨h.mp, h.mpr, fun _ => rfl, fun _ => rfl⟩ +#align equiv.of_iff Equiv.ofIff /-- If `α₁` is equivalent to `α₂` and `β₁` is equivalent to `β₂`, then the type of maps `α₁ → β₁` is equivalent to the type of maps `α₂ → β₂`. -/ @@ -444,6 +516,7 @@ def arrowCongr {α₁ β₁ α₂ β₂ : Sort _} (e₁ : α₁ ≃ α₂) (e₂ left_inv f := funext fun x => by simp only [comp_apply, symm_apply_apply] right_inv f := funext fun x => by simp only [comp_apply, apply_symm_apply] #align equiv.arrow_congr_apply Equiv.arrowCongr_apply +#align equiv.arrow_congr Equiv.arrowCongr theorem arrowCongr_comp {α₁ β₁ γ₁ α₂ β₂ γ₂ : Sort _} (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) (ec : γ₁ ≃ γ₂) (f : α₁ → β₁) (g : β₁ → γ₁) : @@ -472,6 +545,8 @@ because Lean's universe rules will not unify `?l_1` with `imax (1 ?m_1)`. @[simps apply] def arrowCongr' {α₁ β₁ α₂ β₂ : Type _} (hα : α₁ ≃ α₂) (hβ : β₁ ≃ β₂) : (α₁ → β₁) ≃ (α₂ → β₂) := Equiv.arrowCongr hα hβ +#align equiv.arrow_congr' Equiv.arrowCongr' +#align equiv.arrow_congr'_apply Equiv.arrowCongr'_apply @[simp] theorem arrowCongr'_refl {α β : Type _} : arrowCongr' (Equiv.refl α) (Equiv.refl β) = Equiv.refl (α → β) := rfl @@ -489,31 +564,41 @@ def arrowCongr' {α₁ β₁ α₂ β₂ : Type _} (hα : α₁ ≃ α₂) (hβ /-- Conjugate a map `f : α → α` by an equivalence `α ≃ β`. -/ @[simps apply] def conj (e : α ≃ β) : (α → α) ≃ (β → β) := arrowCongr e e +#align equiv.conj Equiv.conj +#align equiv.conj_apply Equiv.conj_apply @[simp] theorem conj_refl : conj (Equiv.refl α) = Equiv.refl (α → α) := rfl +#align equiv.conj_refl Equiv.conj_refl @[simp] theorem conj_symm (e : α ≃ β) : e.conj.symm = e.symm.conj := rfl +#align equiv.conj_symm Equiv.conj_symm @[simp] theorem conj_trans (e₁ : α ≃ β) (e₂ : β ≃ γ) : (e₁.trans e₂).conj = e₁.conj.trans e₂.conj := rfl +#align equiv.conj_trans Equiv.conj_trans -- This should not be a simp lemma as long as `(∘)` is reducible: -- when `(∘)` is reducible, Lean can unify `f₁ ∘ f₂` with any `g` using -- `f₁ := g` and `f₂ := λ x, x`. This causes nontermination. theorem conj_comp (e : α ≃ β) (f₁ f₂ : α → α) : e.conj (f₁ ∘ f₂) = e.conj f₁ ∘ e.conj f₂ := by apply arrowCongr_comp +#align equiv.conj_comp Equiv.conj_comp theorem eq_comp_symm {α β γ} (e : α ≃ β) (f : β → γ) (g : α → γ) : f = g ∘ e.symm ↔ f ∘ e = g := (e.arrowCongr (Equiv.refl γ)).symm_apply_eq.symm +#align equiv.eq_comp_symm Equiv.eq_comp_symm theorem comp_symm_eq {α β γ} (e : α ≃ β) (f : β → γ) (g : α → γ) : g ∘ e.symm = f ↔ g = f ∘ e := (e.arrowCongr (Equiv.refl γ)).eq_symm_apply.symm +#align equiv.comp_symm_eq Equiv.comp_symm_eq theorem eq_symm_comp {α β γ} (e : α ≃ β) (f : γ → α) (g : γ → β) : f = e.symm ∘ g ↔ e ∘ f = g := ((Equiv.refl γ).arrowCongr e).eq_symm_apply +#align equiv.eq_symm_comp Equiv.eq_symm_comp theorem symm_comp_eq {α β γ} (e : α ≃ β) (f : γ → α) (g : γ → β) : e.symm ∘ g = f ↔ g = e ∘ f := ((Equiv.refl γ).arrowCongr e).symm_apply_eq +#align equiv.symm_comp_eq Equiv.symm_comp_eq /-- `PUnit` sorts in any two universes are equivalent. -/ def punitEquivPUnit : PUnit.{v} ≃ PUnit.{w} := @@ -526,6 +611,7 @@ noncomputable def propEquivBool : Prop ≃ Bool where invFun b := b left_inv p := by simp [@Bool.decide_iff p (Classical.propDecidable _)] right_inv b := by cases b <;> simp +#align equiv.Prop_equiv_bool Equiv.propEquivBool section @@ -541,16 +627,23 @@ is equivalent to `β a`. -/ invFun x b := cast (congr_arg β <| Subsingleton.elim a b) x left_inv _ := funext fun b => by rw [Subsingleton.elim b a]; rfl right_inv _ := rfl +#align equiv.Pi_subsingleton_apply Equiv.piSubsingleton_apply +#align equiv.Pi_subsingleton_symm_apply Equiv.piSubsingleton_symm_apply +#align equiv.Pi_subsingleton Equiv.piSubsingleton /-- If `α` has a unique term, then the type of function `α → β` is equivalent to `β`. -/ @[simps (config := { fullyApplied := false }) apply] def funUnique (α β) [Unique.{u} α] : (α → β) ≃ β := piSubsingleton _ default +#align equiv.fun_unique Equiv.funUnique +#align equiv.fun_unique_apply Equiv.funUnique_apply /-- The sort of maps from `PUnit` is equivalent to the codomain. -/ def punitArrowEquiv (α : Sort _) : (PUnit.{u} → α) ≃ α := funUnique PUnit.{u} α +#align equiv.punit_arrow_equiv Equiv.punitArrowEquiv /-- The sort of maps from `True` is equivalent to the codomain. -/ def trueArrowEquiv (α : Sort _) : (True → α) ≃ α := funUnique _ _ +#align equiv.true_arrow_equiv Equiv.trueArrowEquiv /-- The sort of maps from a type that `IsEmpty` is equivalent to `PUnit`. -/ def arrowPUnitOfIsEmpty (α β : Sort _) [IsEmpty α] : (α → β) ≃ PUnit.{u} where @@ -583,6 +676,9 @@ def psigmaEquivSigma {α} (β : α → Type _) : (Σ' i, β i) ≃ Σ i, β i wh invFun a := ⟨a.1, a.2⟩ left_inv _ := rfl right_inv _ := rfl +#align equiv.psigma_equiv_sigma Equiv.psigmaEquivSigma +#align equiv.psigma_equiv_sigma_symm_apply Equiv.psigmaEquivSigma_symm_apply +#align equiv.psigma_equiv_sigma_apply Equiv.psigmaEquivSigma_apply /-- A `PSigma`-type is equivalent to the corresponding `Sigma`-type. -/ @[simps apply symm_apply] @@ -592,6 +688,8 @@ def psigmaEquivSigmaPLift {α} (β : α → Sort _) : (Σ' i, β i) ≃ Σ i : P left_inv _ := rfl right_inv _ := rfl #align equiv.psigma_equiv_sigma_plift Equiv.psigmaEquivSigmaPLift +#align equiv.psigma_equiv_sigma_plift_symm_apply Equiv.psigmaEquivSigmaPLift_symm_apply +#align equiv.psigma_equiv_sigma_plift_apply Equiv.psigmaEquivSigmaPLift_apply /-- A family of equivalences `Π a, β₁ a ≃ β₂ a` generates an equivalence between `Σ' a, β₁ a` and `Σ' a, β₂ a`. -/ @@ -601,6 +699,8 @@ def psigmaCongrRight {β₁ β₂ : α → Sort _} (F : ∀ a, β₁ a ≃ β₂ invFun a := ⟨a.1, (F a.1).symm a.2⟩ left_inv | ⟨a, b⟩ => congr_arg (PSigma.mk a) <| symm_apply_apply (F a) b right_inv | ⟨a, b⟩ => congr_arg (PSigma.mk a) <| apply_symm_apply (F a) b +#align equiv.psigma_congr_right Equiv.psigmaCongrRight +#align equiv.psigma_congr_right_apply Equiv.psigmaCongrRight_apply -- Porting note: simp can now simplify the LHS, so I have removed `@[simp]` theorem psigmaCongrRight_trans {α} {β₁ β₂ β₃ : α → Sort _} @@ -628,6 +728,7 @@ def sigmaCongrRight {α} {β₁ β₂ : α → Type _} (F : ∀ a, β₁ a ≃ left_inv | ⟨a, b⟩ => congr_arg (Sigma.mk a) <| symm_apply_apply (F a) b right_inv | ⟨a, b⟩ => congr_arg (Sigma.mk a) <| apply_symm_apply (F a) b #align equiv.sigma_congr_right Equiv.sigmaCongrRight +#align equiv.sigma_congr_right_apply Equiv.sigmaCongrRight_apply -- Porting note: simp can now simplify the LHS, so I have removed `@[simp]` theorem sigmaCongrRight_trans {α} {β₁ β₂ β₃ : α → Type _} @@ -652,6 +753,7 @@ def psigmaEquivSubtype {α : Type v} (P : α → Prop) : (Σ' i, P i) ≃ Subtyp invFun x := ⟨x.1, x.2⟩ left_inv _ := rfl right_inv _ := rfl +#align equiv.psigma_equiv_subtype Equiv.psigmaEquivSubtype /-- A `Sigma` with `PLift` fibers is equivalent to the subtype. -/ def sigmaPLiftEquivSubtype {α : Type v} (P : α → Prop) : (Σ i, PLift (P i)) ≃ Subtype P := @@ -672,6 +774,7 @@ namespace Perm /-- A family of permutations `Π a, Perm (β a)` generates a permuation `Perm (Σ a, β₁ a)`. -/ @[reducible] def sigmaCongrRight {α} {β : α → Sort _} (F : ∀ a, Perm (β a)) : Perm (Σ a, β a) := Equiv.sigmaCongrRight F +#align equiv.perm.sigma_congr_right Equiv.Perm.sigmaCongrRight @[simp] theorem sigmaCongrRight_trans {α} {β : α → Sort _} (F : ∀ a, Perm (β a)) (G : ∀ a, Perm (β a)) : @@ -705,25 +808,33 @@ end Perm match (motive := ∀ a' (h : a' = a), Sigma.mk a' (h.symm ▸ b) = ⟨a, b⟩) e (e.symm a), e.apply_symm_apply _ with | _, rfl => rfl +#align equiv.sigma_congr_left_apply Equiv.sigmaCongrLeft_apply +#align equiv.sigma_congr_left Equiv.sigmaCongrLeft /-- Transporting a sigma type through an equivalence of the base -/ def sigmaCongrLeft' {α₁ α₂} {β : α₁ → Sort _} (f : α₁ ≃ α₂) : (Σ a : α₁, β a) ≃ Σ a : α₂, β (f.symm a) := (sigmaCongrLeft f.symm).symm +#align equiv.sigma_congr_left' Equiv.sigmaCongrLeft' /-- Transporting a sigma type through an equivalence of the base and a family of equivalences of matching fibers -/ def sigmaCongr {α₁ α₂} {β₁ : α₁ → Sort _} {β₂ : α₂ → Sort _} (f : α₁ ≃ α₂) (F : ∀ a, β₁ a ≃ β₂ (f a)) : Sigma β₁ ≃ Sigma β₂ := (sigmaCongrRight F).trans (sigmaCongrLeft f) +#align equiv.sigma_congr Equiv.sigmaCongr /-- `Sigma` type with a constant fiber is equivalent to the product. -/ @[simps apply symm_apply] def sigmaEquivProd (α β : Type _) : (Σ _ : α, β) ≃ α × β := ⟨fun a => ⟨a.1, a.2⟩, fun a => ⟨a.1, a.2⟩, fun ⟨_, _⟩ => rfl, fun ⟨_, _⟩ => rfl⟩ +#align equiv.sigma_equiv_prod_apply Equiv.sigmaEquivProd_apply +#align equiv.sigma_equiv_prod_symm_apply Equiv.sigmaEquivProd_symm_apply +#align equiv.sigma_equiv_prod Equiv.sigmaEquivProd /-- If each fiber of a `Sigma` type is equivalent to a fixed type, then the sigma type is equivalent to the product. -/ def sigmaEquivProdOfEquiv {α β} {β₁ : α → Sort _} (F : ∀ a, β₁ a ≃ β) : Sigma β₁ ≃ α × β := (sigmaCongrRight F).trans (sigmaEquivProd α β) +#align equiv.sigma_equiv_prod_of_equiv Equiv.sigmaEquivProdOfEquiv /-- Dependent product of types is associative up to an equivalence. -/ def sigmaAssoc {α : Type _} {β : α → Type _} (γ : ∀ a : α, β a → Type _) : @@ -732,6 +843,7 @@ def sigmaAssoc {α : Type _} {β : α → Type _} (γ : ∀ a : α, β a → Typ invFun x := ⟨⟨x.1, x.2.1⟩, x.2.2⟩ left_inv _ := rfl right_inv _ := rfl +#align equiv.sigma_assoc Equiv.sigmaAssoc end @@ -742,22 +854,27 @@ protected theorem exists_unique_congr {p : α → Prop} {q : β → Prop} exact ⟨f a, h.1 ha₁, fun b hb => f.symm_apply_eq.1 (ha₂ (f.symm b) (h.2 (by simpa using hb)))⟩ · rintro ⟨b, hb₁, hb₂⟩ exact ⟨f.symm b, h.2 (by simpa using hb₁), fun y hy => (eq_symm_apply f).2 (hb₂ _ (h.1 hy))⟩ +#align equiv.exists_unique_congr Equiv.exists_unique_congr protected theorem exists_unique_congr_left' {p : α → Prop} (f : α ≃ β) : (∃! x, p x) ↔ ∃! y, p (f.symm y) := Equiv.exists_unique_congr f fun {_} => by simp +#align equiv.exists_unique_congr_left' Equiv.exists_unique_congr_left' protected theorem exists_unique_congr_left {p : β → Prop} (f : α ≃ β) : (∃! x, p (f x)) ↔ ∃! y, p y := (Equiv.exists_unique_congr_left' f.symm).symm +#align equiv.exists_unique_congr_left Equiv.exists_unique_congr_left protected theorem forall_congr {p : α → Prop} {q : β → Prop} (f : α ≃ β) (h : ∀ {x}, p x ↔ q (f x)) : (∀ x, p x) ↔ (∀ y, q y) := by constructor <;> intro h₂ x . rw [← f.right_inv x]; apply h.mp; apply h₂ · apply h.mpr; apply h₂ +#align equiv.forall_congr Equiv.forall_congr protected theorem forall_congr' {p : α → Prop} {q : β → Prop} (f : α ≃ β) (h : ∀ {x}, p (f.symm x) ↔ q x) : (∀ x, p x) ↔ ∀ y, q y := (Equiv.forall_congr f.symm h.symm).symm +#align equiv.forall_congr' Equiv.forall_congr' -- We next build some higher arity versions of `Equiv.forall_congr`. -- Although they appear to just be repeated applications of `Equiv.forall_congr`, @@ -769,31 +886,38 @@ protected theorem forall_congr' {p : α → Prop} {q : β → Prop} (f : α ≃ protected theorem forall₂_congr {p : α₁ → β₁ → Prop} {q : α₂ → β₂ → Prop} (eα : α₁ ≃ α₂) (eβ : β₁ ≃ β₂) (h : ∀ {x y}, p x y ↔ q (eα x) (eβ y)) : (∀ x y, p x y) ↔ ∀ x y, q x y := Equiv.forall_congr _ <| Equiv.forall_congr _ h +#align equiv.forall₂_congr Equiv.forall₂_congr protected theorem forall₂_congr' {p : α₁ → β₁ → Prop} {q : α₂ → β₂ → Prop} (eα : α₁ ≃ α₂) (eβ : β₁ ≃ β₂) (h : ∀ {x y}, p (eα.symm x) (eβ.symm y) ↔ q x y) : (∀ x y, p x y) ↔ ∀ x y, q x y := (Equiv.forall₂_congr eα.symm eβ.symm h.symm).symm +#align equiv.forall₂_congr' Equiv.forall₂_congr' protected theorem forall₃_congr {p : α₁ → β₁ → γ₁ → Prop} {q : α₂ → β₂ → γ₂ → Prop} (eα : α₁ ≃ α₂) (eβ : β₁ ≃ β₂) (eγ : γ₁ ≃ γ₂) (h : ∀ {x y z}, p x y z ↔ q (eα x) (eβ y) (eγ z)) : (∀ x y z, p x y z) ↔ ∀ x y z, q x y z := Equiv.forall₂_congr _ _ <| Equiv.forall_congr _ h +#align equiv.forall₃_congr Equiv.forall₃_congr protected theorem forall₃_congr' {p : α₁ → β₁ → γ₁ → Prop} {q : α₂ → β₂ → γ₂ → Prop} (eα : α₁ ≃ α₂) (eβ : β₁ ≃ β₂) (eγ : γ₁ ≃ γ₂) (h : ∀ {x y z}, p (eα.symm x) (eβ.symm y) (eγ.symm z) ↔ q x y z) : (∀ x y z, p x y z) ↔ ∀ x y z, q x y z := (Equiv.forall₃_congr eα.symm eβ.symm eγ.symm h.symm).symm +#align equiv.forall₃_congr' Equiv.forall₃_congr' protected theorem forall_congr_left' {p : α → Prop} (f : α ≃ β) : (∀ x, p x) ↔ ∀ y, p (f.symm y) := Equiv.forall_congr f <| by simp +#align equiv.forall_congr_left' Equiv.forall_congr_left' protected theorem forall_congr_left {p : β → Prop} (f : α ≃ β) : (∀ x, p (f x)) ↔ ∀ y, p y := (Equiv.forall_congr_left' f.symm).symm +#align equiv.forall_congr_left Equiv.forall_congr_left protected theorem exists_congr_left {α β} (f : α ≃ β) {p : α → Prop} : (∃ a, p a) ↔ ∃ b, p (f.symm b) := ⟨fun ⟨a, h⟩ => ⟨f a, by simpa using h⟩, fun ⟨b, h⟩ => ⟨_, h⟩⟩ +#align equiv.exists_congr_left Equiv.exists_congr_left end Equiv @@ -809,21 +933,25 @@ protected def congr {ra : α → α → Prop} {rb : β → β → Prop} (e : α ((e.apply_symm_apply b₁).symm ▸ (e.apply_symm_apply b₂).symm ▸ h) left_inv := by rintro ⟨a⟩; simp only [Quot.map, Equiv.symm_apply_apply] right_inv := by rintro ⟨a⟩; simp only [Quot.map, Equiv.apply_symm_apply] +#align quot.congr Quot.congr @[simp] theorem congr_mk {ra : α → α → Prop} {rb : β → β → Prop} (e : α ≃ β) (eq : ∀ a₁ a₂ : α, ra a₁ a₂ ↔ rb (e a₁) (e a₂)) (a : α) : Quot.congr e eq (Quot.mk ra a) = Quot.mk rb (e a) := rfl +#align quot.congr_mk Quot.congr_mk /-- Quotients are congruent on equivalences under equality of their relation. An alternative is just to use rewriting with `eq`, but then computational proofs get stuck. -/ protected def congrRight {r r' : α → α → Prop} (eq : ∀ a₁ a₂, r a₁ a₂ ↔ r' a₁ a₂) : Quot r ≃ Quot r' := Quot.congr (Equiv.refl α) eq +#align quot.congr_right Quot.congrRight /-- An equivalence `e : α ≃ β` generates an equivalence between the quotient space of `α` by a relation `ra` and the quotient space of `β` by the image of this relation under `e`. -/ protected def congrLeft {r : α → α → Prop} (e : α ≃ β) : Quot r ≃ Quot fun b b' => r (e.symm b) (e.symm b') := Quot.congr e fun _ _ => by simp only [e.symm_apply_apply] +#align quot.congr_left Quot.congrLeft end Quot @@ -834,15 +962,18 @@ if `ra a₁ a₂ ↔ rb (e a₁) (e a₂). -/ protected def congr {ra : Setoid α} {rb : Setoid β} (e : α ≃ β) (eq : ∀ a₁ a₂, @Setoid.r α ra a₁ a₂ ↔ @Setoid.r β rb (e a₁) (e a₂)) : Quotient ra ≃ Quotient rb := Quot.congr e eq +#align quotient.congr Quotient.congr @[simp] theorem congr_mk {ra : Setoid α} {rb : Setoid β} (e : α ≃ β) (eq : ∀ a₁ a₂ : α, Setoid.r a₁ a₂ ↔ Setoid.r (e a₁) (e a₂)) (a : α) : Quotient.congr e eq (Quotient.mk ra a) = Quotient.mk rb (e a) := rfl +#align quotient.congr_mk Quotient.congr_mk /-- Quotients are congruent on equivalences under equality of their relation. An alternative is just to use rewriting with `eq`, but then computational proofs get stuck. -/ protected def congrRight {r r' : Setoid α} (eq : ∀ a₁ a₂, @Setoid.r α r a₁ a₂ ↔ @Setoid.r α r' a₁ a₂) : Quotient r ≃ Quotient r' := Quot.congrRight eq +#align quotient.congr_right Quotient.congrRight end Quotient diff --git a/Mathlib/Logic/Equiv/Fin.lean b/Mathlib/Logic/Equiv/Fin.lean index 6d924c01cf54a..0e3d3ec1a1a38 100644 --- a/Mathlib/Logic/Equiv/Fin.lean +++ b/Mathlib/Logic/Equiv/Fin.lean @@ -52,6 +52,8 @@ def piFinTwoEquiv (α : Fin 2 → Type u) : (∀ i, α i) ≃ α 0 × α 1 left_inv _ := funext <| Fin.forall_fin_two.2 ⟨rfl, rfl⟩ right_inv := fun _ => rfl #align pi_fin_two_equiv piFinTwoEquiv +#align pi_fin_two_equiv_symm_apply piFinTwoEquiv_symm_apply +#align pi_fin_two_equiv_apply piFinTwoEquiv_apply /- ./././Mathport/Syntax/Translate/Expr.lean:177:8: unsupported: ambiguous notation -/ theorem Fin.preimage_apply_01_prod {α : Fin 2 → Type u} (s : Set (α 0)) (t : Set (α 1)) : @@ -75,6 +77,8 @@ theorem Fin.preimage_apply_01_prod' {α : Type u} (s t : Set α) : def prodEquivPiFinTwo (α β : Type u) : α × β ≃ ∀ i : Fin 2, ![α, β] i := (piFinTwoEquiv (Fin.cons α (Fin.cons β finZeroElim))).symm #align prod_equiv_pi_fin_two prodEquivPiFinTwo +#align prod_equiv_pi_fin_two_apply prodEquivPiFinTwo_apply +#align prod_equiv_pi_fin_two_symm_apply prodEquivPiFinTwo_symm_apply /-- The space of functions `Fin 2 → α` is equivalent to `α × α`. See also `piFinTwoEquiv` and `prodEquivPiFinTwo`. -/ @@ -82,6 +86,8 @@ def prodEquivPiFinTwo (α β : Type u) : α × β ≃ ∀ i : Fin 2, ![α, β] i def finTwoArrowEquiv (α : Type _) : (Fin 2 → α) ≃ α × α := { piFinTwoEquiv fun _ => α with invFun := fun x => ![x.1, x.2] } #align fin_two_arrow_equiv finTwoArrowEquiv +#align fin_two_arrow_equiv_symm_apply finTwoArrowEquiv_symm_apply +#align fin_two_arrow_equiv_apply finTwoArrowEquiv_apply /-- `Π i : Fin 2, α i` is order equivalent to `α 0 × α 1`. See also `OrderIso.finTwoArrowEquiv` for a non-dependent version. -/ @@ -291,6 +297,8 @@ def Equiv.piFinSuccAboveEquiv (α : Fin (n + 1) → Type u) (i : Fin (n + 1)) : left_inv f := by simp [Fin.insertNth_eq_iff] right_inv f := by simp #align equiv.pi_fin_succ_above_equiv Equiv.piFinSuccAboveEquiv +#align equiv.pi_fin_succ_above_equiv_apply Equiv.piFinSuccAboveEquiv_apply +#align equiv.pi_fin_succ_above_equiv_symm_apply Equiv.piFinSuccAboveEquiv_symm_apply /-- Order isomorphism between `Π j : Fin (n + 1), α j` and `α i × Π j : Fin n, α (Fin.succAbove i j)`. -/ @@ -305,6 +313,8 @@ def OrderIso.piFinSuccAboveIso (α : Fin (n + 1) → Type u) [∀ i, LE (α i)] def Equiv.piFinSucc (n : ℕ) (β : Type u) : (Fin (n + 1) → β) ≃ β × (Fin n → β) := Equiv.piFinSuccAboveEquiv (fun _ => β) 0 #align equiv.pi_fin_succ Equiv.piFinSucc +#align equiv.pi_fin_succ_apply Equiv.piFinSucc_apply +#align equiv.pi_fin_succ_symm_apply Equiv.piFinSucc_symm_apply /-- Equivalence between `Fin m ⊕ Fin n` and `Fin (m + n)` -/ def finSumFinEquiv : Sum (Fin m) (Fin n) ≃ Fin (m + n) @@ -478,6 +488,8 @@ def finProdFinEquiv : Fin m × Fin n ≃ Fin (m * n) ) right_inv x := Fin.eq_of_veq <| Nat.mod_add_div _ _ #align fin_prod_fin_equiv finProdFinEquiv +#align fin_prod_fin_equiv_apply_val finProdFinEquiv_apply_val +#align fin_prod_fin_equiv_symm_apply finProdFinEquiv_symm_apply /-- Promote a `Fin n` into a larger `Fin m`, as a subtype where the underlying values are retained. This is the `OrderIso` version of `Fin.castLe`. -/ @@ -490,6 +502,8 @@ def Fin.castLeOrderIso {n m : ℕ} (h : n ≤ m) : Fin n ≃o { i : Fin m // (i right_inv _ := by simp map_rel_iff' := by simp #align fin.cast_le_order_iso Fin.castLeOrderIso +#align fin.cast_le_order_iso_apply Fin.castLeOrderIso_apply +#align fin.cast_le_order_iso_symm_apply Fin.castLeOrderIso_symmApply /-- `Fin 0` is a subsingleton. -/ instance subsingleton_fin_zero : Subsingleton (Fin 0) := diff --git a/Mathlib/Logic/Equiv/LocalEquiv.lean b/Mathlib/Logic/Equiv/LocalEquiv.lean index 806163ae68334..4e67ed9f60ee8 100644 --- a/Mathlib/Logic/Equiv/LocalEquiv.lean +++ b/Mathlib/Logic/Equiv/LocalEquiv.lean @@ -270,6 +270,10 @@ def _root_.Equiv.toLocalEquiv (e : α ≃ β) : LocalEquiv α β where left_inv' x _ := e.left_inv x right_inv' x _ := e.right_inv x #align equiv.to_local_equiv Equiv.toLocalEquiv +#align equiv.to_local_equiv_symm_apply Equiv.toLocalEquiv_symmApply +#align equiv.to_local_equiv_target Equiv.toLocalEquiv_target +#align equiv.to_local_equiv_apply Equiv.toLocalEquiv_apply +#align equiv.to_local_equiv_source Equiv.toLocalEquiv_source instance inhabitedOfEmpty [IsEmpty α] [IsEmpty β] : Inhabited (LocalEquiv α β) := ⟨((Equiv.equivEmpty α).trans (Equiv.equivEmpty β).symm).toLocalEquiv⟩ @@ -289,6 +293,10 @@ def copy (e : LocalEquiv α β) (f : α → β) (hf : ⇑e = f) (g : β → α) left_inv' _ := hs ▸ hf ▸ hg ▸ e.left_inv right_inv' _ := ht ▸ hf ▸ hg ▸ e.right_inv #align local_equiv.copy LocalEquiv.copy +#align local_equiv.copy_source LocalEquiv.copy_source +#align local_equiv.copy_apply LocalEquiv.copy_apply +#align local_equiv.copy_symm_apply LocalEquiv.copy_symmApply +#align local_equiv.copy_target LocalEquiv.copy_target theorem copy_eq (e : LocalEquiv α β) (f : α → β) (hf : ⇑e = f) (g : β → α) (hg : ⇑e.symm = g) (s : Set α) (hs : e.source = s) (t : Set β) (ht : e.target = t) : @@ -386,6 +394,10 @@ def restr (h : e.IsImage s t) : LocalEquiv α β where left_inv' := e.leftInvOn.mono (inter_subset_left _ _) right_inv' := e.rightInvOn.mono (inter_subset_left _ _) #align local_equiv.is_image.restr LocalEquiv.IsImage.restr +#align local_equiv.is_image.restr_apply LocalEquiv.IsImage.restr_apply +#align local_equiv.is_image.restr_source LocalEquiv.IsImage.restr_source +#align local_equiv.is_image.restr_target LocalEquiv.IsImage.restr_target +#align local_equiv.is_image.restr_symm_apply LocalEquiv.IsImage.restr_symmApply theorem image_eq (h : e.IsImage s t) : e '' (e.source ∩ s) = e.target ∩ t := h.restr.image_source_eq_target @@ -400,12 +412,16 @@ theorem iff_preimage_eq : e.IsImage s t ↔ e.source ∩ e ⁻¹' t = e.source #align local_equiv.is_image.iff_preimage_eq LocalEquiv.IsImage.iff_preimage_eq alias iff_preimage_eq ↔ preimage_eq of_preimage_eq +#align local_equiv.is_image.of_preimage_eq LocalEquiv.IsImage.of_preimage_eq +#align local_equiv.is_image.preimage_eq LocalEquiv.IsImage.preimage_eq theorem iff_symm_preimage_eq : e.IsImage s t ↔ e.target ∩ e.symm ⁻¹' s = e.target ∩ t := symm_iff.symm.trans iff_preimage_eq #align local_equiv.is_image.iff_symm_preimage_eq LocalEquiv.IsImage.iff_symm_preimage_eq alias iff_symm_preimage_eq ↔ symm_preimage_eq of_symm_preimage_eq +#align local_equiv.is_image.of_symm_preimage_eq LocalEquiv.IsImage.of_symm_preimage_eq +#align local_equiv.is_image.symm_preimage_eq LocalEquiv.IsImage.symm_preimage_eq theorem of_image_eq (h : e '' (e.source ∩ s) = e.target ∩ t) : e.IsImage s t := of_symm_preimage_eq <| Eq.trans (of_symm_preimage_eq rfl).image_eq.symm h @@ -765,6 +781,10 @@ def transEquiv (e' : β ≃ γ) : LocalEquiv α γ := (e.trans e'.toLocalEquiv).copy _ rfl _ rfl e.source (inter_univ _) (e'.symm ⁻¹' e.target) (univ_inter _) #align local_equiv.trans_equiv LocalEquiv.transEquiv +#align local_equiv.trans_equiv_source LocalEquiv.transEquiv_source +#align local_equiv.trans_equiv_apply LocalEquiv.transEquiv_apply +#align local_equiv.trans_equiv_target LocalEquiv.transEquiv_target +#align local_equiv.trans_equiv_symm_apply LocalEquiv.transEquiv_symmApply theorem transEquiv_eq_trans (e' : β ≃ γ) : e.transEquiv e' = e.trans e'.toLocalEquiv := copy_eq .. @@ -777,6 +797,10 @@ def _root_.Equiv.transLocalEquiv (e : α ≃ β) : LocalEquiv α γ := (e.toLocalEquiv.trans e').copy _ rfl _ rfl (e ⁻¹' e'.source) (univ_inter _) e'.target (inter_univ _) #align equiv.trans_local_equiv Equiv.transLocalEquiv +#align equiv.trans_local_equiv_target Equiv.transLocalEquiv_target +#align equiv.trans_local_equiv_apply Equiv.transLocalEquiv_apply +#align equiv.trans_local_equiv_source Equiv.transLocalEquiv_source +#align equiv.trans_local_equiv_symm_apply Equiv.transLocalEquiv_symmApply theorem _root_.Equiv.transLocalEquiv_eq_trans (e : α ≃ β) : e.transLocalEquiv e' = e.toLocalEquiv.trans e' := @@ -969,6 +993,10 @@ def piecewise (e e' : LocalEquiv α β) (s : Set α) (t : Set β) [∀ x, Decida left_inv' := H.leftInvOn_piecewise H' right_inv' := H.symm.leftInvOn_piecewise H'.symm #align local_equiv.piecewise LocalEquiv.piecewise +#align local_equiv.piecewise_source LocalEquiv.piecewise_source +#align local_equiv.piecewise_target LocalEquiv.piecewise_target +#align local_equiv.piecewise_symm_apply LocalEquiv.piecewise_symmApply +#align local_equiv.piecewise_apply LocalEquiv.piecewise_apply theorem symm_piecewise (e e' : LocalEquiv α β) {s : Set α} {t : Set β} [∀ x, Decidable (x ∈ s)] [∀ y, Decidable (y ∈ t)] (H : e.IsImage s t) (H' : e'.IsImage s t) : @@ -987,6 +1015,10 @@ def disjointUnion (e e' : LocalEquiv α β) (hs : Disjoint e.source e'.source) e'.isImage_source_target_of_disjoint _ hs.symm ht.symm).copy _ rfl _ rfl (e.source ∪ e'.source) (ite_left _ _) (e.target ∪ e'.target) (ite_left _ _) #align local_equiv.disjoint_union LocalEquiv.disjointUnion +#align local_equiv.disjoint_union_source LocalEquiv.disjointUnion_source +#align local_equiv.disjoint_union_target LocalEquiv.disjointUnion_target +#align local_equiv.disjoint_union_symm_apply LocalEquiv.disjointUnion_symmApply +#align local_equiv.disjoint_union_apply LocalEquiv.disjointUnion_apply theorem disjointUnion_eq_piecewise (e e' : LocalEquiv α β) (hs : Disjoint e.source e'.source) (ht : Disjoint e.target e'.target) [∀ x, Decidable (x ∈ e.source)] @@ -1013,6 +1045,10 @@ protected def pi : LocalEquiv (∀ i, αi i) (∀ i, βi i) where left_inv' _ hf := funext fun i => (ei i).left_inv (hf i trivial) right_inv' _ hf := funext fun i => (ei i).right_inv (hf i trivial) #align local_equiv.pi LocalEquiv.pi +#align local_equiv.pi_symm_apply LocalEquiv.pi_symmApply +#align local_equiv.pi_source LocalEquiv.pi_source +#align local_equiv.pi_apply LocalEquiv.pi_apply +#align local_equiv.pi_target LocalEquiv.pi_target end Pi @@ -1035,6 +1071,10 @@ noncomputable def BijOn.toLocalEquiv [Nonempty α] (f : α → β) (s : Set α) left_inv' := hf.invOn_invFunOn.1 right_inv' := hf.invOn_invFunOn.2 #align set.bij_on.to_local_equiv Set.BijOn.toLocalEquiv +#align set.bij_on.to_local_equiv_target Set.BijOn.toLocalEquiv_target +#align set.bij_on.to_local_equiv_symm_apply Set.BijOn.toLocalEquiv_symmApply +#align set.bij_on.to_local_equiv_apply Set.BijOn.toLocalEquiv_apply +#align set.bij_on.to_local_equiv_source Set.BijOn.toLocalEquiv_source /-- A map injective on a subset of its domain provides a local equivalence. -/ @[simp, mfld_simps] diff --git a/Mathlib/Logic/Equiv/Nat.lean b/Mathlib/Logic/Equiv/Nat.lean index 7385837c696a5..511c15175b6ad 100644 --- a/Mathlib/Logic/Equiv/Nat.lean +++ b/Mathlib/Logic/Equiv/Nat.lean @@ -34,6 +34,8 @@ def boolProdNatEquivNat : Bool × ℕ ≃ ℕ where left_inv := fun ⟨b, n⟩ => by simp only [bodd_bit, div2_bit, uncurry_apply_pair, boddDiv2_eq] right_inv n := by simp only [bit_decomp, boddDiv2_eq, uncurry_apply_pair] #align equiv.bool_prod_nat_equiv_nat Equiv.boolProdNatEquivNat +#align equiv.bool_prod_nat_equiv_nat_symm_apply Equiv.boolProdNatEquivNat_symm_apply +#align equiv.bool_prod_nat_equiv_nat_apply Equiv.boolProdNatEquivNat_apply /-- An equivalence between `ℕ ⊕ ℕ` and `ℕ`, by mapping `(Sum.inl x)` to `2 * x` and `(Sum.inr x)` to `2 * x + 1`. @@ -42,6 +44,7 @@ def boolProdNatEquivNat : Bool × ℕ ≃ ℕ where def natSumNatEquivNat : ℕ ⊕ ℕ ≃ ℕ := (boolProdEquivSum ℕ).symm.trans boolProdNatEquivNat #align equiv.nat_sum_nat_equiv_nat Equiv.natSumNatEquivNat +#align equiv.nat_sum_nat_equiv_nat_symm_apply Equiv.natSumNatEquivNat_symm_apply set_option linter.deprecated false in @[simp] diff --git a/Mathlib/Logic/Equiv/Option.lean b/Mathlib/Logic/Equiv/Option.lean index 9eeee9ae478b6..8504934c5bce4 100644 --- a/Mathlib/Logic/Equiv/Option.lean +++ b/Mathlib/Logic/Equiv/Option.lean @@ -41,6 +41,7 @@ def optionCongr (e : α ≃ β) : Option α ≃ Option β where left_inv x := (Option.map_map _ _ _).trans <| e.symm_comp_self.symm ▸ congr_fun Option.map_id x right_inv x := (Option.map_map _ _ _).trans <| e.self_comp_symm.symm ▸ congr_fun Option.map_id x #align equiv.option_congr Equiv.optionCongr +#align equiv.option_congr_apply Equiv.optionCongr_apply @[simp] theorem optionCongr_refl : optionCongr (Equiv.refl α) = Equiv.refl _ := diff --git a/Mathlib/Logic/Equiv/Set.lean b/Mathlib/Logic/Equiv/Set.lean index 205a4beeb4715..5fcfbb96c5e82 100644 --- a/Mathlib/Logic/Equiv/Set.lean +++ b/Mathlib/Logic/Equiv/Set.lean @@ -177,6 +177,7 @@ def setProdEquivSigma {α β : Type _} (s : Set (α × β)) : def setCongr {α : Type _} {s t : Set α} (h : s = t) : s ≃ t := subtypeEquivProp h #align equiv.set_congr Equiv.setCongr +#align equiv.set_congr_apply Equiv.setCongr_apply -- We could construct this using `Equiv.Set.image e s e.injective`, -- but this definition provides an explicit inverse. @@ -193,6 +194,8 @@ def image {α β : Type _} (e : α ≃ β) (s : Set α) : left_inv x := by simp right_inv y := by simp #align equiv.image Equiv.image +#align equiv.image_symm_apply_coe Equiv.image_symm_apply_coe +#align equiv.image_apply_coe Equiv.image_apply_coe namespace Set @@ -471,6 +474,8 @@ protected def univPi {α : Type _} {β : α → Type _} (s : ∀ a, Set (β a)) ext a rfl #align equiv.set.univ_pi Equiv.Set.univPi +#align equiv.set.univ_pi_symm_apply_coe Equiv.Set.univPi_symm_apply_coe +#align equiv.set.univ_pi_apply_coe Equiv.Set.univPi_apply_coe /-- If a function `f` is injective on a set `s`, then `s` is equivalent to `f '' s`. -/ protected noncomputable def imageOfInjOn {α β} (f : α → β) (s : Set α) (H : InjOn f s) : @@ -488,6 +493,7 @@ protected noncomputable def imageOfInjOn {α β} (f : α → β) (s : Set α) (H protected noncomputable def image {α β} (f : α → β) (s : Set α) (H : Injective f) : s ≃ f '' s := Equiv.Set.imageOfInjOn f s (H.injOn s) #align equiv.set.image Equiv.Set.image +#align equiv.set.image_apply Equiv.Set.image_apply @[simp] protected theorem image_symm_apply {α β} (f : α → β) (s : Set α) (H : Injective f) (x : α) @@ -508,6 +514,8 @@ theorem image_symm_preimage {α β} {f : α → β} (hf : Injective f) (u s : Se protected def congr {α β : Type _} (e : α ≃ β) : Set α ≃ Set β := ⟨fun s => e '' s, fun t => e.symm '' t, symm_image_image e, symm_image_image e.symm⟩ #align equiv.set.congr Equiv.Set.congr +#align equiv.set.congr_apply Equiv.Set.congr_apply +#align equiv.set.congr_symm_apply Equiv.Set.congr_symm_apply /-- The set `{x ∈ s | t x}` is equivalent to the set of `x : s` such that `t x`. -/ protected def sep {α : Type u} (s : Set α) (t : α → Prop) : @@ -540,6 +548,8 @@ noncomputable def rangeSplittingImageEquiv {α β : Type _} (f : α → β) (s : simp [apply_rangeSplitting f] right_inv x := by simp [apply_rangeSplitting f] #align equiv.set.range_splitting_image_equiv Equiv.Set.rangeSplittingImageEquiv +#align equiv.set.range_splitting_image_equiv_symm_apply_coe Equiv.Set.rangeSplittingImageEquiv_symm_apply_coe +#align equiv.set.range_splitting_image_equiv_apply_coe_coe Equiv.Set.rangeSplittingImageEquiv_apply_coe_coe end Set @@ -560,6 +570,8 @@ def ofLeftInverse {α β : Sort _} (f : α → β) (f_inv : Nonempty α → β right_inv := fun ⟨b, a, ha⟩ => Subtype.eq <| show f (f_inv ⟨a⟩ b) = b from Eq.trans (congr_arg f <| ha ▸ hf _ a) ha #align equiv.of_left_inverse Equiv.ofLeftInverse +#align equiv.of_left_inverse_apply_coe Equiv.ofLeftInverse_apply_coe +#align equiv.of_left_inverse_symm_apply Equiv.ofLeftInverse_symm_apply /-- If `f : α → β` has a left-inverse, then `α` is computably equivalent to the range of `f`. @@ -575,6 +587,7 @@ abbrev ofLeftInverse' {α β : Sort _} (f : α → β) (f_inv : β → α) (hf : noncomputable def ofInjective {α β} (f : α → β) (hf : Injective f) : α ≃ range f := Equiv.ofLeftInverse f (fun _ => Function.invFun f) fun _ => Function.leftInverse_invFun hf #align equiv.of_injective Equiv.ofInjective +#align equiv.of_injective_apply Equiv.ofInjective_apply theorem apply_ofInjective_symm {α β} {f : α → β} (hf : Injective f) (b : range f) : f ((ofInjective f hf).symm b) = b := @@ -644,6 +657,9 @@ the type of all preimages of points under `f` and the total space `α`. -/ def sigmaPreimageEquiv {α β} (f : α → β) : (Σb, f ⁻¹' {b}) ≃ α := sigmaFiberEquiv f #align equiv.sigma_preimage_equiv Equiv.sigmaPreimageEquiv +#align equiv.sigma_preimage_equiv_symm_apply_snd_coe Equiv.sigmaPreimageEquiv_symm_apply_snd_coe +#align equiv.sigma_preimage_equiv_apply Equiv.sigmaPreimageEquiv_apply +#align equiv.sigma_preimage_equiv_symm_apply_fst Equiv.sigmaPreimageEquiv_symm_apply_fst -- See also `Equiv.ofFiberEquiv`. /-- A family of equivalences between preimages of points gives an equivalence between domains. -/ @@ -651,6 +667,8 @@ def sigmaPreimageEquiv {α β} (f : α → β) : (Σb, f ⁻¹' {b}) ≃ α := def ofPreimageEquiv {α β γ} {f : α → γ} {g : β → γ} (e : ∀ c, f ⁻¹' {c} ≃ g ⁻¹' {c}) : α ≃ β := Equiv.ofFiberEquiv e #align equiv.of_preimage_equiv Equiv.ofPreimageEquiv +#align equiv.of_preimage_equiv_apply Equiv.ofPreimageEquiv_apply +#align equiv.of_preimage_equiv_symm_apply Equiv.ofPreimageEquiv_symm_apply theorem ofPreimageEquiv_map {α β γ} {f : α → γ} {g : β → γ} (e : ∀ c, f ⁻¹' {c} ≃ g ⁻¹' {c}) (a : α) : g (ofPreimageEquiv e a) = f a := diff --git a/Mathlib/Logic/Function/Basic.lean b/Mathlib/Logic/Function/Basic.lean index cb7059f91103d..35c504d29652f 100644 --- a/Mathlib/Logic/Function/Basic.lean +++ b/Mathlib/Logic/Function/Basic.lean @@ -29,31 +29,39 @@ variable {α β γ : Sort _} {f : α → β} /-- Evaluate a function at an argument. Useful if you want to talk about the partially applied `Function.eval x : (∀ x, β x) → β x`. -/ @[reducible, simp] def eval {β : α → Sort _} (x : α) (f : ∀ x, β x) : β x := f x +#align function.eval Function.eval theorem eval_apply {β : α → Sort _} (x : α) (f : ∀ x, β x) : eval x f = f x := rfl +#align function.eval_apply Function.eval_apply theorem const_def {y : β} : (fun _ : α ↦ y) = const α y := rfl +#align function.const_def Function.const_def @[simp] theorem const_comp {f : α → β} {c : γ} : const β c ∘ f = const α c := rfl +#align function.const_comp Function.const_comp @[simp] theorem comp_const {f : β → γ} {b : β} : f ∘ const α b = const α (f b) := rfl +#align function.comp_const Function.comp_const theorem const_injective [Nonempty α] : Injective (const α : β → α → β) := fun y₁ y₂ h ↦ let ⟨x⟩ := ‹Nonempty α› congr_fun h x +#align function.const_injective Function.const_injective @[simp] theorem const_inj [Nonempty α] {y₁ y₂ : β} : const α y₁ = const α y₂ ↔ y₁ = y₂ := ⟨fun h ↦ const_injective h, fun h ↦ h ▸ rfl⟩ +#align function.const_inj Function.const_inj theorem id_def : @id α = fun x ↦ x := rfl +#align function.id_def Function.id_def lemma hfunext {α α': Sort u} {β : α → Sort v} {β' : α' → Sort v} {f : ∀a, β a} {f' : ∀a, β' a} (hα : α = α') (h : ∀a a', HEq a a' → HEq (f a) (f' a')) : HEq f f' := by @@ -65,12 +73,15 @@ lemma hfunext {α α': Sort u} {β : α → Sort v} {β' : α' → Sort v} {f : apply heq_of_eq funext a exact eq_of_heq (this a) +#align function.hfunext Function.hfunext theorem funext_iff {β : α → Sort _} {f₁ f₂ : ∀ x : α, β x} : f₁ = f₂ ↔ ∀ a, f₁ a = f₂ a := Iff.intro (fun h _ ↦ h ▸ rfl) funext +#align function.funext_iff Function.funext_iff theorem ne_iff {β : α → Sort _} {f₁ f₂ : ∀ a, β a} : f₁ ≠ f₂ ↔ ∃ a, f₁ a ≠ f₂ a := funext_iff.not.trans not_forall +#align function.ne_iff Function.ne_iff protected theorem Bijective.injective {f : α → β} (hf : Bijective f) : Injective f := hf.1 #align function.bijective.injective Function.Bijective.injective @@ -79,6 +90,7 @@ protected theorem Bijective.surjective {f : α → β} (hf : Bijective f) : Surj theorem Injective.eq_iff (I : Injective f) {a b : α} : f a = f b ↔ a = b := ⟨@I _ _, congr_arg f⟩ +#align function.injective.eq_iff Function.Injective.eq_iff theorem Injective.beq_eq [BEq α] [LawfulBEq α] [BEq β] [LawfulBEq β] (I : Injective f) {a b : α} : (f a == f b) = (a == b) := by @@ -86,6 +98,7 @@ theorem Injective.beq_eq [BEq α] [LawfulBEq α] [BEq β] [LawfulBEq β] theorem Injective.eq_iff' (I : Injective f) {a b : α} {c : β} (h : f b = c) : f a = c ↔ a = b := h ▸ I.eq_iff +#align function.injective.eq_iff' Function.Injective.eq_iff' theorem Injective.ne (hf : Injective f) {a₁ a₂ : α} : a₁ ≠ a₂ → f a₁ ≠ f a₂ := mt fun h ↦ hf h @@ -93,21 +106,26 @@ theorem Injective.ne (hf : Injective f) {a₁ a₂ : α} : a₁ ≠ a₂ → f a theorem Injective.ne_iff (hf : Injective f) {x y : α} : f x ≠ f y ↔ x ≠ y := ⟨mt <| congr_arg f, hf.ne⟩ +#align function.injective.ne_iff Function.Injective.ne_iff theorem Injective.ne_iff' (hf : Injective f) {x y : α} {z : β} (h : f y = z) : f x ≠ z ↔ x ≠ y := h ▸ hf.ne_iff +#align function.injective.ne_iff' Function.Injective.ne_iff' /-- If the co-domain `β` of an injective function `f : α → β` has decidable equality, then the domain `α` also has decidable equality. -/ protected def Injective.decidableEq [DecidableEq β] (I : Injective f) : DecidableEq α := fun _ _ ↦ decidable_of_iff _ I.eq_iff +#align function.injective.decidable_eq Function.Injective.decidableEq theorem Injective.of_comp {g : γ → α} (I : Injective (f ∘ g)) : Injective g := fun x y h ↦ I <| show f (g x) = f (g y) from congr_arg f h +#align function.injective.of_comp Function.Injective.of_comp theorem Injective.of_comp_iff {f : α → β} (hf : Injective f) (g : γ → α) : Injective (f ∘ g) ↔ Injective g := ⟨Injective.of_comp, hf.comp⟩ +#align function.injective.of_comp_iff Function.Injective.of_comp_iff theorem Injective.of_comp_iff' (f : α → β) {g : γ → α} (hg : Bijective g) : Injective (f ∘ g) ↔ Injective f := @@ -115,14 +133,17 @@ theorem Injective.of_comp_iff' (f : α → β) {g : γ → α} (hg : Bijective g let ⟨_, hy⟩ := hg.surjective y hx ▸ hy ▸ λ hf => h hf ▸ rfl, λ h => h.comp hg.injective⟩ +#align function.injective.of_comp_iff' Function.Injective.of_comp_iff' /-- Composition by an injective function on the left is itself injective. -/ theorem Injective.comp_left {g : β → γ} (hg : Function.Injective g) : Function.Injective ((· ∘ ·) g : (α → β) → α → γ) := fun _ _ hgf ↦ funext fun i ↦ hg <| (congr_fun hgf i : _) +#align function.injective.comp_left Function.Injective.comp_left theorem injective_of_subsingleton [Subsingleton α] (f : α → β) : Injective f := fun _ _ _ ↦ Subsingleton.elim _ _ +#align function.injective_of_subsingleton Function.injective_of_subsingleton lemma Injective.dite (p : α → Prop) [DecidablePred p] {f : {a : α // p a} → β} {f' : {a : α // ¬ p a} → β} @@ -136,14 +157,17 @@ by intros x₁ x₂ h · rw [dif_pos h₁, dif_neg h₂] at h; exact (im_disj h).elim · rw [dif_neg h₁, dif_pos h₂] at h; exact (im_disj h.symm).elim · rw [dif_neg h₁, dif_neg h₂] at h; injection (hf' h) +#align function.injective.dite Function.Injective.dite theorem Surjective.of_comp {g : γ → α} (S : Surjective (f ∘ g)) : Surjective f := fun y ↦ let ⟨x, h⟩ := S y ⟨g x, h⟩ +#align function.surjective.of_comp Function.Surjective.of_comp theorem Surjective.of_comp_iff (f : α → β) {g : γ → α} (hg : Surjective g) : Surjective (f ∘ g) ↔ Surjective f := ⟨Surjective.of_comp, fun h ↦ h.comp hg⟩ +#align function.surjective.of_comp_iff Function.Surjective.of_comp_iff theorem Surjective.of_comp_iff' (hf : Bijective f) (g : γ → α) : Surjective (f ∘ g) ↔ Surjective g := @@ -151,6 +175,7 @@ theorem Surjective.of_comp_iff' (hf : Bijective f) (g : γ → α) : let ⟨x', hx'⟩ := h (f x) ⟨x', hf.injective hx'⟩, hf.surjective.comp⟩ +#align function.surjective.of_comp_iff' Function.Surjective.of_comp_iff' instance decidableEqPfun (p : Prop) [Decidable p] (α : p → Type _) [∀ hp, DecidableEq (α hp)] : DecidableEq (∀ hp, α hp) @@ -161,14 +186,17 @@ protected theorem Surjective.forall (hf : Surjective f) {p : β → Prop} : ⟨fun h x ↦ h (f x), fun h y ↦ let ⟨x, hx⟩ := hf y hx ▸ h x⟩ +#align function.surjective.forall Function.Surjective.forall protected theorem Surjective.forall₂ (hf : Surjective f) {p : β → β → Prop} : (∀ y₁ y₂, p y₁ y₂) ↔ ∀ x₁ x₂, p (f x₁) (f x₂) := hf.forall.trans $ forall_congr' fun _ ↦ hf.forall +#align function.surjective.forall₂ Function.Surjective.forall₂ protected theorem Surjective.forall₃ (hf : Surjective f) {p : β → β → β → Prop} : (∀ y₁ y₂ y₃, p y₁ y₂ y₃) ↔ ∀ x₁ x₂ x₃, p (f x₁) (f x₂) (f x₃) := hf.forall.trans $ forall_congr' fun _ ↦ hf.forall₂ +#align function.surjective.forall₃ Function.Surjective.forall₃ protected theorem Surjective.exists (hf : Surjective f) {p : β → Prop} : (∃ y, p y) ↔ ∃ x, p (f x) := @@ -176,26 +204,32 @@ protected theorem Surjective.exists (hf : Surjective f) {p : β → Prop} : let ⟨x, hx⟩ := hf y ⟨x, hx.symm ▸ hy⟩, fun ⟨x, hx⟩ ↦ ⟨f x, hx⟩⟩ +#align function.surjective.exists Function.Surjective.exists protected theorem Surjective.exists₂ (hf : Surjective f) {p : β → β → Prop} : (∃ y₁ y₂, p y₁ y₂) ↔ ∃ x₁ x₂, p (f x₁) (f x₂) := hf.exists.trans <| exists_congr fun _ ↦ hf.exists +#align function.surjective.exists₂ Function.Surjective.exists₂ protected theorem Surjective.exists₃ (hf : Surjective f) {p : β → β → β → Prop} : (∃ y₁ y₂ y₃, p y₁ y₂ y₃) ↔ ∃ x₁ x₂ x₃, p (f x₁) (f x₂) (f x₃) := hf.exists.trans <| exists_congr fun _ ↦ hf.exists₂ +#align function.surjective.exists₃ Function.Surjective.exists₃ theorem Surjective.injective_comp_right (hf : Surjective f) : Injective fun g : β → γ ↦ g ∘ f := fun _ _ h ↦ funext <| hf.forall.2 <| congr_fun h +#align function.surjective.injective_comp_right Function.Surjective.injective_comp_right protected theorem Surjective.right_cancellable (hf : Surjective f) {g₁ g₂ : β → γ} : g₁ ∘ f = g₂ ∘ f ↔ g₁ = g₂ := hf.injective_comp_right.eq_iff +#align function.surjective.right_cancellable Function.Surjective.right_cancellable theorem surjective_of_right_cancellable_Prop (h : ∀ g₁ g₂ : β → Prop, g₁ ∘ f = g₂ ∘ f → g₁ = g₂) : Surjective f := by specialize h (fun y ↦ ∃ x, f x = y) (fun _ ↦ True) (funext fun x ↦ eq_true ⟨_, rfl⟩) intro y; rw [congr_fun h y]; trivial +#align function.surjective_of_right_cancellable_Prop Function.surjective_of_right_cancellable_Prop theorem bijective_iff_existsUnique (f : α → β) : Bijective f ↔ ∀ b : β, ∃! a : α, f a = b := ⟨fun hf b ↦ @@ -224,16 +258,19 @@ theorem Bijective.existsUnique_iff {f : α → β} (hf : Bijective f) {p : β theorem Bijective.of_comp_iff (f : α → β) {g : γ → α} (hg : Bijective g) : Bijective (f ∘ g) ↔ Bijective f := and_congr (Injective.of_comp_iff' _ hg) (Surjective.of_comp_iff _ hg.surjective) +#align function.bijective.of_comp_iff Function.Bijective.of_comp_iff theorem Bijective.of_comp_iff' {f : α → β} (hf : Bijective f) (g : γ → α) : Function.Bijective (f ∘ g) ↔ Function.Bijective g := and_congr (Injective.of_comp_iff hf.injective _) (Surjective.of_comp_iff' hf _) +#align function.bijective.of_comp_iff' Function.Bijective.of_comp_iff' /-- **Cantor's diagonal argument** implies that there are no surjective functions from `α` to `Set α`. -/ theorem cantor_surjective {α} (f : α → Set α) : ¬Surjective f | h => let ⟨D, e⟩ := h (λ a => ¬ f a a) (@iff_not_self (f D D)) $ iff_of_eq (congr_fun e D) +#align function.cantor_surjective Function.cantor_surjective /-- **Cantor's diagonal argument** implies that there are no injective functions from `Set α` to `α`. -/ @@ -241,6 +278,7 @@ theorem cantor_injective {α : Type _} (f : Set α → α) : ¬Injective f | i => cantor_surjective (λ a b => ∀ U, a = f U → U b) $ RightInverse.surjective (λ U => funext $ λ _a => propext ⟨λ h => h U rfl, λ h' _U e => i e ▸ h'⟩) +#align function.cantor_injective Function.cantor_injective /-- There is no surjection from `α : Type u` into `Type u`. This theorem demonstrates why `Type : Type` would be inconsistent in Lean. -/ @@ -256,12 +294,14 @@ theorem not_surjective_Type {α : Type u} (f : α → Type max u v) : ¬Surjecti assumption · congr exact cantor_injective g hg +#align function.not_surjective_Type Function.not_surjective_Type /-- `g` is a partial inverse to `f` (an injective but not necessarily surjective function) if `g y = some x` implies `f x = y`, and `g y = none` implies that `y` is not in the range of `f`. -/ def IsPartialInv {α β} (f : α → β) (g : β → Option α) : Prop := ∀ x y, g y = some x ↔ f x = y +#align function.is_partial_inv Function.IsPartialInv theorem isPartialInv_left {α β} {f : α → β} {g} (H : IsPartialInv f g) (x) : g (f x) = some x := (H _ _).2 rfl @@ -279,6 +319,7 @@ theorem injective_of_isPartialInv_right {α β} {f : α → β} {g} (H : IsParti theorem LeftInverse.comp_eq_id {f : α → β} {g : β → α} (h : LeftInverse f g) : f ∘ g = id := funext h +#align function.left_inverse.comp_eq_id Function.LeftInverse.comp_eq_id theorem leftInverse_iff_comp {f : α → β} {g : β → α} : LeftInverse f g ↔ f ∘ g = id := ⟨LeftInverse.comp_eq_id, congr_fun⟩ @@ -286,6 +327,7 @@ theorem leftInverse_iff_comp {f : α → β} {g : β → α} : LeftInverse f g theorem RightInverse.comp_eq_id {f : α → β} {g : β → α} (h : RightInverse f g) : g ∘ f = id := funext h +#align function.right_inverse.comp_eq_id Function.RightInverse.comp_eq_id theorem rightInverse_iff_comp {f : α → β} {g : β → α} : RightInverse f g ↔ g ∘ f = id := ⟨RightInverse.comp_eq_id, congr_fun⟩ @@ -294,10 +336,12 @@ theorem rightInverse_iff_comp {f : α → β} {g : β → α} : RightInverse f g theorem LeftInverse.comp {f : α → β} {g : β → α} {h : β → γ} {i : γ → β} (hf : LeftInverse f g) (hh : LeftInverse h i) : LeftInverse (h ∘ f) (g ∘ i) := fun a ↦ show h (f (g (i a))) = a by rw [hf (i a), hh a] +#align function.left_inverse.comp Function.LeftInverse.comp theorem RightInverse.comp {f : α → β} {g : β → α} {h : β → γ} {i : γ → β} (hf : RightInverse f g) (hh : RightInverse h i) : RightInverse (h ∘ f) (g ∘ i) := LeftInverse.comp hh hf +#align function.right_inverse.comp Function.RightInverse.comp theorem LeftInverse.rightInverse {f : α → β} {g : β → α} (h : LeftInverse g f) : RightInverse f g := h @@ -311,9 +355,11 @@ theorem RightInverse.leftInverse {f : α → β} {g : β → α} (h : RightInver theorem LeftInverse.surjective {f : α → β} {g : β → α} (h : LeftInverse f g) : Surjective f := h.rightInverse.surjective +#align function.left_inverse.surjective Function.LeftInverse.surjective theorem RightInverse.injective {f : α → β} {g : β → α} (h : RightInverse f g) : Injective f := h.leftInverse.injective +#align function.right_inverse.injective Function.RightInverse.injective theorem LeftInverse.rightInverse_of_injective {f : α → β} {g : β → α} (h : LeftInverse f g) (hf : Injective f) : RightInverse f g := @@ -348,6 +394,7 @@ attribute [local instance] Classical.propDecidable a given injective function `f`. -/ noncomputable def partialInv {α β} (f : α → β) (b : β) : Option α := if h : ∃ a, f a = b then some (Classical.choose h) else none +#align function.partial_inv Function.partialInv theorem partialInv_of_injective {α β} {f : α → β} (I : Injective f) : IsPartialInv f (partialInv f) | a, b => @@ -380,6 +427,7 @@ attribute [local instance] Classical.propDecidable -- Explicit Sort so that `α` isn't inferred to be Prop via `exists_prop_decidable` noncomputable def invFun {α : Sort u} {β} [Nonempty α] (f : α → β) : β → α := fun y ↦ if h : (∃ x, f x = y) then h.choose else Classical.arbitrary α +#align function.inv_fun Function.invFun theorem invFun_eq (h : ∃ a, f a = b) : f (invFun f b) = b := by simp only [invFun, dif_pos h, h.choose_spec] @@ -432,6 +480,7 @@ variable {α : Sort u} {β : Sort v} {γ : Sort w} {f : α → β} `α` to be inhabited.) -/ noncomputable def surjInv {f : α → β} (h : Surjective f) (b : β) : α := Classical.choose (h b) +#align function.surj_inv Function.surjInv theorem surjInv_eq (h : Surjective f) (b) : f (surjInv h b) = b := Classical.choose_spec (h b) @@ -456,6 +505,7 @@ theorem surjective_iff_hasRightInverse : Surjective f ↔ HasRightInverse f := theorem bijective_iff_has_inverse : Bijective f ↔ ∃ g, LeftInverse g f ∧ RightInverse g f := ⟨fun hf ↦ ⟨_, leftInverse_surjInv hf, rightInverse_surjInv hf.2⟩, fun ⟨_, gl, gr⟩ ↦ ⟨gl.injective, gr.surjective⟩⟩ +#align function.bijective_iff_has_inverse Function.bijective_iff_has_inverse theorem injective_surjInv (h : Surjective f) : Injective (surjInv h) := (rightInverse_surjInv h).injective @@ -464,16 +514,19 @@ theorem injective_surjInv (h : Surjective f) : Injective (surjInv h) := theorem surjective_to_subsingleton [na : Nonempty α] [Subsingleton β] (f : α → β) : Surjective f := fun _ ↦ let ⟨a⟩ := na; ⟨a, Subsingleton.elim _ _⟩ +#align function.surjective_to_subsingleton Function.surjective_to_subsingleton /-- Composition by an surjective function on the left is itself surjective. -/ theorem Surjective.comp_left {g : β → γ} (hg : Surjective g) : Surjective ((· ∘ ·) g : (α → β) → α → γ) := fun f ↦ ⟨surjInv hg ∘ f, funext fun _ ↦ rightInverse_surjInv _ _⟩ +#align function.surjective.comp_left Function.Surjective.comp_left /-- Composition by an bijective function on the left is itself bijective. -/ theorem Bijective.comp_left {g : β → γ} (hg : Bijective g) : Bijective ((· ∘ ·) g : (α → β) → α → γ) := ⟨hg.injective.comp_left, hg.surjective.comp_left⟩ +#align function.bijective.comp_left Function.Bijective.comp_left end SurjInv @@ -486,6 +539,7 @@ variable {α : Sort u} {β : α → Sort v} {α' : Sort w} [DecidableEq α] [Dec /-- Replacing the value of a function at a given point by a given value. -/ def update (f : ∀ a, β a) (a' : α) (v : β a') (a : α) : β a := if h : a = a' then Eq.ndrec v h.symm else f a +#align function.update Function.update /-- On non-dependent functions, `Function.update` can be expressed as an `ite` -/ theorem update_apply {β : Sort _} (f : α → β) (a' : α) (b : β) (a : α) : @@ -497,70 +551,88 @@ by have h2 : (h : a = a') → Eq.rec (motive := λ _ _ => β) b h.symm = b := (λ _ : a = a' => b) := funext h2 let f := λ x => dite (a = a') x (λ (_: ¬ a = a') => (f a)) exact congrArg f h3 +#align function.update_apply Function.update_apply @[simp] theorem update_same (a : α) (v : β a) (f : ∀ a, β a) : update f a v a = v := dif_pos rfl +#align function.update_same Function.update_same theorem surjective_eval {α : Sort u} {β : α → Sort v} [h : ∀ a, Nonempty (β a)] (a : α) : Surjective (eval a : (∀ a, β a) → β a) := fun b ↦ ⟨@update _ _ (Classical.decEq α) (fun a ↦ (h a).some) a b, @update_same _ _ (Classical.decEq α) _ _ _⟩ +#align function.surjective_eval Function.surjective_eval theorem update_injective (f : ∀ a, β a) (a' : α) : Injective (update f a') := fun v v' h ↦ by have := congr_fun h a' rwa [update_same, update_same] at this +#align function.update_injective Function.update_injective @[simp] theorem update_noteq {a a' : α} (h : a ≠ a') (v : β a') (f : ∀ a, β a) : update f a' v a = f a := dif_neg h +#align function.update_noteq Function.update_noteq lemma forall_update_iff (f : ∀a, β a) {a : α} {b : β a} (p : ∀a, β a → Prop) : (∀ x, p x (update f a b x)) ↔ p a b ∧ ∀ x, x ≠ a → p x (f x) := by rw [← and_forall_ne a, update_same] simp (config := { contextual := true }) +#align function.forall_update_iff Function.forall_update_iff theorem exists_update_iff (f : ∀ a, β a) {a : α} {b : β a} (p : ∀ a, β a → Prop) : (∃ x, p x (update f a b x)) ↔ p a b ∨ ∃ (x : _)(_ : x ≠ a), p x (f x) := by rw [← not_forall_not, forall_update_iff f fun a b ↦ ¬p a b] simp [-not_and, not_and_or] +#align function.exists_update_iff Function.exists_update_iff theorem update_eq_iff {a : α} {b : β a} {f g : ∀ a, β a} : update f a b = g ↔ b = g a ∧ ∀ (x) (_ : x ≠ a), f x = g x := funext_iff.trans <| forall_update_iff _ fun x y ↦ y = g x +#align function.update_eq_iff Function.update_eq_iff theorem eq_update_iff {a : α} {b : β a} {f g : ∀ a, β a} : g = update f a b ↔ g a = b ∧ ∀ (x) (_ : x ≠ a), g x = f x := funext_iff.trans <| forall_update_iff _ fun x y ↦ g x = y +#align function.eq_update_iff Function.eq_update_iff @[simp] lemma update_eq_self_iff : update f a b = f ↔ b = f a := by simp [update_eq_iff] @[simp] lemma eq_update_self_iff : f = update f a b ↔ f a = b := by simp [eq_update_iff] +#align function.eq_update_self_iff Function.eq_update_self_iff +#align function.update_eq_self_iff Function.update_eq_self_iff lemma ne_update_self_iff : f ≠ update f a b ↔ f a ≠ b := eq_update_self_iff.not lemma update_ne_self_iff : update f a b ≠ f ↔ b ≠ f a := update_eq_self_iff.not +#align function.update_ne_self_iff Function.update_ne_self_iff +#align function.ne_update_self_iff Function.ne_update_self_iff @[simp] theorem update_eq_self (a : α) (f : ∀ a, β a) : update f a (f a) = f := update_eq_iff.2 ⟨rfl, fun _ _ ↦ rfl⟩ +#align function.update_eq_self Function.update_eq_self theorem update_comp_eq_of_forall_ne' {α'} (g : ∀ a, β a) {f : α' → α} {i : α} (a : β i) (h : ∀ x, f x ≠ i) : (fun j ↦ (update g i a) (f j)) = fun j ↦ g (f j) := funext fun _ ↦ update_noteq (h _) _ _ +#align function.update_comp_eq_of_forall_ne' Function.update_comp_eq_of_forall_ne' /-- Non-dependent version of `Function.update_comp_eq_of_forall_ne'` -/ theorem update_comp_eq_of_forall_ne {α β : Sort _} (g : α' → β) {f : α → α'} {i : α'} (a : β) (h : ∀ x, f x ≠ i) : update g i a ∘ f = g ∘ f := update_comp_eq_of_forall_ne' g a h +#align function.update_comp_eq_of_forall_ne Function.update_comp_eq_of_forall_ne theorem update_comp_eq_of_injective' (g : ∀ a, β a) {f : α' → α} (hf : Function.Injective f) (i : α') (a : β (f i)) : (fun j ↦ update g (f i) a (f j)) = update (fun i ↦ g (f i)) i a := eq_update_iff.2 ⟨update_same _ _ _, fun _ hj ↦ update_noteq (hf.ne hj) _ _⟩ +#align function.update_comp_eq_of_injective' Function.update_comp_eq_of_injective' /-- Non-dependent version of `Function.update_comp_eq_of_injective'` -/ theorem update_comp_eq_of_injective {β : Sort _} (g : α' → β) {f : α → α'} (hf : Function.Injective f) (i : α) (a : β) : Function.update g (f i) a ∘ f = Function.update (g ∘ f) i a := update_comp_eq_of_injective' g hf i a +#align function.update_comp_eq_of_injective Function.update_comp_eq_of_injective theorem apply_update {ι : Sort _} [DecidableEq ι] {α β : ι → Sort _} (f : ∀ i, α i → β i) (g : ∀ i, α i) (i : ι) (v : α i) (j : ι) : @@ -569,6 +641,7 @@ theorem apply_update {ι : Sort _} [DecidableEq ι] {α β : ι → Sort _} (f : · subst j simp · simp [h] +#align function.apply_update Function.apply_update theorem apply_update₂ {ι : Sort _} [DecidableEq ι] {α β γ : ι → Sort _} (f : ∀ i, α i → β i → γ i) (g : ∀ i, α i) (h : ∀ i, β i) (i : ι) (v : α i) (w : β i) (j : ι) : @@ -577,11 +650,13 @@ theorem apply_update₂ {ι : Sort _} [DecidableEq ι] {α β γ : ι → Sort _ · subst j simp · simp [h] +#align function.apply_update₂ Function.apply_update₂ theorem comp_update {α' : Sort _} {β : Sort _} (f : α' → β) (g : α → α') (i : α) (v : α') : f ∘ update g i v = update (f ∘ g) i (f v) := funext <| apply_update _ _ _ _ +#align function.comp_update Function.comp_update theorem update_comm {α} [DecidableEq α] {β : α → Sort _} {a b : α} (h : a ≠ b) (v : β a) (w : β b) (f : ∀ a, β a) : update (update f a v) b w = update (update f b w) a v := by @@ -593,12 +668,14 @@ theorem update_comm {α} [DecidableEq α] {β : α → Sort _} {a b : α} (h : a · rw [dif_pos h₁, dif_pos h₁, dif_neg h₂] · rw [dif_neg h₁, dif_neg h₁, dif_pos h₂] · rw [dif_neg h₁, dif_neg h₁, dif_neg h₂] +#align function.update_comm Function.update_comm @[simp] theorem update_idem {α} [DecidableEq α] {β : α → Sort _} {a : α} (v w : β a) (f : ∀ a, β a) : update (update f a v) a w = update f a w := by funext b by_cases b = a <;> simp [update, h] +#align function.update_idem Function.update_idem end Update @@ -617,39 +694,47 @@ Mostly useful when `f` is injective, or more generally when `g.factors_through f -- Explicit Sort so that `α` isn't inferred to be Prop via `exists_prop_decidable` def extend {α : Sort u} {β γ} (f : α → β) (g : α → γ) (e' : β → γ) : β → γ := fun b ↦ if h : ∃ a, f a = b then g (Classical.choose h) else e' b +#align function.extend Function.extend /-- g factors through f : `f a = f b → g a = g b` -/ def FactorsThrough (g : α → γ) (f : α → β) : Prop := ∀ ⦃a b⦄, f a = f b → g a = g b +#align function.factors_through Function.FactorsThrough theorem extend_def (f : α → β) (g : α → γ) (e' : β → γ) (b : β) [Decidable (∃ a, f a = b)] : extend f g e' b = if h : ∃ a, f a = b then g (Classical.choose h) else e' b := by unfold extend congr +#align function.extend_def Function.extend_def lemma Injective.FactorsThrough (hf : Injective f) (g : α → γ) : g.FactorsThrough f := fun _ _ h => congr_arg g (hf h) +#align function.injective.factors_through Function.Injective.FactorsThrough lemma FactorsThrough.extend_apply {g : α → γ} (hf : g.FactorsThrough f) (e' : β → γ) (a : α) : extend f g e' (f a) = g a := by simp only [extend_def, dif_pos, exists_apply_eq_apply] exact hf (Classical.choose_spec (exists_apply_eq_apply f a)) +#align function.factors_through.extend_apply Function.FactorsThrough.extend_apply @[simp] theorem Injective.extend_apply (hf : Injective f) (g : α → γ) (e' : β → γ) (a : α) : extend f g e' (f a) = g a := (hf.FactorsThrough g).extend_apply e' a +#align function.injective.extend_apply Function.Injective.extend_apply @[simp] theorem extend_apply' (g : α → γ) (e' : β → γ) (b : β) (hb : ¬∃ a, f a = b) : extend f g e' b = e' b := by simp [Function.extend_def, hb] +#align function.extend_apply' Function.extend_apply' lemma factorsThrough_iff (g : α → γ) [Nonempty γ] : g.FactorsThrough f ↔ ∃ (e : β → γ), g = e ∘ f := ⟨fun hf => ⟨extend f g (const β (Classical.arbitrary γ)), funext (fun x => by simp only [comp_apply, hf.extend_apply])⟩, fun h _ _ hf => by rw [Classical.choose_spec h, comp_apply, comp_apply, hf]⟩ +#align function.factors_through_iff Function.factorsThrough_iff lemma FactorsThrough.apply_extend {δ} {g : α → γ} (hf : FactorsThrough g f) (F : γ → δ) (e' : β → γ) (b : β) : @@ -666,10 +751,12 @@ lemma FactorsThrough.apply_extend {δ} {g : α → γ} (hf : FactorsThrough g f) exact hf h case neg => rw [extend_apply' _ _ _ hb, extend_apply' _ _ _ hb, comp] +#align function.factors_through.apply_extend Function.FactorsThrough.apply_extend lemma Injective.apply_extend {δ} (hf : Injective f) (F : γ → δ) (g : α → γ) (e' : β → γ) (b : β) : F (extend f g e' b) = extend f (F ∘ g) (F ∘ e') b := (hf.FactorsThrough g).apply_extend F e' b +#align function.injective.apply_extend Function.Injective.apply_extend theorem extend_injective (hf : Injective f) (e' : β → γ) : Injective fun g ↦ extend f g e' := by intro g₁ g₂ hg @@ -677,40 +764,49 @@ theorem extend_injective (hf : Injective f) (e' : β → γ) : Injective fun g have H := congr_fun hg (f x) simp only [hf.extend_apply] at H exact H +#align function.extend_injective Function.extend_injective lemma FactorsThrough.extend_comp {g : α → γ} (e' : β → γ) (hf : FactorsThrough g f) : extend f g e' ∘ f = g := funext $ fun a => hf.extend_apply e' a +#align function.factors_through.extend_comp Function.FactorsThrough.extend_comp @[simp] theorem extend_comp (hf : Injective f) (g : α → γ) (e' : β → γ) : extend f g e' ∘ f = g := funext fun a ↦ hf.extend_apply g e' a +#align function.extend_comp Function.extend_comp theorem Injective.surjective_comp_right' (hf : Injective f) (g₀ : β → γ) : Surjective fun g : β → γ ↦ g ∘ f := fun g ↦ ⟨extend f g g₀, extend_comp hf _ _⟩ +#align function.injective.surjective_comp_right' Function.Injective.surjective_comp_right' theorem Injective.surjective_comp_right [Nonempty γ] (hf : Injective f) : Surjective fun g : β → γ ↦ g ∘ f := hf.surjective_comp_right' fun _ ↦ Classical.choice ‹_› +#align function.injective.surjective_comp_right Function.Injective.surjective_comp_right theorem Bijective.comp_right (hf : Bijective f) : Bijective fun g : β → γ ↦ g ∘ f := ⟨hf.surjective.injective_comp_right, fun g ↦ ⟨g ∘ surjInv hf.surjective, by simp only [comp.assoc g _ f, (leftInverse_surjInv hf).comp_eq_id, comp.right_id]⟩⟩ +#align function.bijective.comp_right Function.Bijective.comp_right end Extend theorem uncurry_def {α β γ} (f : α → β → γ) : uncurry f = fun p ↦ f p.1 p.2 := rfl +#align function.uncurry_def Function.uncurry_def @[simp] theorem uncurry_apply_pair {α β γ} (f : α → β → γ) (x : α) (y : β) : uncurry f (x, y) = f x y := rfl +#align function.uncurry_apply_pair Function.uncurry_apply_pair @[simp] theorem curry_apply {α β γ} (f : α × β → γ) (x : α) (y : β) : curry f x y = f (x, y) := rfl +#align function.curry_apply Function.curry_apply section Bicomp @@ -720,20 +816,24 @@ variable {α β γ δ ε : Type _} If both arguments of `f` have the same type and `g = h`, then `bicompl f g g = f on g`. -/ def bicompl (f : γ → δ → ε) (g : α → γ) (h : β → δ) (a b) := f (g a) (h b) +#align function.bicompl Function.bicompl /-- Compose an unary function `f` with a binary function `g`. -/ def bicompr (f : γ → δ) (g : α → β → γ) (a b) := f (g a b) +#align function.bicompr Function.bicompr -- Suggested local notation: local notation f " ∘₂ " g => bicompr f g theorem uncurry_bicompr (f : α → β → γ) (g : γ → δ) : uncurry (g ∘₂ f) = g ∘ uncurry f := rfl +#align function.uncurry_bicompr Function.uncurry_bicompr theorem uncurry_bicompl (f : γ → δ → ε) (g : α → γ) (h : β → δ) : uncurry (bicompl f g h) = uncurry f ∘ Prod.map g h := rfl +#align function.uncurry_bicompl Function.uncurry_bicompl end Bicomp @@ -749,6 +849,7 @@ class HasUncurry (α : Type _) (β : outParam (Type _)) (γ : outParam (Type _)) `f : α → β → γ → δ` will be turned into `↿f : α × β × γ → δ`. One can also add instances for bundled maps.-/ uncurry : α → β → γ +#align function.has_uncurry Function.HasUncurry notation:arg "↿" x:arg => HasUncurry.uncurry x @@ -763,9 +864,11 @@ end Uncurry /-- A function is involutive, if `f ∘ f = id`. -/ def Involutive {α} (f : α → α) : Prop := ∀ x, f (f x) = x +#align function.involutive Function.Involutive theorem involutive_iff_iter_2_eq_id {α} {f : α → α} : Involutive f ↔ f^[2] = id := funext_iff.symm +#align function.involutive_iff_iter_2_eq_id Function.involutive_iff_iter_2_eq_id theorem _root_.Bool.involutive_not : Involutive not := Bool.not_not @@ -777,6 +880,7 @@ variable {α : Sort u} {f : α → α} (h : Involutive f) @[simp] theorem comp_self : f ∘ f = id := funext h +#align function.involutive.comp_self Function.Involutive.comp_self protected theorem leftInverse : LeftInverse f f := h #align function.involutive.left_inverse Function.Involutive.leftInverse @@ -796,10 +900,12 @@ protected theorem bijective : Bijective f := ⟨h.injective, h.surjective⟩ /-- Involuting an `ite` of an involuted value `x : α` negates the `Prop` condition in the `ite`. -/ protected theorem ite_not (P : Prop) [Decidable P] (x : α) : f (ite P x (f x)) = ite (¬P) x (f x) := by rw [apply_ite f, h, ite_not] +#align function.involutive.ite_not Function.Involutive.ite_not /-- An involution commutes across an equality. Compare to `Function.Injective.eq_iff`. -/ protected theorem eq_iff {x y : α} : f x = y ↔ x = f y := h.injective.eq_iff' (h y) +#align function.involutive.eq_iff Function.Involutive.eq_iff end Involutive @@ -808,6 +914,7 @@ Mathematically this should be thought of as the corresponding function `α × β -/ def Injective2 {α β γ} (f : α → β → γ) : Prop := ∀ ⦃a₁ a₂ b₁ b₂⦄, f a₁ b₁ = f a₂ b₂ → a₁ = a₂ ∧ b₁ = b₂ +#align function.injective2 Function.Injective2 namespace Injective2 @@ -816,28 +923,34 @@ variable {α β γ : Sort _} {f : α → β → γ} /-- A binary injective function is injective when only the left argument varies. -/ protected theorem left (hf : Injective2 f) (b : β) : Function.Injective fun a ↦ f a b := fun _ _ h ↦ (hf h).left +#align function.injective2.left Function.Injective2.left /-- A binary injective function is injective when only the right argument varies. -/ protected theorem right (hf : Injective2 f) (a : α) : Function.Injective (f a) := fun _ _ h ↦ (hf h).right +#align function.injective2.right Function.Injective2.right protected theorem uncurry {α β γ : Type _} {f : α → β → γ} (hf : Injective2 f) : Function.Injective (uncurry f) := fun ⟨_, _⟩ ⟨_, _⟩ h ↦ (hf h).elim (congr_arg₂ _) +#align function.injective2.uncurry Function.Injective2.uncurry /-- As a map from the left argument to a unary function, `f` is injective. -/ theorem left' (hf : Injective2 f) [Nonempty β] : Function.Injective f := fun a₁ a₂ h ↦ let ⟨b⟩ := ‹Nonempty β› hf.left b <| (congr_fun h b : _) +#align function.injective2.left' Function.Injective2.left' /-- As a map from the right argument to a unary function, `f` is injective. -/ theorem right' (hf : Injective2 f) [Nonempty α] : Function.Injective fun b a ↦ f a b := fun b₁ b₂ h ↦ let ⟨a⟩ := ‹Nonempty α› hf.right a <| (congr_fun h a : _) +#align function.injective2.right' Function.Injective2.right' theorem eq_iff (hf : Injective2 f) {a₁ a₂ b₁ b₂} : f a₁ b₁ = f a₂ b₂ ↔ a₁ = a₂ ∧ b₁ = b₂ := ⟨fun h ↦ hf h, fun ⟨h1, h2⟩ ↦ congr_arg₂ f h1 h2⟩ +#align function.injective2.eq_iff Function.Injective2.eq_iff end Injective2 @@ -850,13 +963,16 @@ interesting in the case where `α` is a proposition, in which case `f` is necess constant function, so that `sometimes f = f a` for all `a`. -/ noncomputable def sometimes {α β} [Nonempty β] (f : α → β) : β := if h : Nonempty α then f (Classical.choice h) else Classical.choice ‹_› +#align function.sometimes Function.sometimes theorem sometimes_eq {p : Prop} {α} [Nonempty α] (f : p → α) (a : p) : sometimes f = f a := dif_pos ⟨a⟩ +#align function.sometimes_eq Function.sometimes_eq theorem sometimes_spec {p : Prop} {α} [Nonempty α] (P : α → Prop) (f : p → α) (a : p) (h : P (f a)) : P (sometimes f) := by rwa [sometimes_eq] +#align function.sometimes_spec Function.sometimes_spec end Sometimes @@ -866,6 +982,7 @@ end Function def Set.piecewise {α : Type u} {β : α → Sort v} (s : Set α) (f g : ∀ i, β i) [∀ j, Decidable (j ∈ s)] : ∀ i, β i := fun i ↦ if i ∈ s then f i else g i +#align set.piecewise Set.piecewise /-! ### Bijectivity of `eq.rec`, `eq.mp`, `eq.mpr`, and `cast` -/ @@ -873,20 +990,24 @@ def Set.piecewise {α : Type u} {β : α → Sort v} (s : Set α) (f g : ∀ i, theorem eq_rec_on_bijective {α : Sort _} {C : α → Sort _} : ∀ {a a' : α} (h : a = a'), Function.Bijective (@Eq.ndrec _ _ C · _ h) | _, _, rfl => ⟨fun _ _ ↦ id, fun x ↦ ⟨x, rfl⟩⟩ +#align eq_rec_on_bijective eq_rec_on_bijective theorem eq_mp_bijective {α β : Sort _} (h : α = β) : Function.Bijective (Eq.mp h) := by -- TODO: mathlib3 uses `eq_rec_on_bijective`, difference in elaboration here -- due to `@[macro_inline] possibly? cases h refine ⟨fun _ _ ↦ id, fun x ↦ ⟨x, rfl⟩⟩ +#align eq_mp_bijective eq_mp_bijective theorem eq_mpr_bijective {α β : Sort _} (h : α = β) : Function.Bijective (Eq.mpr h) := by cases h refine ⟨fun _ _ ↦ id, fun x ↦ ⟨x, rfl⟩⟩ +#align eq_mpr_bijective eq_mpr_bijective theorem cast_bijective {α β : Sort _} (h : α = β) : Function.Bijective (cast h) := by cases h refine ⟨fun _ _ ↦ id, fun x ↦ ⟨x, rfl⟩⟩ +#align cast_bijective cast_bijective /-! Note these lemmas apply to `Type _` not `Sort*`, as the latter interferes with `simp`, and is trivial anyway.-/ @@ -896,36 +1017,44 @@ is trivial anyway.-/ theorem eq_rec_inj {α : Sort _} {a a' : α} (h : a = a') {C : α → Type _} (x y : C a) : (Eq.ndrec x h : C a') = Eq.ndrec y h ↔ x = y := (eq_rec_on_bijective h).injective.eq_iff +#align eq_rec_inj eq_rec_inj @[simp] theorem cast_inj {α β : Type _} (h : α = β) {x y : α} : cast h x = cast h y ↔ x = y := (cast_bijective h).injective.eq_iff +#align cast_inj cast_inj theorem Function.LeftInverse.eq_rec_eq {α β : Sort _} {γ : β → Sort v} {f : α → β} {g : β → α} (h : Function.LeftInverse g f) (C : ∀ a : α, γ (f a)) (a : α) : -- TODO: mathlib3 uses `(congr_arg f (h a)).rec (C (g (f a)))` for LHS @Eq.rec β (f (g (f a))) (fun x _ ↦ γ x) (C (g (f a))) (f a) (congr_arg f (h a)) = C a := eq_of_heq <| (eq_rec_heq _ _).trans <| by rw [h] +#align function.left_inverse.eq_rec_eq Function.LeftInverse.eq_rec_eq theorem Function.LeftInverse.eq_rec_on_eq {α β : Sort _} {γ : β → Sort v} {f : α → β} {g : β → α} (h : Function.LeftInverse g f) (C : ∀ a : α, γ (f a)) (a : α) : -- TODO: mathlib3 uses `(congr_arg f (h a)).recOn (C (g (f a)))` for LHS @Eq.recOn β (f (g (f a))) (fun x _ ↦ γ x) (f a) (congr_arg f (h a)) (C (g (f a))) = C a := h.eq_rec_eq _ _ +#align function.left_inverse.eq_rec_on_eq Function.LeftInverse.eq_rec_on_eq theorem Function.LeftInverse.cast_eq {α β : Sort _} {γ : β → Sort v} {f : α → β} {g : β → α} (h : Function.LeftInverse g f) (C : ∀ a : α, γ (f a)) (a : α) : cast (congr_arg (fun a ↦ γ (f a)) (h a)) (C (g (f a))) = C a := by rw [cast_eq_iff_heq, h] +#align function.left_inverse.cast_eq Function.LeftInverse.cast_eq /-- A set of functions "separates points" if for each pair of distinct points there is a function taking different values on them. -/ def Set.SeparatesPoints {α β : Type _} (A : Set (α → β)) : Prop := ∀ ⦃x y : α⦄, x ≠ y → ∃ f ∈ A, (f x : β) ≠ f y +#align set.separates_points Set.SeparatesPoints theorem IsSymmOp.flip_eq {α β} (op) [IsSymmOp α β op] : flip op = op := funext fun a ↦ funext fun b ↦ (IsSymmOp.symm_op a b).symm +#align is_symm_op.flip_eq IsSymmOp.flip_eq theorem InvImage.equivalence {α : Sort u} {β : Sort v} (r : β → β → Prop) (f : α → β) (h : Equivalence r) : Equivalence (InvImage r f) := ⟨fun _ ↦ h.1 _, fun w ↦ h.symm w, fun h₁ h₂ ↦ InvImage.trans r f (fun _ _ _ ↦ h.trans) h₁ h₂⟩ +#align inv_image.equivalence InvImage.equivalence diff --git a/Mathlib/Logic/Function/Conjugate.lean b/Mathlib/Logic/Function/Conjugate.lean index 282614ef362d6..82936d38c1742 100644 --- a/Mathlib/Logic/Function/Conjugate.lean +++ b/Mathlib/Logic/Function/Conjugate.lean @@ -34,6 +34,7 @@ We use `∀ x, f (ga x) = gb (f x)` as the definition, so given `h : Function.Se -/ def Semiconj (f : α → β) (ga : α → α) (gb : β → β) : Prop := ∀ x, f (ga x) = gb (f x) +#align function.semiconj Function.Semiconj namespace Semiconj @@ -41,32 +42,40 @@ variable {f fab : α → β} {fbc : β → γ} {ga ga' : α → α} {gb gb' : β protected theorem comp_eq (h : Semiconj f ga gb) : f ∘ ga = gb ∘ f := funext h +#align function.semiconj.comp_eq Function.Semiconj.comp_eq protected theorem eq (h : Semiconj f ga gb) (x : α) : f (ga x) = gb (f x) := h x +#align function.semiconj.eq Function.Semiconj.eq theorem comp_right (h : Semiconj f ga gb) (h' : Semiconj f ga' gb') : Semiconj f (ga ∘ ga') (gb ∘ gb') := fun x ↦ by simp only [comp_apply, h.eq, h'.eq] +#align function.semiconj.comp_right Function.Semiconj.comp_right theorem comp_left (hab : Semiconj fab ga gb) (hbc : Semiconj fbc gb gc) : Semiconj (fbc ∘ fab) ga gc := fun x ↦ by simp only [comp_apply, hab.eq, hbc.eq] +#align function.semiconj.comp_left Function.Semiconj.comp_left theorem id_right : Semiconj f id id := fun _ ↦ rfl +#align function.semiconj.id_right Function.Semiconj.id_right theorem id_left : Semiconj id ga ga := fun _ ↦ rfl +#align function.semiconj.id_left Function.Semiconj.id_left theorem inverses_right (h : Semiconj f ga gb) (ha : RightInverse ga' ga) (hb : LeftInverse gb' gb) : Semiconj f ga' gb' := fun x ↦ by rw [← hb (f (ga' x)), ← h.eq, ha x] +#align function.semiconj.inverses_right Function.Semiconj.inverses_right theorem option_map {f : α → β} {ga : α → α} {gb : β → β} (h : Semiconj f ga gb) : Semiconj (Option.map f) (Option.map ga) (Option.map gb) | none => rfl | some _ => congr_arg some <| h _ +#align function.semiconj.option_map Function.Semiconj.option_map end Semiconj @@ -77,10 +86,12 @@ Given `h : Function.commute f g` and `a : α`, we have `h a : f (g a) = g (f a)` -/ def Commute (f g : α → α) : Prop := Semiconj f g g +#align function.commute Function.Commute /-- Reinterpret `Function.Semiconj f g g` as `Function.Commute f g`. These two predicates are definitionally equal but have different dot-notation lemmas. -/ theorem Semiconj.commute {f g : α → α} (h : Semiconj f g g) : Commute f g := h +#align function.semiconj.commute Function.Semiconj.commute namespace Commute @@ -93,25 +104,32 @@ theorem semiconj (h : Commute f g) : Semiconj f g g := h @[refl] theorem refl (f : α → α) : Commute f f := fun _ ↦ Eq.refl _ +#align function.commute.refl Function.Commute.refl @[symm] theorem symm (h : Commute f g) : Commute g f := fun x ↦ (h x).symm +#align function.commute.symm Function.Commute.symm theorem comp_right (h : Commute f g) (h' : Commute f g') : Commute f (g ∘ g') := Semiconj.comp_right h h' +#align function.commute.comp_right Function.Commute.comp_right theorem comp_left (h : Commute f g) (h' : Commute f' g) : Commute (f ∘ f') g := (h.symm.comp_right h'.symm).symm +#align function.commute.comp_left Function.Commute.comp_left theorem id_right : Commute f id := Semiconj.id_right +#align function.commute.id_right Function.Commute.id_right theorem id_left : Commute id f := Semiconj.id_left +#align function.commute.id_left Function.Commute.id_left theorem option_map {f g : α → α} : Commute f g → Commute (Option.map f) (Option.map g) := Semiconj.option_map +#align function.commute.option_map Function.Commute.option_map end Commute @@ -122,6 +140,7 @@ semiconjugates `(*)` to `(*)`. -/ def Semiconj₂ (f : α → β) (ga : α → α → α) (gb : β → β → β) : Prop := ∀ x y, f (ga x y) = gb (f x) (f y) +#align function.semiconj₂ Function.Semiconj₂ namespace Semiconj₂ @@ -129,14 +148,18 @@ variable {f : α → β} {ga : α → α → α} {gb : β → β → β} protected theorem eq (h : Semiconj₂ f ga gb) (x y : α) : f (ga x y) = gb (f x) (f y) := h x y +#align function.semiconj₂.eq Function.Semiconj₂.eq protected theorem comp_eq (h : Semiconj₂ f ga gb) : bicompr f ga = bicompl gb f f := funext fun x ↦ funext <| h x +#align function.semiconj₂.comp_eq Function.Semiconj₂.comp_eq theorem id_left (op : α → α → α) : Semiconj₂ id op op := fun _ _ ↦ rfl +#align function.semiconj₂.id_left Function.Semiconj₂.id_left theorem comp {f' : β → γ} {gc : γ → γ → γ} (hf' : Semiconj₂ f' gb gc) (hf : Semiconj₂ f ga gb) : Semiconj₂ (f' ∘ f) ga gc := fun x y ↦ by simp only [hf'.eq, hf.eq, comp_apply] +#align function.semiconj₂.comp Function.Semiconj₂.comp theorem isAssociative_right [IsAssociative α ga] (h : Semiconj₂ f ga gb) (h_surj : Surjective f) : IsAssociative β gb := diff --git a/Mathlib/Logic/Function/Iterate.lean b/Mathlib/Logic/Function/Iterate.lean index 05284494ae640..eff1fadec765e 100644 --- a/Mathlib/Logic/Function/Iterate.lean +++ b/Mathlib/Logic/Function/Iterate.lean @@ -44,66 +44,79 @@ variable (f : α → α) @[simp] theorem iterate_zero : f^[0] = id := rfl +#align function.iterate_zero Function.iterate_zero theorem iterate_zero_apply (x : α) : (f^[0]) x = x := rfl +#align function.iterate_zero_apply Function.iterate_zero_apply @[simp] theorem iterate_succ (n : ℕ) : f^[n.succ] = f^[n] ∘ f := rfl +#align function.iterate_succ Function.iterate_succ theorem iterate_succ_apply (n : ℕ) (x : α) : (f^[n.succ]) x = (f^[n]) (f x) := rfl +#align function.iterate_succ_apply Function.iterate_succ_apply @[simp] theorem iterate_id (n : ℕ) : (id : α → α)^[n] = id := Nat.recOn n rfl fun n ihn ↦ by rw [iterate_succ, ihn, comp.left_id] +#align function.iterate_id Function.iterate_id theorem iterate_add (m : ℕ) : ∀ n : ℕ, f^[m + n] = f^[m] ∘ f^[n] | 0 => rfl | Nat.succ n => by rw [Nat.add_succ, iterate_succ, iterate_succ, iterate_add m n]; rfl +#align function.iterate_add Function.iterate_add theorem iterate_add_apply (m n : ℕ) (x : α) : (f^[m + n]) x = (f^[m]) ((f^[n]) x) := by rw [iterate_add f m n] rfl +#align function.iterate_add_apply Function.iterate_add_apply @[simp] theorem iterate_one : f^[1] = f := funext fun _ ↦ rfl +#align function.iterate_one Function.iterate_one theorem iterate_mul (m : ℕ) : ∀ n, f^[m * n] = f^[m]^[n] | 0 => by simp only [Nat.mul_zero, iterate_zero] | n + 1 => by simp only [Nat.mul_succ, Nat.mul_one, iterate_one, iterate_add, iterate_mul m n] +#align function.iterate_mul Function.iterate_mul variable {f} theorem iterate_fixed {x} (h : f x = x) (n : ℕ) : (f^[n]) x = x := Nat.recOn n rfl fun n ihn ↦ by rw [iterate_succ_apply, h, ihn] +#align function.iterate_fixed Function.iterate_fixed theorem Injective.iterate (Hinj : Injective f) (n : ℕ) : Injective (f^[n]) := Nat.recOn n injective_id fun _ ihn ↦ ihn.comp Hinj +#align function.injective.iterate Function.Injective.iterate theorem Surjective.iterate (Hsurj : Surjective f) (n : ℕ) : Surjective (f^[n]) := Nat.recOn n surjective_id fun _ ihn ↦ ihn.comp Hsurj +#align function.surjective.iterate Function.Surjective.iterate theorem Bijective.iterate (Hbij : Bijective f) (n : ℕ) : Bijective (f^[n]) := ⟨Hbij.1.iterate n, Hbij.2.iterate n⟩ +#align function.bijective.iterate Function.Bijective.iterate namespace Semiconj theorem iterate_right {f : α → β} {ga : α → α} {gb : β → β} (h : Semiconj f ga gb) (n : ℕ) : Semiconj f (ga^[n]) (gb^[n]) := Nat.recOn n id_right fun _ ihn ↦ ihn.comp_right h +#align function.semiconj.iterate_right Function.Semiconj.iterate_right theorem iterate_left {g : ℕ → α → α} (H : ∀ n, Semiconj f (g n) (g <| n + 1)) (n k : ℕ) : Semiconj (f^[n]) (g k) (g <| n + k) := by induction' n with n ihn generalizing k · rw [Nat.zero_add] exact id_left - · rw [Nat.succ_eq_add_one, Nat.add_right_comm, Nat.add_assoc] exact (H k).comp_left (ihn (k + 1)) - +#align function.semiconj.iterate_left Function.Semiconj.iterate_left end Semiconj @@ -113,56 +126,68 @@ variable {g : α → α} theorem iterate_right (h : Commute f g) (n : ℕ) : Commute f (g^[n]) := Semiconj.iterate_right h n +#align function.commute.iterate_right Function.Commute.iterate_right theorem iterate_left (h : Commute f g) (n : ℕ) : Commute (f^[n]) g := (h.symm.iterate_right n).symm +#align function.commute.iterate_left Function.Commute.iterate_left theorem iterate_iterate (h : Commute f g) (m n : ℕ) : Commute (f^[m]) (g^[n]) := (h.iterate_left m).iterate_right n +#align function.commute.iterate_iterate Function.Commute.iterate_iterate theorem iterate_eq_of_map_eq (h : Commute f g) (n : ℕ) {x} (hx : f x = g x) : (f^[n]) x = (g^[n]) x := Nat.recOn n rfl fun n ihn ↦ by simp only [iterate_succ_apply, hx, (h.iterate_left n).eq, ihn, ((refl g).iterate_right n).eq] +#align function.commute.iterate_eq_of_map_eq Function.Commute.iterate_eq_of_map_eq theorem comp_iterate (h : Commute f g) (n : ℕ) : (f ∘ g)^[n] = f^[n] ∘ g^[n] := by induction' n with n ihn · rfl - funext x simp only [ihn, (h.iterate_right n).eq, iterate_succ, comp_apply] +#align function.commute.comp_iterate Function.Commute.comp_iterate variable (f) theorem iterate_self (n : ℕ) : Commute (f^[n]) f := (refl f).iterate_left n +#align function.commute.iterate_self Function.Commute.iterate_self theorem self_iterate (n : ℕ) : Commute f (f^[n]) := (refl f).iterate_right n +#align function.commute.self_iterate Function.Commute.self_iterate theorem iterate_iterate_self (m n : ℕ) : Commute (f^[m]) (f^[n]) := (refl f).iterate_iterate m n +#align function.commute.iterate_iterate_self Function.Commute.iterate_iterate_self end Commute theorem Semiconj₂.iterate {f : α → α} {op : α → α → α} (hf : Semiconj₂ f op op) (n : ℕ) : Semiconj₂ (f^[n]) op op := Nat.recOn n (Semiconj₂.id_left op) fun _ ihn ↦ ihn.comp hf +#align function.semiconj₂.iterate Function.Semiconj₂.iterate variable (f) theorem iterate_succ' (n : ℕ) : f^[n.succ] = f ∘ f^[n] := by rw [iterate_succ, (Commute.self_iterate f n).comp_eq] +#align function.iterate_succ' Function.iterate_succ' theorem iterate_succ_apply' (n : ℕ) (x : α) : (f^[n.succ]) x = f ((f^[n]) x) := by rw [iterate_succ'] rfl +#align function.iterate_succ_apply' Function.iterate_succ_apply' theorem iterate_pred_comp_of_pos {n : ℕ} (hn : 0 < n) : f^[n.pred] ∘ f = f^[n] := by rw [← iterate_succ, Nat.succ_pred_eq_of_pos hn] +#align function.iterate_pred_comp_of_pos Function.iterate_pred_comp_of_pos theorem comp_iterate_pred_of_pos {n : ℕ} (hn : 0 < n) : f ∘ f^[n.pred] = f^[n] := by rw [← iterate_succ', Nat.succ_pred_eq_of_pos hn] +#align function.comp_iterate_pred_of_pos Function.comp_iterate_pred_of_pos /-- A recursor for the iterate of a function. -/ def Iterate.rec (p : α → Sort _) {f : α → α} (h : ∀ a, p a → p (f a)) {a : α} (ha : p a) (n : ℕ) : @@ -170,10 +195,12 @@ def Iterate.rec (p : α → Sort _) {f : α → α} (h : ∀ a, p a → p (f a)) match n with | 0 => ha | m+1 => Iterate.rec p h (h _ ha) m +#align function.iterate.rec Function.Iterate.rec theorem Iterate.rec_zero (p : α → Sort _) {f : α → α} (h : ∀ a, p a → p (f a)) {a : α} (ha : p a) : Iterate.rec p h ha 0 = ha := rfl +#align function.iterate.rec_zero Function.Iterate.rec_zero variable {f} {m n : ℕ} {a : α} @@ -182,16 +209,20 @@ theorem LeftInverse.iterate {g : α → α} (hg : LeftInverse g f) (n : ℕ) : Nat.recOn n (fun _ ↦ rfl) fun n ihn ↦ by rw [iterate_succ', iterate_succ] exact ihn.comp hg +#align function.left_inverse.iterate Function.LeftInverse.iterate theorem RightInverse.iterate {g : α → α} (hg : RightInverse g f) (n : ℕ) : RightInverse (g^[n]) (f^[n]) := LeftInverse.iterate hg n +#align function.right_inverse.iterate Function.RightInverse.iterate theorem iterate_comm (f : α → α) (m n : ℕ) : f^[n]^[m] = f^[m]^[n] := (iterate_mul _ _ _).symm.trans (Eq.trans (by rw [Nat.mul_comm]) (iterate_mul _ _ _)) +#align function.iterate_comm Function.iterate_comm theorem iterate_commute (m n : ℕ) : Commute (fun f : α → α ↦ f^[m]) fun f ↦ f^[n] := fun f ↦ iterate_comm f m n +#align function.iterate_commute Function.iterate_commute lemma iterate_add_eq_iterate (hf : Injective f) : (f^[m + n]) a = (f^[n]) a ↔ (f^[m]) a = a := Iff.trans (by rw [←iterate_add_apply, Nat.add_comm]) (hf.iterate n).eq_iff @@ -217,9 +248,11 @@ theorem foldl_const (f : α → α) (a : α) (l : List β) : induction' l with b l H generalizing a · rfl · rw [length_cons, foldl, iterate_succ_apply, H] +#align list.foldl_const List.foldl_const theorem foldr_const (f : β → β) (b : β) : ∀ l : List α, l.foldr (fun _ ↦ f) b = (f^[l.length]) b | [] => rfl | a :: l => by rw [length_cons, foldr, foldr_const f b l, iterate_succ_apply'] +#align list.foldr_const List.foldr_const end List diff --git a/Mathlib/Logic/IsEmpty.lean b/Mathlib/Logic/IsEmpty.lean index c3f0be0e259f4..1588b5b670107 100644 --- a/Mathlib/Logic/IsEmpty.lean +++ b/Mathlib/Logic/IsEmpty.lean @@ -26,6 +26,7 @@ variable {α β γ : Sort _} /-- `IsEmpty α` expresses that `α` is empty. -/ class IsEmpty (α : Sort _) : Prop where protected false : α → False +#align is_empty IsEmpty instance : IsEmpty Empty := ⟨Empty.elim⟩ @@ -86,6 +87,7 @@ example [h : Nonempty α] [IsEmpty β] : IsEmpty (α → β) := by infer_instanc @[elab_as_elim] def isEmptyElim [IsEmpty α] {p : α → Sort _} (a : α) : p a := (IsEmpty.false a).elim +#align is_empty_elim isEmptyElim theorem isEmpty_iff : IsEmpty α ↔ α → False := ⟨@IsEmpty.false α, IsEmpty.mk⟩ @@ -99,24 +101,29 @@ open Function @[elab_as_elim] protected def elim {α : Sort u} (_ : IsEmpty α) {p : α → Sort _} (a : α) : p a := isEmptyElim a +#align is_empty.elim IsEmpty.elim /-- Non-dependent version of `IsEmpty.elim`. Helpful if the elaborator cannot elaborate `h.elim a` correctly. -/ protected def elim' {β : Sort _} (h : IsEmpty α) (a : α) : β := (h.false a).elim +#align is_empty.elim' IsEmpty.elim' protected theorem prop_iff {p : Prop} : IsEmpty p ↔ ¬p := isEmpty_iff +#align is_empty.prop_iff IsEmpty.prop_iff variable [IsEmpty α] @[simp] theorem forall_iff {p : α → Prop} : (∀ a, p a) ↔ True := iff_true_intro isEmptyElim +#align is_empty.forall_iff IsEmpty.forall_iff @[simp] theorem exists_iff {p : α → Prop} : (∃ a, p a) ↔ False := iff_false_intro fun ⟨x, _⟩ ↦ IsEmpty.false x +#align is_empty.exists_iff IsEmpty.exists_iff -- see Note [lower instance priority] instance (priority := 100) : Subsingleton α := @@ -127,6 +134,7 @@ end IsEmpty @[simp] theorem not_nonempty_iff : ¬Nonempty α ↔ IsEmpty α := ⟨fun h ↦ ⟨fun x ↦ h ⟨x⟩⟩, fun h1 h2 ↦ h2.elim h1.elim⟩ +#align not_nonempty_iff not_nonempty_iff @[simp] theorem not_isEmpty_iff : ¬IsEmpty α ↔ Nonempty α := diff --git a/Mathlib/Logic/Lemmas.lean b/Mathlib/Logic/Lemmas.lean index 4847e02e2a2ab..8e5281110bf2a 100644 --- a/Mathlib/Logic/Lemmas.lean +++ b/Mathlib/Logic/Lemmas.lean @@ -23,6 +23,8 @@ would result in less delta-reduced statements. alias heq_iff_eq ↔ HEq.eq Eq.heq +#align heq.eq HEq.eq +#align eq.heq Eq.heq -- Porting note: we may need to modify `alias` so we can make aliases protected at creation. -- attribute [protected] HEq.eq Eq.heq @@ -33,37 +35,47 @@ theorem dite_dite_distrib_left {a : p → α} {b : ¬p → q → α} {c : ¬p (dite p a fun hp ↦ dite q (b hp) (c hp)) = dite q (fun hq ↦ (dite p a) fun hp ↦ b hp hq) fun hq ↦ (dite p a) fun hp ↦ c hp hq := by split_ifs <;> rfl +#align dite_dite_distrib_left dite_dite_distrib_left theorem dite_dite_distrib_right {a : p → q → α} {b : p → ¬q → α} {c : ¬p → α} : dite p (fun hp ↦ dite q (a hp) (b hp)) c = dite q (fun hq ↦ dite p (fun hp ↦ a hp hq) c) fun hq ↦ dite p (fun hp ↦ b hp hq) c := by split_ifs <;> rfl +#align dite_dite_distrib_right dite_dite_distrib_right theorem ite_dite_distrib_left {a : α} {b : q → α} {c : ¬q → α} : ite p a (dite q b c) = dite q (fun hq ↦ ite p a <| b hq) fun hq ↦ ite p a <| c hq := dite_dite_distrib_left +#align ite_dite_distrib_left ite_dite_distrib_left theorem ite_dite_distrib_right {a : q → α} {b : ¬q → α} {c : α} : ite p (dite q a b) c = dite q (fun hq ↦ ite p (a hq) c) fun hq ↦ ite p (b hq) c := dite_dite_distrib_right +#align ite_dite_distrib_right ite_dite_distrib_right theorem dite_ite_distrib_left {a : p → α} {b : ¬p → α} {c : ¬p → α} : (dite p a fun hp ↦ ite q (b hp) (c hp)) = ite q (dite p a b) (dite p a c) := dite_dite_distrib_left +#align dite_ite_distrib_left dite_ite_distrib_left theorem dite_ite_distrib_right {a : p → α} {b : p → α} {c : ¬p → α} : dite p (fun hp ↦ ite q (a hp) (b hp)) c = ite q (dite p a c) (dite p b c) := dite_dite_distrib_right +#align dite_ite_distrib_right dite_ite_distrib_right theorem ite_ite_distrib_left : ite p a (ite q b c) = ite q (ite p a b) (ite p a c) := dite_dite_distrib_left +#align ite_ite_distrib_left ite_ite_distrib_left theorem ite_ite_distrib_right : ite p (ite q a b) c = ite q (ite p a c) (ite p b c) := dite_dite_distrib_right +#align ite_ite_distrib_right ite_ite_distrib_right lemma Prop.forall {f : Prop → Prop} : (∀ p, f p) ↔ f True ∧ f False := ⟨fun h ↦ ⟨h _, h _⟩, by rintro ⟨h₁, h₀⟩ p; by_cases hp : p <;> simp only [hp] <;> assumption⟩ +#align Prop.forall Prop.forall lemma Prop.exists {f : Prop → Prop} : (∃ p, f p) ↔ f True ∨ f False := ⟨fun ⟨p, h⟩ ↦ by refine' (em p).imp _ _ <;> intro H <;> convert h <;> simp [H], by rintro (h | h) <;> exact ⟨_, h⟩⟩ +#align Prop.exists Prop.exists diff --git a/Mathlib/Logic/Nonempty.lean b/Mathlib/Logic/Nonempty.lean index 23d2d497b3cff..9a0b28b68a434 100644 --- a/Mathlib/Logic/Nonempty.lean +++ b/Mathlib/Logic/Nonempty.lean @@ -34,33 +34,41 @@ instance (priority := 20) One.nonempty [One α] : Nonempty α := theorem exists_true_iff_nonempty {α : Sort _} : (∃ _ : α, True) ↔ Nonempty α := Iff.intro (fun ⟨a, _⟩ ↦ ⟨a⟩) fun ⟨a⟩ ↦ ⟨a, trivial⟩ +#align exists_true_iff_nonempty exists_true_iff_nonempty @[simp] theorem nonempty_Prop {p : Prop} : Nonempty p ↔ p := Iff.intro (fun ⟨h⟩ ↦ h) fun h ↦ ⟨h⟩ +#align nonempty_Prop nonempty_Prop theorem not_nonempty_iff_imp_false {α : Sort _} : ¬Nonempty α ↔ α → False := ⟨fun h a ↦ h ⟨a⟩, fun h ⟨a⟩ ↦ h a⟩ +#align not_nonempty_iff_imp_false not_nonempty_iff_imp_false @[simp] theorem nonempty_sigma : Nonempty (Σa : α, γ a) ↔ ∃ a : α, Nonempty (γ a) := Iff.intro (fun ⟨⟨a, c⟩⟩ ↦ ⟨a, ⟨c⟩⟩) fun ⟨a, ⟨c⟩⟩ ↦ ⟨⟨a, c⟩⟩ +#align nonempty_sigma nonempty_sigma @[simp] theorem nonempty_psigma {α} {β : α → Sort _} : Nonempty (PSigma β) ↔ ∃ a : α, Nonempty (β a) := Iff.intro (fun ⟨⟨a, c⟩⟩ ↦ ⟨a, ⟨c⟩⟩) fun ⟨a, ⟨c⟩⟩ ↦ ⟨⟨a, c⟩⟩ +#align nonempty_psigma nonempty_psigma @[simp] theorem nonempty_subtype {α} {p : α → Prop} : Nonempty (Subtype p) ↔ ∃ a : α, p a := Iff.intro (fun ⟨⟨a, h⟩⟩ ↦ ⟨a, h⟩) fun ⟨a, h⟩ ↦ ⟨⟨a, h⟩⟩ +#align nonempty_subtype nonempty_subtype @[simp] theorem nonempty_prod : Nonempty (α × β) ↔ Nonempty α ∧ Nonempty β := Iff.intro (fun ⟨⟨a, b⟩⟩ ↦ ⟨⟨a⟩, ⟨b⟩⟩) fun ⟨⟨a⟩, ⟨b⟩⟩ ↦ ⟨⟨a, b⟩⟩ +#align nonempty_prod nonempty_prod @[simp] theorem nonempty_pprod {α β} : Nonempty (PProd α β) ↔ Nonempty α ∧ Nonempty β := Iff.intro (fun ⟨⟨a, b⟩⟩ ↦ ⟨⟨a⟩, ⟨b⟩⟩) fun ⟨⟨a⟩, ⟨b⟩⟩ ↦ ⟨⟨a, b⟩⟩ +#align nonempty_pprod nonempty_pprod @[simp] theorem nonempty_sum : Nonempty (Sum α β) ↔ Nonempty α ∨ Nonempty β := @@ -73,6 +81,7 @@ theorem nonempty_sum : Nonempty (Sum α β) ↔ Nonempty α ∨ Nonempty β := match h with | Or.inl ⟨a⟩ => ⟨Sum.inl a⟩ | Or.inr ⟨b⟩ => ⟨Sum.inr b⟩ +#align nonempty_sum nonempty_sum @[simp] theorem nonempty_psum {α β} : Nonempty (PSum α β) ↔ Nonempty α ∨ Nonempty β := @@ -85,22 +94,27 @@ theorem nonempty_psum {α β} : Nonempty (PSum α β) ↔ Nonempty α ∨ Nonemp match h with | Or.inl ⟨a⟩ => ⟨PSum.inl a⟩ | Or.inr ⟨b⟩ => ⟨PSum.inr b⟩ +#align nonempty_psum nonempty_psum @[simp] theorem nonempty_ulift : Nonempty (ULift α) ↔ Nonempty α := Iff.intro (fun ⟨⟨a⟩⟩ ↦ ⟨a⟩) fun ⟨a⟩ ↦ ⟨⟨a⟩⟩ +#align nonempty_ulift nonempty_ulift @[simp] theorem nonempty_plift {α} : Nonempty (PLift α) ↔ Nonempty α := Iff.intro (fun ⟨⟨a⟩⟩ ↦ ⟨a⟩) fun ⟨a⟩ ↦ ⟨⟨a⟩⟩ +#align nonempty_plift nonempty_plift @[simp] theorem Nonempty.forall {α} {p : Nonempty α → Prop} : (∀ h : Nonempty α, p h) ↔ ∀ a, p ⟨a⟩ := Iff.intro (fun h _ ↦ h _) fun h ⟨a⟩ ↦ h a +#align nonempty.forall Nonempty.forall @[simp] theorem Nonempty.exists {α} {p : Nonempty α → Prop} : (∃ h : Nonempty α, p h) ↔ ∃ a, p ⟨a⟩ := Iff.intro (fun ⟨⟨a⟩, h⟩ ↦ ⟨a, h⟩) fun ⟨a, h⟩ ↦ ⟨⟨a⟩, h⟩ +#align nonempty.exists Nonempty.exists /-- Using `Classical.choice`, lifts a (`Prop`-valued) `Nonempty` instance to a (`Type`-valued) `Inhabited` instance. `Classical.inhabited_of_nonempty` already exists, in @@ -108,32 +122,39 @@ theorem Nonempty.exists {α} {p : Nonempty α → Prop} : (∃ h : Nonempty α, which makes it unsuitable for some applications. -/ noncomputable def Classical.inhabited_of_nonempty' {α} [h : Nonempty α] : Inhabited α := ⟨Classical.choice h⟩ +#align classical.inhabited_of_nonempty' Classical.inhabited_of_nonempty' /-- Using `Classical.choice`, extracts a term from a `Nonempty` type. -/ @[reducible] protected noncomputable def Nonempty.some {α} (h : Nonempty α) : α := Classical.choice h +#align nonempty.some Nonempty.some /-- Using `Classical.choice`, extracts a term from a `Nonempty` type. -/ @[reducible] protected noncomputable def Classical.arbitrary (α) [h : Nonempty α] : α := Classical.choice h +#align classical.arbitrary Classical.arbitrary /-- Given `f : α → β`, if `α` is nonempty then `β` is also nonempty. `Nonempty` cannot be a `functor`, because `Functor` is restricted to `Type`. -/ theorem Nonempty.map {α β} (f : α → β) : Nonempty α → Nonempty β | ⟨h⟩ => ⟨f h⟩ +#align nonempty.map Nonempty.map protected theorem Nonempty.map2 {α β γ : Sort _} (f : α → β → γ) : Nonempty α → Nonempty β → Nonempty γ | ⟨x⟩, ⟨y⟩ => ⟨f x y⟩ +#align nonempty.map2 Nonempty.map2 protected theorem Nonempty.congr {α β} (f : α → β) (g : β → α) : Nonempty α ↔ Nonempty β := ⟨Nonempty.map f, Nonempty.map g⟩ +#align nonempty.congr Nonempty.congr theorem Nonempty.elim_to_inhabited {α : Sort _} [h : Nonempty α] {p : Prop} (f : Inhabited α → p) : p := h.elim <| f ∘ Inhabited.mk +#align nonempty.elim_to_inhabited Nonempty.elim_to_inhabited protected instance Prod.Nonempty {α β} [h : Nonempty α] [h2 : Nonempty β] : Nonempty (α × β) := h.elim fun g ↦ h2.elim fun g2 ↦ ⟨⟨g, g2⟩⟩ @@ -144,12 +165,15 @@ protected instance Pi.Nonempty {ι : Sort _} {α : ι → Sort _} [∀ i, Nonemp theorem Classical.nonempty_pi {ι} {α : ι → Sort _} : Nonempty (∀ i, α i) ↔ ∀ i, Nonempty (α i) := ⟨fun ⟨f⟩ a ↦ ⟨f a⟩, @Pi.Nonempty _ _⟩ +#align classical.nonempty_pi Classical.nonempty_pi theorem subsingleton_of_not_nonempty {α : Sort _} (h : ¬Nonempty α) : Subsingleton α := ⟨fun x ↦ False.elim <| not_nonempty_iff_imp_false.mp h x⟩ +#align subsingleton_of_not_nonempty subsingleton_of_not_nonempty theorem Function.Surjective.nonempty [h : Nonempty β] {f : α → β} (hf : Function.Surjective f) : Nonempty α := let ⟨y⟩ := h let ⟨x, _⟩ := hf y ⟨x⟩ +#align function.surjective.nonempty Function.Surjective.nonempty diff --git a/Mathlib/Logic/Nontrivial.lean b/Mathlib/Logic/Nontrivial.lean index cdcc6c9990222..71f3f164fbf7a 100644 --- a/Mathlib/Logic/Nontrivial.lean +++ b/Mathlib/Logic/Nontrivial.lean @@ -33,12 +33,15 @@ this is equivalent to `0 ≠ 1`. In vector spaces, this is equivalent to positiv class Nontrivial (α : Type _) : Prop where /-- In a nontrivial type, there exists a pair of distinct terms. -/ exists_pair_ne : ∃ x y : α, x ≠ y +#align nontrivial Nontrivial theorem nontrivial_iff : Nontrivial α ↔ ∃ x y : α, x ≠ y := ⟨fun h ↦ h.exists_pair_ne, fun h ↦ ⟨h⟩⟩ +#align nontrivial_iff nontrivial_iff theorem exists_pair_ne (α : Type _) [Nontrivial α] : ∃ x y : α, x ≠ y := Nontrivial.exists_pair_ne +#align exists_pair_ne exists_pair_ne -- See Note [decidable namespace] protected theorem Decidable.exists_ne [Nontrivial α] [DecidableEq α] (x : α) : ∃ y, y ≠ x := by @@ -47,32 +50,40 @@ protected theorem Decidable.exists_ne [Nontrivial α] [DecidableEq α] (x : α) · rw [← hx] at h exact ⟨y', h.symm⟩ · exact ⟨y, Ne.symm hx⟩ +#align decidable.exists_ne Decidable.exists_ne theorem exists_ne [Nontrivial α] (x : α) : ∃ y, y ≠ x := by letI := Classical.decEq α; exact Decidable.exists_ne x +#align exists_ne exists_ne -- `x` and `y` are explicit here, as they are often needed to guide typechecking of `h`. theorem nontrivial_of_ne (x y : α) (h : x ≠ y) : Nontrivial α := ⟨⟨x, y, h⟩⟩ +#align nontrivial_of_ne nontrivial_of_ne -- `x` and `y` are explicit here, as they are often needed to guide typechecking of `h`. theorem nontrivial_of_lt [Preorder α] (x y : α) (h : x < y) : Nontrivial α := ⟨⟨x, y, ne_of_lt h⟩⟩ +#align nontrivial_of_lt nontrivial_of_lt theorem exists_pair_lt (α : Type _) [Nontrivial α] [LinearOrder α] : ∃ x y : α, x < y := by rcases exists_pair_ne α with ⟨x, y, hxy⟩ cases lt_or_gt_of_ne hxy <;> exact ⟨_, _, ‹_›⟩ +#align exists_pair_lt exists_pair_lt theorem nontrivial_iff_lt [LinearOrder α] : Nontrivial α ↔ ∃ x y : α, x < y := ⟨fun h ↦ @exists_pair_lt α h _, fun ⟨x, y, h⟩ ↦ nontrivial_of_lt x y h⟩ +#align nontrivial_iff_lt nontrivial_iff_lt theorem nontrivial_iff_exists_ne (x : α) : Nontrivial α ↔ ∃ y, y ≠ x := ⟨fun h ↦ @exists_ne α h x, fun ⟨_, hy⟩ ↦ nontrivial_of_ne _ _ hy⟩ +#align nontrivial_iff_exists_ne nontrivial_iff_exists_ne theorem Subtype.nontrivial_iff_exists_ne (p : α → Prop) (x : Subtype p) : Nontrivial (Subtype p) ↔ ∃ (y : α) (_ : p y), y ≠ x := by simp only [_root_.nontrivial_iff_exists_ne x, Subtype.exists, Ne.def, Subtype.ext_iff] +#align subtype.nontrivial_iff_exists_ne Subtype.nontrivial_iff_exists_ne instance : Nontrivial Prop := ⟨⟨True, False, true_ne_false⟩⟩ @@ -96,28 +107,35 @@ noncomputable def nontrivialPSumUnique (α : Type _) [Inhabited α] : uniq := fun x : α ↦ by by_contra H exact h ⟨_, _, H⟩ } +#align nontrivial_psum_unique nontrivialPSumUnique theorem subsingleton_iff : Subsingleton α ↔ ∀ x y : α, x = y := ⟨by intro h exact Subsingleton.elim, fun h ↦ ⟨h⟩⟩ +#align subsingleton_iff subsingleton_iff theorem not_nontrivial_iff_subsingleton : ¬Nontrivial α ↔ Subsingleton α := by simp only [nontrivial_iff, subsingleton_iff, not_exists, Ne.def, _root_.not_not] +#align not_nontrivial_iff_subsingleton not_nontrivial_iff_subsingleton theorem not_nontrivial (α) [Subsingleton α] : ¬Nontrivial α := fun ⟨⟨x, y, h⟩⟩ ↦ h <| Subsingleton.elim x y +#align not_nontrivial not_nontrivial theorem not_subsingleton (α) [Nontrivial α] : ¬Subsingleton α := fun _ => not_nontrivial _ ‹_› +#align not_subsingleton not_subsingleton /-- A type is either a subsingleton or nontrivial. -/ theorem subsingleton_or_nontrivial (α : Type _) : Subsingleton α ∨ Nontrivial α := by rw [← not_nontrivial_iff_subsingleton, or_comm] exact Classical.em _ +#align subsingleton_or_nontrivial subsingleton_or_nontrivial theorem false_of_nontrivial_of_subsingleton (α : Type _) [Nontrivial α] [Subsingleton α] : False := not_nontrivial _ ‹_› +#align false_of_nontrivial_of_subsingleton false_of_nontrivial_of_subsingleton instance Option.nontrivial [Nonempty α] : Nontrivial (Option α) := by inhabit α @@ -128,6 +146,7 @@ protected theorem Function.Injective.nontrivial [Nontrivial α] {f : α → β} (hf : Function.Injective f) : Nontrivial β := let ⟨x, y, h⟩ := exists_pair_ne α ⟨⟨f x, f y, hf.ne h⟩⟩ +#align function.injective.nontrivial Function.Injective.nontrivial /-- Pullback a `nontrivial` instance along a surjective function. -/ protected theorem Function.Surjective.nontrivial [Nontrivial β] {f : α → β} @@ -139,6 +158,7 @@ protected theorem Function.Surjective.nontrivial [Nontrivial β] {f : α → β} refine fun H ↦ h ?_ rw [← hx', ← hy', H] exact ⟨⟨x', y', this⟩⟩ +#align function.surjective.nontrivial Function.Surjective.nontrivial /-- An injective function from a nontrivial type has an argument at which it does not take a given value. -/ @@ -148,6 +168,7 @@ protected theorem Function.Injective.exists_ne [Nontrivial α] {f : α → β} by_cases h:f x₂ = y · exact ⟨x₁, (hf.ne_iff' h).2 hx⟩ · exact ⟨x₂, h⟩ +#align function.injective.exists_ne Function.Injective.exists_ne instance nontrivial_prod_right [Nonempty α] [Nontrivial β] : Nontrivial (α × β) := @@ -165,6 +186,7 @@ theorem nontrivial_at (i' : I) [inst : ∀ i, Nonempty (f i)] [Nontrivial (f i') Nontrivial (∀ i : I, f i) := by letI := Classical.decEq (∀ i : I, f i) exact (Function.update_injective (fun i ↦ Classical.choice (inst i)) i').nontrivial +#align pi.nontrivial_at Pi.nontrivial_at /-- As a convenience, provide an instance automatically if `(f default)` is nontrivial. @@ -181,6 +203,7 @@ instance Function.nontrivial [h : Nonempty α] [Nontrivial β] : Nontrivial (α protected theorem Subsingleton.le [Preorder α] [Subsingleton α] (x y : α) : x ≤ y := le_of_eq (Subsingleton.elim x y) +#align subsingleton.le Subsingleton.le namespace Bool diff --git a/Mathlib/Logic/Pairwise.lean b/Mathlib/Logic/Pairwise.lean index 26b2ff5632c76..38ed8e491de1f 100644 --- a/Mathlib/Logic/Pairwise.lean +++ b/Mathlib/Logic/Pairwise.lean @@ -50,6 +50,7 @@ theorem Function.injective_iff_pairwise_ne : Injective f ↔ Pairwise ((· ≠ #align function.injective_iff_pairwise_ne Function.injective_iff_pairwise_ne alias Function.injective_iff_pairwise_ne ↔ Function.Injective.pairwise_ne _ +#align function.injective.pairwise_ne Function.Injective.pairwise_ne namespace Set diff --git a/Mathlib/Logic/Relation.lean b/Mathlib/Logic/Relation.lean index 35453cf54cac7..629469db692ec 100644 --- a/Mathlib/Logic/Relation.lean +++ b/Mathlib/Logic/Relation.lean @@ -56,6 +56,7 @@ section NeImp variable {r : α → α → Prop} theorem IsRefl.reflexive [IsRefl α r] : Reflexive r := fun x ↦ IsRefl.refl x +#align is_refl.reflexive IsRefl.reflexive /-- To show a reflexive relation `r : α → α → Prop` holds over `x y : α`, it suffices to show it holds when `x ≠ y`. -/ @@ -63,32 +64,40 @@ theorem Reflexive.rel_of_ne_imp (h : Reflexive r) {x y : α} (hr : x ≠ y → r by_cases hxy : x = y · exact hxy ▸ h x · exact hr hxy +#align reflexive.rel_of_ne_imp Reflexive.rel_of_ne_imp /-- If a reflexive relation `r : α → α → Prop` holds over `x y : α`, then it holds whether or not `x ≠ y`. -/ theorem Reflexive.ne_imp_iff (h : Reflexive r) {x y : α} : x ≠ y → r x y ↔ r x y := ⟨h.rel_of_ne_imp, fun hr _ ↦ hr⟩ +#align reflexive.ne_imp_iff Reflexive.ne_imp_iff /-- If a reflexive relation `r : α → α → Prop` holds over `x y : α`, then it holds whether or not `x ≠ y`. Unlike `Reflexive.ne_imp_iff`, this uses `[IsRefl α r]`. -/ theorem reflexive_ne_imp_iff [IsRefl α r] {x y : α} : x ≠ y → r x y ↔ r x y := IsRefl.reflexive.ne_imp_iff +#align reflexive_ne_imp_iff reflexive_ne_imp_iff protected theorem Symmetric.iff (H : Symmetric r) (x y : α) : r x y ↔ r y x := ⟨fun h ↦ H h, fun h ↦ H h⟩ +#align symmetric.iff Symmetric.iff theorem Symmetric.flip_eq (h : Symmetric r) : flip r = r := funext₂ fun _ _ ↦ propext <| h.iff _ _ +#align symmetric.flip_eq Symmetric.flip_eq theorem Symmetric.swap_eq : Symmetric r → swap r = r := Symmetric.flip_eq +#align symmetric.swap_eq Symmetric.swap_eq theorem flip_eq_iff : flip r = r ↔ Symmetric r := ⟨fun h _ _ ↦ (congr_fun₂ h _ _).mp, Symmetric.flip_eq⟩ +#align flip_eq_iff flip_eq_iff theorem swap_eq_iff : swap r = r ↔ Symmetric r := flip_eq_iff +#align swap_eq_iff swap_eq_iff end NeImp @@ -97,14 +106,18 @@ section Comap variable {r : β → β → Prop} theorem Reflexive.comap (h : Reflexive r) (f : α → β) : Reflexive (r on f) := fun a ↦ h (f a) +#align reflexive.comap Reflexive.comap theorem Symmetric.comap (h : Symmetric r) (f : α → β) : Symmetric (r on f) := fun _ _ hab ↦ h hab +#align symmetric.comap Symmetric.comap theorem Transitive.comap (h : Transitive r) (f : α → β) : Transitive (r on f) := fun _ _ _ hab hbc ↦ h hab hbc +#align transitive.comap Transitive.comap theorem Equivalence.comap (h : Equivalence r) (f : α → β) : Equivalence (r on f) := ⟨h.reflexive.comap f, @(h.symmetric.comap f), @(h.transitive.comap f)⟩ +#align equivalence.comap Equivalence.comap end Comap @@ -120,6 +133,7 @@ term of `β` related to both. -/ def Comp (r : α → β → Prop) (p : β → γ → Prop) (a : α) (c : γ) : Prop := ∃ b, r a b ∧ p b c +#align relation.comp Relation.Comp @[inherit_doc] local infixr:80 " ∘r " => Relation.Comp @@ -127,18 +141,22 @@ local infixr:80 " ∘r " => Relation.Comp theorem comp_eq : r ∘r (· = ·) = r := funext fun _ ↦ funext fun b ↦ propext <| Iff.intro (fun ⟨_, h, Eq⟩ ↦ Eq ▸ h) fun h ↦ ⟨b, h, rfl⟩ +#align relation.comp_eq Relation.comp_eq theorem eq_comp : (· = ·) ∘r r = r := funext fun a ↦ funext fun _ ↦ propext <| Iff.intro (fun ⟨_, Eq, h⟩ ↦ Eq.symm ▸ h) fun h ↦ ⟨a, rfl, h⟩ +#align relation.eq_comp Relation.eq_comp theorem iff_comp {r : Prop → α → Prop} : (· ↔ ·) ∘r r = r := by have : (· ↔ ·) = (· = ·) := by funext a b; exact iff_eq_eq rw [this, eq_comp] +#align relation.iff_comp Relation.iff_comp theorem comp_iff {r : α → Prop → Prop} : r ∘r (· ↔ ·) = r := by have : (· ↔ ·) = (· = ·) := by funext a b; exact iff_eq_eq rw [this, comp_eq] +#align relation.comp_iff Relation.comp_iff theorem comp_assoc : (r ∘r p) ∘r q = r ∘r p ∘r q := by funext a d @@ -146,6 +164,7 @@ theorem comp_assoc : (r ∘r p) ∘r q = r ∘r p ∘r q := by constructor exact fun ⟨c, ⟨b, hab, hbc⟩, hcd⟩ ↦ ⟨b, hab, c, hbc, hcd⟩ exact fun ⟨b, hab, c, hbc, hcd⟩ ↦ ⟨c, ⟨b, hab, hbc⟩, hcd⟩ +#align relation.comp_assoc Relation.comp_assoc theorem flip_comp : flip (r ∘r p) = flip p ∘r flip r := by funext c a @@ -153,6 +172,7 @@ theorem flip_comp : flip (r ∘r p) = flip p ∘r flip r := by constructor exact fun ⟨b, hab, hbc⟩ ↦ ⟨b, hbc, hab⟩ exact fun ⟨b, hbc, hab⟩ ↦ ⟨b, hab, hbc⟩ +#align relation.flip_comp Relation.flip_comp end Comp @@ -165,6 +185,7 @@ variable (rα : α → α → Prop) (rβ : β → β → Prop) (f : α → β) of some `a' : α` under `f`, and `a'` and `a` are related by `rα`. -/ def Fibration := ∀ ⦃a b⦄, rβ b (f a) → ∃ a', rα a' a ∧ f a' = b +#align relation.fibration Relation.Fibration variable {rα rβ} @@ -175,6 +196,7 @@ theorem _root_.Acc.of_fibration (fib : Fibration rα rβ f) {a} (ha : Acc rα a) refine' Acc.intro (f a) fun b hr ↦ _ obtain ⟨a', hr', rfl⟩ := fib hr exact ih a' hr' +#align acc.of_fibration Acc.of_fibration theorem _root_.Acc.of_downward_closed (dc : ∀ {a b}, rβ b (f a) → ∃ c, f c = b) (a : α) (ha : Acc (InvImage rβ f) a) : Acc rβ (f a) := @@ -182,6 +204,7 @@ theorem _root_.Acc.of_downward_closed (dc : ∀ {a b}, rβ b (f a) → ∃ c, f let ⟨a', he⟩ := dc h -- porting note: Lean 3 did not need the motive ⟨a', he.substr (p := λ x => rβ x (f a)) h, he⟩ +#align acc.of_downward_closed Acc.of_downward_closed end Fibration @@ -192,6 +215,7 @@ related by `r`. -/ protected def Map (r : α → β → Prop) (f : α → γ) (g : β → δ) : γ → δ → Prop := fun c d ↦ ∃ a b, r a b ∧ f a = c ∧ g b = d +#align relation.map Relation.Map variable {r : α → α → Prop} {a b c d : α} @@ -200,6 +224,8 @@ variable {r : α → α → Prop} {a b c d : α} inductive ReflTransGen (r : α → α → Prop) (a : α) : α → Prop | refl : ReflTransGen r a a | tail {b c} : ReflTransGen r a b → r b c → ReflTransGen r a c +#align relation.refl_trans_gen Relation.ReflTransGen +#align relation.refl_trans_gen.cases_tail_iff Relation.ReflTransGen.cases_tail_iff attribute [refl] ReflTransGen.refl @@ -208,6 +234,7 @@ attribute [refl] ReflTransGen.refl inductive ReflGen (r : α → α → Prop) (a : α) : α → Prop | refl : ReflGen r a a | single {b} : r a b → ReflGen r a b +#align relation.refl_gen Relation.ReflGen #align relation.refl_gen_iff Relation.reflGen_iff @@ -216,6 +243,7 @@ inductive ReflGen (r : α → α → Prop) (a : α) : α → Prop inductive TransGen (r : α → α → Prop) (a : α) : α → Prop | single {b} : r a b → TransGen r a b | tail {b c} : TransGen r a b → r b c → TransGen r a c +#align relation.trans_gen Relation.TransGen #align relation.trans_gen_iff Relation.transGen_iff @@ -231,6 +259,7 @@ theorem to_reflTransGen : ∀ {a b}, ReflGen r a b → ReflTransGen r a b theorem mono {p : α → α → Prop} (hp : ∀ a b, r a b → p a b) : ∀ {a b}, ReflGen r a b → ReflGen p a b | a, _, ReflGen.refl => by rfl | a, b, single h => single (hp a b h) +#align relation.refl_gen.mono Relation.ReflGen.mono instance : IsRefl α (ReflGen r) := ⟨@refl α r⟩ @@ -244,23 +273,28 @@ theorem trans (hab : ReflTransGen r a b) (hbc : ReflTransGen r b c) : ReflTransG induction hbc case refl => assumption case tail c d _ hcd hac => exact hac.tail hcd +#align relation.refl_trans_gen.trans Relation.ReflTransGen.trans theorem single (hab : r a b) : ReflTransGen r a b := refl.tail hab +#align relation.refl_trans_gen.single Relation.ReflTransGen.single theorem head (hab : r a b) (hbc : ReflTransGen r b c) : ReflTransGen r a c := by induction hbc case refl => exact refl.tail hab case tail c d _ hcd hac => exact hac.tail hcd +#align relation.refl_trans_gen.head Relation.ReflTransGen.head theorem symmetric (h : Symmetric r) : Symmetric (ReflTransGen r) := by intro x y h induction' h with z w _ b c · rfl · apply Relation.ReflTransGen.head (h b) c +#align relation.refl_trans_gen.symmetric Relation.ReflTransGen.symmetric theorem cases_tail : ReflTransGen r a b → b = a ∨ ∃ c, ReflTransGen r a c ∧ r c b := (cases_tail_iff r a b).1 +#align relation.refl_trans_gen.cases_tail Relation.ReflTransGen.cases_tail @[elab_as_elim] theorem head_induction_on {P : ∀ a : α, ReflTransGen r a b → Prop} {a : α} (h : ReflTransGen r a b) @@ -273,6 +307,7 @@ theorem head_induction_on {P : ∀ a : α, ReflTransGen r a b → Prop} {a : α} refine @ih (λ {a : α} (hab : ReflTransGen r a b) => P a (ReflTransGen.tail hab hbc)) ?_ ?_ { exact head hbc _ refl } { exact fun h1 h2 ↦ head h1 (h2.tail hbc) } +#align relation.refl_trans_gen.head_induction_on Relation.ReflTransGen.head_induction_on @[elab_as_elim] theorem trans_induction_on {P : ∀ {a b : α}, ReflTransGen r a b → Prop} {a b : α} @@ -282,6 +317,7 @@ theorem trans_induction_on {P : ∀ {a b : α}, ReflTransGen r a b → Prop} {a induction h case refl => exact ih₁ a case tail b c hab hbc ih => exact ih₃ hab (single hbc) ih (ih₂ hbc) +#align relation.refl_trans_gen.trans_induction_on Relation.ReflTransGen.trans_induction_on theorem cases_head (h : ReflTransGen r a b) : a = b ∨ ∃ c, r a c ∧ ReflTransGen r c b := by induction h using Relation.ReflTransGen.head_induction_on @@ -289,12 +325,14 @@ theorem cases_head (h : ReflTransGen r a b) : a = b ∨ ∃ c, r a c ∧ ReflTra rfl · right exact ⟨_, by assumption, by assumption⟩; +#align relation.refl_trans_gen.cases_head Relation.ReflTransGen.cases_head theorem cases_head_iff : ReflTransGen r a b ↔ a = b ∨ ∃ c, r a c ∧ ReflTransGen r c b := by use cases_head rintro (rfl | ⟨c, hac, hcb⟩) · rfl · exact head hac hcb +#align relation.refl_trans_gen.cases_head_iff Relation.ReflTransGen.cases_head_iff theorem total_of_right_unique (U : Relator.RightUnique r) (ab : ReflTransGen r a b) (ac : ReflTransGen r a c) : ReflTransGen r b c ∨ ReflTransGen r c b := by @@ -306,6 +344,7 @@ theorem total_of_right_unique (U : Relator.RightUnique r) (ab : ReflTransGen r a · cases U bd be exact Or.inl ec · exact Or.inr (IH.tail bd) +#align relation.refl_trans_gen.total_of_right_unique Relation.ReflTransGen.total_of_right_unique end ReflTransGen @@ -322,6 +361,7 @@ theorem trans_left (hab : TransGen r a b) (hbc : ReflTransGen r b c) : TransGen induction hbc case refl => assumption case tail c d _ hcd hac => exact hac.tail hcd +#align relation.trans_gen.trans_left Relation.TransGen.trans_left instance : Trans (TransGen r) (ReflTransGen r) (TransGen r) := ⟨trans_left⟩ @@ -329,20 +369,24 @@ instance : Trans (TransGen r) (ReflTransGen r) (TransGen r) := @[trans] theorem trans (hab : TransGen r a b) (hbc : TransGen r b c) : TransGen r a c := trans_left hab hbc.to_reflTransGen +#align relation.trans_gen.trans Relation.TransGen.trans instance : Trans (TransGen r) (TransGen r) (TransGen r) := ⟨trans⟩ theorem head' (hab : r a b) (hbc : ReflTransGen r b c) : TransGen r a c := trans_left (single hab) hbc +#align relation.trans_gen.head' Relation.TransGen.head' theorem tail' (hab : ReflTransGen r a b) (hbc : r b c) : TransGen r a c := by induction hab generalizing c case refl => exact single hbc case tail _ _ _ hdb IH => exact tail (IH hdb) hbc +#align relation.trans_gen.tail' Relation.TransGen.tail' theorem head (hab : r a b) (hbc : TransGen r b c) : TransGen r a c := head' hab hbc.to_reflTransGen +#align relation.trans_gen.head Relation.TransGen.head @[elab_as_elim] theorem head_induction_on {P : ∀ a : α, TransGen r a b → Prop} {a : α} (h : TransGen r a b) @@ -355,6 +399,7 @@ theorem head_induction_on {P : ∀ a : α, TransGen r a b → Prop} {a : α} (h refine @h_ih (λ {a : α} (hab : @TransGen α r a b) => P a (TransGen.tail hab hbc)) ?_ ?_; exact fun h ↦ ih h (single hbc) (base hbc) exact fun hab hbc ↦ ih hab _ +#align relation.trans_gen.head_induction_on Relation.TransGen.head_induction_on @[elab_as_elim] theorem trans_induction_on {P : ∀ {a b : α}, TransGen r a b → Prop} {a b : α} (h : TransGen r a b) @@ -364,11 +409,13 @@ theorem trans_induction_on {P : ∀ {a b : α}, TransGen r a b → Prop} {a b : induction h with | single h => exact base h | tail hab hbc h_ih => exact ih hab (single hbc) h_ih (base hbc) +#align relation.trans_gen.trans_induction_on Relation.TransGen.trans_induction_on theorem trans_right (hab : ReflTransGen r a b) (hbc : TransGen r b c) : TransGen r a c := by induction hbc case single c hbc => exact tail' hab hbc case tail c d _ hcd hac => exact hac.tail hcd +#align relation.trans_gen.trans_right Relation.TransGen.trans_right instance : Trans (ReflTransGen r) (TransGen r) (TransGen r) := ⟨trans_right⟩ @@ -378,6 +425,7 @@ theorem tail'_iff : TransGen r a c ↔ ∃ b, ReflTransGen r a b ∧ r b c := by cases' h with _ hac b _ hab hbc · exact ⟨_, by rfl, hac⟩ · exact ⟨_, hab.to_reflTransGen, hbc⟩ +#align relation.trans_gen.tail'_iff Relation.TransGen.tail'_iff theorem head'_iff : TransGen r a c ↔ ∃ b, r a b ∧ ReflTransGen r b c := by refine' ⟨fun h ↦ _, fun ⟨b, hab, hbc⟩ ↦ head' hab hbc⟩ @@ -386,6 +434,7 @@ theorem head'_iff : TransGen r a c ↔ ∃ b, r a b ∧ ReflTransGen r b c := by case tail b c _ hbc IH => rcases IH with ⟨d, had, hdb⟩ exact ⟨_, had, hdb.tail hbc⟩ +#align relation.trans_gen.head'_iff Relation.TransGen.head'_iff end TransGen @@ -394,6 +443,7 @@ theorem _root_.Acc.TransGen (h : Acc r a) : Acc (TransGen r) a := by refine' Acc.intro x fun y hy ↦ _ cases' hy with _ hyx z _ hyz hzx exacts[H y hyx, (H z hzx).inv hyz] +#align acc.trans_gen Acc.TransGen theorem _root_.acc_transGen_iff : Acc (TransGen r) a ↔ Acc r a := ⟨Subrelation.accessible TransGen.single, Acc.TransGen⟩ @@ -429,24 +479,29 @@ theorem TransGen.lift {p : β → β → Prop} {a b : α} (f : α → β) (h : induction hab case single c hac => exact TransGen.single (h a c hac) case tail c d _ hcd hac => exact TransGen.tail hac (h c d hcd) +#align relation.trans_gen.lift Relation.TransGen.lift theorem TransGen.lift' {p : β → β → Prop} {a b : α} (f : α → β) (h : ∀ a b, r a b → TransGen p (f a) (f b)) (hab : TransGen r a b) : TransGen p (f a) (f b) := by simpa [transGen_idem] using hab.lift f h +#align relation.trans_gen.lift' Relation.TransGen.lift' theorem TransGen.closed {p : α → α → Prop} : (∀ a b, r a b → TransGen p a b) → TransGen r a b → TransGen p a b := TransGen.lift' id +#align relation.trans_gen.closed Relation.TransGen.closed theorem TransGen.mono {p : α → α → Prop} : (∀ a b, r a b → p a b) → TransGen r a b → TransGen p a b := TransGen.lift id +#align relation.trans_gen.mono Relation.TransGen.mono theorem TransGen.swap (h : TransGen r b a) : TransGen (swap r) a b := by induction' h with b h b c _ hbc ih · exact TransGen.single h · exact ih.head hbc +#align relation.trans_gen.swap Relation.TransGen.swap theorem transGen_swap : TransGen (swap r) a b ↔ TransGen r b a := ⟨TransGen.swap, TransGen.swap⟩ @@ -475,10 +530,12 @@ theorem reflTransGen_iff_eq_or_transGen : ReflTransGen r a b ↔ b = a ∨ Trans theorem ReflTransGen.lift {p : β → β → Prop} {a b : α} (f : α → β) (h : ∀ a b, r a b → p (f a) (f b)) (hab : ReflTransGen r a b) : ReflTransGen p (f a) (f b) := ReflTransGen.trans_induction_on hab (fun _ ↦ refl) (ReflTransGen.single ∘ h _ _) fun _ _ ↦ trans +#align relation.refl_trans_gen.lift Relation.ReflTransGen.lift theorem ReflTransGen.mono {p : α → α → Prop} : (∀ a b, r a b → p a b) → ReflTransGen r a b → ReflTransGen p a b := ReflTransGen.lift id +#align relation.refl_trans_gen.mono Relation.ReflTransGen.mono theorem reflTransGen_eq_self (refl : Reflexive r) (trans : Transitive r) : ReflTransGen r = r := funext fun a ↦ funext fun b ↦ propext <| @@ -502,11 +559,13 @@ instance : IsTrans α (ReflTransGen r) := theorem refl_trans_gen_idem : ReflTransGen (ReflTransGen r) = ReflTransGen r := reflTransGen_eq_self reflexive_reflTransGen transitive_reflTransGen +#align relation.refl_trans_gen_idem Relation.refl_trans_gen_idem theorem ReflTransGen.lift' {p : β → β → Prop} {a b : α} (f : α → β) (h : ∀ a b, r a b → ReflTransGen p (f a) (f b)) (hab : ReflTransGen r a b) : ReflTransGen p (f a) (f b) := by simpa [refl_trans_gen_idem] using hab.lift f h +#align relation.refl_trans_gen.lift' Relation.ReflTransGen.lift' theorem reflTransGen_closed {p : α → α → Prop} : (∀ a b, r a b → ReflTransGen p a b) → ReflTransGen r a b → ReflTransGen p a b := @@ -517,6 +576,7 @@ theorem ReflTransGen.swap (h : ReflTransGen r b a) : ReflTransGen (swap r) a b : induction' h with b c _ hbc ih · rfl · exact ih.head hbc +#align relation.refl_trans_gen.swap Relation.ReflTransGen.swap theorem reflTransGen_swap : ReflTransGen (swap r) a b ↔ ReflTransGen r b a := ⟨ReflTransGen.swap, ReflTransGen.swap⟩ @@ -532,6 +592,7 @@ in a term rewriting system, then *confluence* is the property that if (see `relation.church_rosser`). -/ def Join (r : α → α → Prop) : α → α → Prop := fun a b ↦ ∃ c, r a c ∧ r b c +#align relation.join Relation.Join section Join @@ -558,24 +619,30 @@ theorem church_rosser (h : ∀ a b c, r a b → r a c → ∃ d, ReflGen r b d cases' hba with _ hba · exact ⟨b, hea, hcb⟩ · exact ⟨a, hea, hcb.tail hba⟩ +#align relation.church_rosser Relation.church_rosser theorem join_of_single (h : Reflexive r) (hab : r a b) : Join r a b := ⟨b, hab, h b⟩ +#align relation.join_of_single Relation.join_of_single theorem symmetric_join : Symmetric (Join r) := fun _ _ ⟨c, hac, hcb⟩ ↦ ⟨c, hcb, hac⟩ +#align relation.symmetric_join Relation.symmetric_join theorem reflexive_join (h : Reflexive r) : Reflexive (Join r) := fun a ↦ ⟨a, h a, h a⟩ +#align relation.reflexive_join Relation.reflexive_join theorem transitive_join (ht : Transitive r) (h : ∀ a b c, r a b → r a c → Join r b c) : Transitive (Join r) := fun _ b _ ⟨x, hax, hbx⟩ ⟨y, hby, hcy⟩ ↦ let ⟨z, hxz, hyz⟩ := h b x y hbx hby ⟨z, ht hax hxz, ht hcy hyz⟩ +#align relation.transitive_join Relation.transitive_join theorem equivalence_join (hr : Reflexive r) (ht : Transitive r) (h : ∀ a b c, r a b → r a c → Join r b c) : Equivalence (Join r) := ⟨reflexive_join hr, @symmetric_join _ _, @transitive_join _ _ ht h⟩ +#align relation.equivalence_join Relation.equivalence_join theorem equivalence_join_reflTransGen (h : ∀ a b c, r a b → r a c → ∃ d, ReflGen r b d ∧ ReflTransGen r c d) : @@ -586,6 +653,7 @@ theorem equivalence_join_reflTransGen theorem join_of_equivalence {r' : α → α → Prop} (hr : Equivalence r) (h : ∀ a b, r' a b → r a b) : Join r' a b → r a b | ⟨_, hac, hbc⟩ => hr.trans (h _ _ hac) (hr.symm <| h _ _ hbc) +#align relation.join_of_equivalence Relation.join_of_equivalence theorem reflTransGen_of_transitive_reflexive {r' : α → α → Prop} (hr : Reflexive r) (ht : Transitive r) (h : ∀ a b, r' a b → r a b) (h' : ReflTransGen r' a b) : r a b := by @@ -630,5 +698,6 @@ theorem EqvGen.mono {r p : α → α → Prop} (hrp : ∀ a b, r a b → p a b) | refl => exact EqvGen.refl _ | symm a b _ ih => exact EqvGen.symm _ _ ih | trans a b c _ _ hab hbc => exact EqvGen.trans _ _ _ hab hbc +#align eqv_gen.mono EqvGen.mono end EqvGen diff --git a/Mathlib/Logic/Relator.lean b/Mathlib/Logic/Relator.lean index 8b08b9922905f..e7ccaeb1835bc 100644 --- a/Mathlib/Logic/Relator.lean +++ b/Mathlib/Logic/Relator.lean @@ -35,6 +35,7 @@ variable (R : α → β → Prop) (S : γ → δ → Prop) relation on functions `LiftFun : (f : α → γ) (g : β → δ) : Prop'. -/ def LiftFun (f : α → γ) (g : β → δ) : Prop := ∀⦃a b⦄, R a b → S (f a) (g b) +#align relator.lift_fun Relator.LiftFun /-- `(R ⇒ S) f g` means `LiftFun R S f g`. -/ infixr:40 " ⇒ " => LiftFun @@ -47,77 +48,96 @@ variable {α : Type u₁} {β : Type u₂} (R : α → β → Prop) /-- A relation is "right total" if every element appears on the right. -/ def RightTotal : Prop := ∀ b, ∃ a, R a b +#align relator.right_total Relator.RightTotal /-- A relation is "left total" if every element appears on the left. -/ def LeftTotal : Prop := ∀ a, ∃ b, R a b +#align relator.left_total Relator.LeftTotal /-- A relation is "bi-total" if it is both right total and left total. -/ def BiTotal : Prop := LeftTotal R ∧ RightTotal R +#align relator.bi_total Relator.BiTotal /-- A relation is "left unique" if every element on the right is paired with at most one element on the left. -/ def LeftUnique : Prop := ∀ ⦃a b c⦄, R a c → R b c → a = b +#align relator.left_unique Relator.LeftUnique /-- A relation is "right unique" if every element on the left is paired with at most one element on the right. -/ def RightUnique : Prop := ∀ ⦃a b c⦄, R a b → R a c → b = c +#align relator.right_unique Relator.RightUnique /-- A relation is "bi-unique" if it is both left unique and right unique. -/ def BiUnique : Prop := LeftUnique R ∧ RightUnique R +#align relator.bi_unique Relator.BiUnique variable {R} lemma RightTotal.rel_forall (h : RightTotal R) : ((R ⇒ (· → ·)) ⇒ (· → ·)) (λ p => ∀i, p i) (λ q => ∀i, q i) := λ _ _ Hrel H b => Exists.elim (h b) (λ _ Rab => Hrel Rab (H _)) +#align relator.right_total.rel_forall Relator.RightTotal.rel_forall lemma LeftTotal.rel_exists (h : LeftTotal R) : ((R ⇒ (· → ·)) ⇒ (· → ·)) (λ p => ∃i, p i) (λ q => ∃i, q i) := λ _ _ Hrel ⟨a, pa⟩ => (h a).imp $ λ _ Rab => Hrel Rab pa +#align relator.left_total.rel_exists Relator.LeftTotal.rel_exists lemma BiTotal.rel_forall (h : BiTotal R) : ((R ⇒ Iff) ⇒ Iff) (λ p => ∀i, p i) (λ q => ∀i, q i) := λ _ _ Hrel => ⟨λ H b => Exists.elim (h.right b) (λ _ Rab => (Hrel Rab).mp (H _)), λ H a => Exists.elim (h.left a) (λ _ Rab => (Hrel Rab).mpr (H _))⟩ +#align relator.bi_total.rel_forall Relator.BiTotal.rel_forall lemma BiTotal.rel_exists (h : BiTotal R) : ((R ⇒ Iff) ⇒ Iff) (λ p => ∃i, p i) (λ q => ∃i, q i) := λ _ _ Hrel => ⟨λ ⟨a, pa⟩ => (h.left a).imp $ λ _ Rab => (Hrel Rab).1 pa, λ ⟨b, qb⟩ => (h.right b).imp $ λ _ Rab => (Hrel Rab).2 qb⟩ +#align relator.bi_total.rel_exists Relator.BiTotal.rel_exists lemma left_unique_of_rel_eq {eq' : β → β → Prop} (he : (R ⇒ (R ⇒ Iff)) Eq eq') : LeftUnique R := λ a b c (ac : R a c) (bc : R b c) => (he ac bc).mpr ((he bc bc).mp rfl) +#align relator.left_unique_of_rel_eq Relator.left_unique_of_rel_eq end lemma rel_imp : (Iff ⇒ (Iff ⇒ Iff)) (· → ·) (· → ·) := λ _ _ h _ _ l => imp_congr h l +#align relator.rel_imp Relator.rel_imp lemma rel_not : (Iff ⇒ Iff) Not Not := λ _ _ h => not_congr h +#align relator.rel_not Relator.rel_not lemma bi_total_eq {α : Type u₁} : Relator.BiTotal (@Eq α) := { left := λ a => ⟨a, rfl⟩, right := λ a => ⟨a, rfl⟩ } +#align relator.bi_total_eq Relator.bi_total_eq variable {α : Type _} {β : Type _} {γ : Type _} {δ : Type _} variable {r : α → β → Prop} {p : β → γ → Prop} {q : γ → δ → Prop} lemma LeftUnique.flip (h : LeftUnique r) : RightUnique (flip r) := λ _ _ _ h₁ h₂ => h h₁ h₂ +#align relator.left_unique.flip Relator.LeftUnique.flip lemma rel_and : ((·↔·) ⇒ (·↔·) ⇒ (·↔·)) (·∧·) (·∧·) := λ _ _ h₁ _ _ h₂ => and_congr h₁ h₂ +#align relator.rel_and Relator.rel_and lemma rel_or : ((·↔·) ⇒ (·↔·) ⇒ (·↔·)) (·∨·) (·∨·) := λ _ _ h₁ _ _ h₂ => or_congr h₁ h₂ +#align relator.rel_or Relator.rel_or lemma rel_iff : ((·↔·) ⇒ (·↔·) ⇒ (·↔·)) (·↔·) (·↔·) := λ _ _ h₁ _ _ h₂ => iff_congr h₁ h₂ +#align relator.rel_iff Relator.rel_iff lemma rel_eq {r : α → β → Prop} (hr : BiUnique r) : (r ⇒ r ⇒ (·↔·)) (·=·) (·=·) := λ _ _ h₁ _ _ h₂ => ⟨λ h => hr.right h₁ $ h.symm ▸ h₂, λ h => hr.left h₁ $ h.symm ▸ h₂⟩ +#align relator.rel_eq Relator.rel_eq end Relator diff --git a/Mathlib/Logic/Unique.lean b/Mathlib/Logic/Unique.lean index cec86d6763cf2..b60f887994afc 100644 --- a/Mathlib/Logic/Unique.lean +++ b/Mathlib/Logic/Unique.lean @@ -57,6 +57,9 @@ for good definitional properties of the default term. -/ structure Unique (α : Sort u) extends Inhabited α where /-- In a `Unique` type, every term is equal to the default element (from `Inhabited`). -/ uniq : ∀ a : α, a = default +#align unique Unique +#align unique.ext_iff Unique.ext_iff +#align unique.ext Unique.ext attribute [class] Unique -- The simplifier can already prove this using `eq_iff_true_of_subsingleton` @@ -65,6 +68,7 @@ attribute [nolint simpNF] Unique.mk.injEq theorem unique_iff_exists_unique (α : Sort u) : Nonempty (Unique α) ↔ ∃! _ : α, True := ⟨fun ⟨u⟩ ↦ ⟨u.default, trivial, fun a _ ↦ u.uniq a⟩, fun ⟨a, _, h⟩ ↦ ⟨⟨⟨a⟩, fun _ ↦ h _ trivial⟩⟩⟩ +#align unique_iff_exists_unique unique_iff_exists_unique theorem unique_subtype_iff_exists_unique {α} (p : α → Prop) : Nonempty (Unique (Subtype p)) ↔ ∃! a, p a := @@ -72,6 +76,7 @@ theorem unique_subtype_iff_exists_unique {α} (p : α → Prop) : fun ⟨a, ha, he⟩ ↦ ⟨⟨⟨⟨a, ha⟩⟩, fun ⟨b, hb⟩ ↦ by congr exact he b hb⟩⟩⟩ +#align unique_subtype_iff_exists_unique unique_subtype_iff_exists_unique /-- Given an explicit `a : α` with `Subsingleton α`, we can construct a `Unique α` instance. This is a def because the typeclass search cannot @@ -83,6 +88,7 @@ See note [reducible non-instances]. -/ def uniqueOfSubsingleton {α : Sort _} [Subsingleton α] (a : α) : Unique α where default := a uniq _ := Subsingleton.elim _ _ +#align unique_of_subsingleton uniqueOfSubsingleton instance PUnit.unique : Unique PUnit.{u} where default := PUnit.unit @@ -102,12 +108,14 @@ theorem PUnit.default_eq_unit : (default : PUnit) = PUnit.unit := def uniqueProp {p : Prop} (h : p) : Unique.{0} p where default := h uniq _ := rfl +#align unique_prop uniqueProp instance : Unique True := uniqueProp trivial theorem Fin.eq_zero : ∀ n : Fin 1, n = 0 | ⟨_, hn⟩ => Fin.eq_of_veq (Nat.eq_zero_of_le_zero (Nat.le_of_lt_succ hn)) +#align fin.eq_zero Fin.eq_zero instance {n : ℕ} : Inhabited (Fin n.succ) := ⟨0⟩ @@ -118,6 +126,7 @@ instance inhabitedFinOneAdd (n : ℕ) : Inhabited (Fin (1 + n)) := @[simp] theorem Fin.default_eq_zero (n : ℕ) : (default : Fin n.succ) = 0 := rfl +#align fin.default_eq_zero Fin.default_eq_zero instance Fin.unique : Unique (Fin 1) where uniq := Fin.eq_zero @@ -136,9 +145,11 @@ instance (priority := 100) : Inhabited α := theorem eq_default (a : α) : a = default := uniq _ a +#align unique.eq_default Unique.eq_default theorem default_eq (a : α) : default = a := (uniq _ a).symm +#align unique.default_eq Unique.default_eq -- see Note [lower instance priority] instance (priority := 100) : Subsingleton α := @@ -146,15 +157,18 @@ instance (priority := 100) : Subsingleton α := theorem forall_iff {p : α → Prop} : (∀ a, p a) ↔ p default := ⟨fun h ↦ h _, fun h x ↦ by rwa [Unique.eq_default x]⟩ +#align unique.forall_iff Unique.forall_iff theorem exists_iff {p : α → Prop} : Exists p ↔ p default := ⟨fun ⟨a, ha⟩ ↦ eq_default a ▸ ha, Exists.intro default⟩ +#align unique.exists_iff Unique.exists_iff end @[ext] protected theorem subsingleton_unique' : ∀ h₁ h₂ : Unique α, h₁ = h₂ | ⟨⟨x⟩, h⟩, ⟨⟨y⟩, _⟩ => by congr; rw [h x, h y] +#align unique.subsingleton_unique' Unique.subsingleton_unique' instance subsingleton_unique : Subsingleton (Unique α) := ⟨Unique.subsingleton_unique'⟩ @@ -164,6 +178,7 @@ a loop in the class inheritance graph. -/ @[reducible] def mk' (α : Sort u) [h₁ : Inhabited α] [Subsingleton α] : Unique α := { h₁ with uniq := fun _ ↦ Subsingleton.elim _ _ } +#align unique.mk' Unique.mk' end Unique @@ -171,15 +186,18 @@ theorem unique_iff_subsingleton_and_nonempty (α : Sort u) : Nonempty (Unique α) ↔ Subsingleton α ∧ Nonempty α := ⟨fun ⟨u⟩ ↦ by constructor <;> exact inferInstance, fun ⟨hs, hn⟩ ↦ ⟨by inhabit α; exact Unique.mk' α⟩⟩ +#align unique_iff_subsingleton_and_nonempty unique_iff_subsingleton_and_nonempty @[simp] theorem Pi.default_def {β : α → Sort v} [∀ a, Inhabited (β a)] : @default (∀ a, β a) _ = fun a : α ↦ @default (β a) _ := rfl +#align pi.default_def Pi.default_def theorem Pi.default_apply {β : α → Sort v} [∀ a, Inhabited (β a)] (a : α) : @default (∀ a, β a) _ a = default := rfl +#align pi.default_apply Pi.default_apply instance Pi.unique {β : α → Sort v} [∀ a, Unique (β a)] : Unique (∀ a, β a) where uniq := fun _ ↦ funext fun _ ↦ Unique.eq_default _ @@ -193,10 +211,12 @@ theorem eq_const_of_unique [Unique α] (f : α → β) : f = Function.const α ( ext x rw [Subsingleton.elim x default] rfl +#align eq_const_of_unique eq_const_of_unique theorem heq_const_of_unique [Unique α] {β : α → Sort v} (f : ∀ a, β a) : HEq f (Function.const α (f default)) := (Function.hfunext rfl) fun i _ _ ↦ by rw [Subsingleton.elim i default]; rfl +#align heq_const_of_unique heq_const_of_unique namespace Function @@ -206,25 +226,30 @@ variable {f : α → β} is a subsingleton as well. -/ protected theorem Injective.subsingleton (hf : Injective f) [Subsingleton β] : Subsingleton α := ⟨fun _ _ ↦ hf <| Subsingleton.elim _ _⟩ +#align function.injective.subsingleton Function.Injective.subsingleton /-- If the domain of a surjective function is a subsingleton, then the codomain is a subsingleton as well. -/ protected theorem Surjective.subsingleton [Subsingleton α] (hf : Surjective f) : Subsingleton β := ⟨hf.forall₂.2 fun x y ↦ congr_arg f <| Subsingleton.elim x y⟩ +#align function.surjective.subsingleton Function.Surjective.subsingleton /-- If the domain of a surjective function is a singleton, then the codomain is a singleton as well. -/ protected def Surjective.unique (f : α → β) (hf : Surjective f) [Unique.{u} α] : Unique β := @Unique.mk' _ ⟨f default⟩ hf.subsingleton +#align function.surjective.unique Function.Surjective.unique /-- If `α` is inhabited and admits an injective map to a subsingleton type, then `α` is `Unique`. -/ protected def Injective.unique [Inhabited α] [Subsingleton β] (hf : Injective f) : Unique α := @Unique.mk' _ _ hf.subsingleton +#align function.injective.unique Function.Injective.unique /-- If a constant function is surjective, then the codomain is a singleton. -/ def Surjective.uniqueOfSurjectiveConst (α : Type _) {β : Type _} (b : β) (h : Function.Surjective (Function.const α b)) : Unique β := @uniqueOfSubsingleton _ (subsingleton_of_forall_eq b <| h.forall.mpr fun _ ↦ rfl) b +#align function.surjective.unique_of_surjective_const Function.Surjective.uniqueOfSurjectiveConst end Function @@ -233,6 +258,7 @@ attribute [simp] eq_iff_true_of_subsingleton in theorem Unique.bijective {A B} [Unique A] [Unique B] {f : A → B} : Function.Bijective f := by rw [Function.bijective_iff_has_inverse] refine' ⟨default, _, _⟩ <;> intro x <;> simp +#align unique.bijective Unique.bijective namespace Option diff --git a/Mathlib/Order/Antisymmetrization.lean b/Mathlib/Order/Antisymmetrization.lean index b295cfe16153e..40dd73d19207f 100644 --- a/Mathlib/Order/Antisymmetrization.lean +++ b/Mathlib/Order/Antisymmetrization.lean @@ -75,6 +75,7 @@ theorem antisymmRel_iff_eq [IsRefl α r] [IsAntisymm α r] {a b : α} : Antisymm #align antisymm_rel_iff_eq antisymmRel_iff_eq alias antisymmRel_iff_eq ↔ AntisymmRel.eq _ +#align antisymm_rel.eq AntisymmRel.eq end Relation @@ -87,6 +88,7 @@ variable (α) (r : α → α → Prop) [IsPreorder α r] def AntisymmRel.setoid : Setoid α := ⟨AntisymmRel r, antisymmRel_refl _, AntisymmRel.symm, AntisymmRel.trans⟩ #align antisymm_rel.setoid AntisymmRel.setoid +#align antisymm_rel.setoid_r AntisymmRel.setoid_r /-- The partial order derived from a preorder by making pairwise comparable elements equal. This is the quotient by `fun a b => a ≤ b ∧ b ≤ a`. -/ @@ -254,6 +256,7 @@ noncomputable def OrderEmbedding.ofAntisymmetrization : Antisymmetrization α ( inj' _ _ := Quotient.out_inj.1 map_rel_iff' := ofAntisymmetrization_le_ofAntisymmetrization_iff #align order_embedding.of_antisymmetrization OrderEmbedding.ofAntisymmetrization +#align order_embedding.of_antisymmetrization_apply OrderEmbedding.ofAntisymmetrization_apply /-- `Antisymmetrization` and `orderDual` commute. -/ def OrderIso.dualAntisymmetrization : diff --git a/Mathlib/Order/Atoms.lean b/Mathlib/Order/Atoms.lean index 0e1e2a7f5dae1..49c59e13b46cb 100644 --- a/Mathlib/Order/Atoms.lean +++ b/Mathlib/Order/Atoms.lean @@ -102,6 +102,8 @@ theorem bot_covby_iff : ⊥ ⋖ a ↔ IsAtom a := by #align bot_covby_iff bot_covby_iff alias bot_covby_iff ↔ Covby.is_atom IsAtom.bot_covby +#align covby.is_atom Covby.is_atom +#align is_atom.bot_covby IsAtom.bot_covby end IsAtom @@ -130,8 +132,10 @@ theorem isAtom_dual_iff_isCoatom [OrderTop α] {a : α} : #align is_atom_dual_iff_is_coatom isAtom_dual_iff_isCoatom alias isCoatom_dual_iff_isAtom ↔ _ IsAtom.dual +#align is_atom.dual IsAtom.dual alias isAtom_dual_iff_isCoatom ↔ _ IsCoatom.dual +#align is_coatom.dual IsCoatom.dual variable [OrderTop α] {a x : α} @@ -169,6 +173,8 @@ theorem covby_top_iff : a ⋖ ⊤ ↔ IsCoatom a := #align covby_top_iff covby_top_iff alias covby_top_iff ↔ Covby.is_coatom IsCoatom.covby_top +#align covby.is_coatom Covby.is_coatom +#align is_coatom.covby_top IsCoatom.covby_top end IsCoatom @@ -229,6 +235,7 @@ class IsAtomic [OrderBot α] : Prop where /--Every element other than `⊥` has an atom below it. -/ eq_bot_or_exists_atom_le : ∀ b : α, b = ⊥ ∨ ∃ a : α, IsAtom a ∧ a ≤ b #align is_atomic IsAtomic +#align is_atomic_iff IsAtomic_iff /-- A lattice is coatomic iff every element other than `⊤` has a coatom above it. -/ @[mk_iff] @@ -236,6 +243,7 @@ class IsCoatomic [OrderTop α] : Prop where /--Every element other than `⊤` has an atom above it. -/ eq_top_or_exists_le_coatom : ∀ b : α, b = ⊤ ∨ ∃ a : α, IsCoatom a ∧ b ≤ a #align is_coatomic IsCoatomic +#align is_coatomic_iff IsCoatomic_iff export IsAtomic (eq_bot_or_exists_atom_le) @@ -554,6 +562,8 @@ def equivBool {α} [DecidableEq α] [LE α] [BoundedOrder α] [IsSimpleOrder α] left_inv x := by rcases eq_bot_or_eq_top x with (rfl | rfl) <;> simp [bot_ne_top] right_inv x := by cases x <;> simp [bot_ne_top] #align is_simple_order.equiv_bool IsSimpleOrder.equivBool +#align is_simple_order.equiv_bool_symm_apply IsSimpleOrder.equivBool_symm_apply +#align is_simple_order.equiv_bool_apply IsSimpleOrder.equivBool_apply /-- Every simple lattice over a partial order is order-isomorphic to `Bool`. -/ def orderIsoBool : α ≃o Bool := diff --git a/Mathlib/Order/Basic.lean b/Mathlib/Order/Basic.lean index 5ffd6bc687d84..48a19def4a2cb 100644 --- a/Mathlib/Order/Basic.lean +++ b/Mathlib/Order/Basic.lean @@ -157,6 +157,7 @@ alias ne_of_lt ← LT.lt.ne alias lt_asymm ← LT.lt.asymm LT.lt.not_lt alias le_of_eq ← Eq.le +#align eq.le Eq.le -- Porting note: no `decidable_classical` linter -- attribute [nolint decidable_classical] LE.le.lt_or_eq_dec @@ -216,12 +217,16 @@ alias lt_of_lt_of_eq ← LT.lt.trans_eq alias lt_of_lt_of_eq' ← LT.lt.trans_eq' alias le_of_eq_of_le ← Eq.trans_le +#align eq.trans_le Eq.trans_le alias le_of_eq_of_le' ← Eq.trans_ge +#align eq.trans_ge Eq.trans_ge alias lt_of_eq_of_lt ← Eq.trans_lt +#align eq.trans_lt Eq.trans_lt alias lt_of_eq_of_lt' ← Eq.trans_gt +#align eq.trans_gt Eq.trans_gt end @@ -824,10 +829,13 @@ theorem strongLT_of_le_of_strongLT (hab : a ≤ b) (hbc : b ≺ c) : a ≺ c := #align strong_lt_of_le_of_strong_lt strongLT_of_le_of_strongLT alias le_of_strongLT ← StrongLT.le +#align strong_lt.le StrongLT.le alias lt_of_strongLT ← StrongLT.lt +#align strong_lt.lt StrongLT.lt alias strongLT_of_strongLT_of_le ← StrongLT.trans_le +#align strong_lt.trans_le StrongLT.trans_le alias strongLT_of_le_of_strongLT ← LE.le.trans_strongLT diff --git a/Mathlib/Order/BooleanAlgebra.lean b/Mathlib/Order/BooleanAlgebra.lean index da962609d1142..6222271aabe3a 100644 --- a/Mathlib/Order/BooleanAlgebra.lean +++ b/Mathlib/Order/BooleanAlgebra.lean @@ -223,9 +223,11 @@ theorem disjoint_sdiff_self_right : Disjoint x (y \ x) := lemma le_sdiff : x ≤ y \ z ↔ x ≤ y ∧ Disjoint x z := ⟨fun h ↦ ⟨h.trans sdiff_le, disjoint_sdiff_self_left.mono_left h⟩, fun h ↦ by rw [←h.2.sdiff_eq_left]; exact sdiff_le_sdiff_right h.1⟩ +#align le_sdiff le_sdiff @[simp] lemma sdiff_eq_left : x \ y = x ↔ Disjoint x y := ⟨fun h ↦ disjoint_sdiff_self_left.mono_left h.ge, Disjoint.sdiff_eq_left⟩ +#align sdiff_eq_left sdiff_eq_left /- TODO: we could make an alternative constructor for `GeneralizedBooleanAlgebra` using `Disjoint x (y \ x)` and `x ⊔ (y \ x) = y` as axioms. -/ diff --git a/Mathlib/Order/BoundedOrder.lean b/Mathlib/Order/BoundedOrder.lean index 45eef8aef3101..95c48cffe6ebb 100644 --- a/Mathlib/Order/BoundedOrder.lean +++ b/Mathlib/Order/BoundedOrder.lean @@ -155,8 +155,10 @@ theorem not_isTop_iff_ne_top : ¬IsTop a ↔ a ≠ ⊤ := #align not_is_top_iff_ne_top not_isTop_iff_ne_top alias isMax_iff_eq_top ↔ IsMax.eq_top _ +#align is_max.eq_top IsMax.eq_top alias isTop_iff_eq_top ↔ IsTop.eq_top _ +#align is_top.eq_top IsTop.eq_top @[simp] theorem top_le_iff : ⊤ ≤ a ↔ a = ⊤ := @@ -361,8 +363,10 @@ theorem not_isBot_iff_ne_bot : ¬IsBot a ↔ a ≠ ⊥ := #align not_is_bot_iff_ne_bot not_isBot_iff_ne_bot alias isMin_iff_eq_bot ↔ IsMin.eq_bot _ +#align is_min.eq_bot IsMin.eq_bot alias isBot_iff_eq_bot ↔ IsBot.eq_bot _ +#align is_bot.eq_bot IsBot.eq_bot @[simp] theorem le_bot_iff : a ≤ ⊥ ↔ a = ⊥ := diff --git a/Mathlib/Order/Closure.lean b/Mathlib/Order/Closure.lean index 2ad246a7d96ca..81751a0e62c9d 100644 --- a/Mathlib/Order/Closure.lean +++ b/Mathlib/Order/Closure.lean @@ -90,6 +90,7 @@ def id : ClosureOperator α where le_closure' _ := le_rfl idempotent' _ := rfl #align closure_operator.id ClosureOperator.id +#align closure_operator.id_apply ClosureOperator.id_apply instance : Inhabited (ClosureOperator α) := ⟨id α⟩ @@ -111,6 +112,7 @@ def mk' (f : α → α) (hf₁ : Monotone f) (hf₂ : ∀ x, x ≤ f x) (hf₃ : le_closure' := hf₂ idempotent' x := (hf₃ x).antisymm (hf₁ (hf₂ x)) #align closure_operator.mk' ClosureOperator.mk' +#align closure_operator.mk'_apply ClosureOperator.mk'_apply /-- Convenience constructor for a closure operator using the weaker minimality axiom: `x ≤ f y → f x ≤ f y`, which is sometimes easier to prove in practice. -/ @@ -122,6 +124,7 @@ def mk₂ (f : α → α) (hf : ∀ x, x ≤ f x) (hmin : ∀ ⦃x y⦄, x ≤ f le_closure' := hf idempotent' _ := (hmin le_rfl).antisymm (hf _) #align closure_operator.mk₂ ClosureOperator.mk₂ +#align closure_operator.mk₂_apply ClosureOperator.mk₂_apply /-- Expanded out version of `mk₂`. `p` implies being closed. This constructor should be used when you already know a sufficient condition for being closed and using `mem_mk₃_closed` will avoid you @@ -131,6 +134,7 @@ def mk₃ (f : α → α) (p : α → Prop) (hf : ∀ x, x ≤ f x) (hfp : ∀ x (hmin : ∀ ⦃x y⦄, x ≤ y → p y → f x ≤ y) : ClosureOperator α := mk₂ f hf fun _ y hxy => hmin hxy (hfp y) #align closure_operator.mk₃ ClosureOperator.mk₃ +#align closure_operator.mk₃_apply ClosureOperator.mk₃_apply /-- This lemma shows that the image of `x` of a closure operator built from the `mk₃` constructor respects `p`, the property that was fed into it. -/ @@ -317,6 +321,7 @@ protected def id [Preorder α] : LowerAdjoint (id : α → α) toFun x := x gc' := GaloisConnection.id #align lower_adjoint.id LowerAdjoint.id +#align lower_adjoint.id_to_fun LowerAdjoint.id_toFun variable {α} @@ -370,6 +375,7 @@ def closureOperator : ClosureOperator α where le_closure' := l.le_closure idempotent' x := l.gc.u_l_u_eq_u (l x) #align lower_adjoint.closure_operator LowerAdjoint.closureOperator +#align lower_adjoint.closure_operator_apply LowerAdjoint.closureOperator_apply theorem idempotent (x : α) : u (l (u (l x))) = u (l x) := l.closureOperator.idempotent _ @@ -547,6 +553,7 @@ def GaloisConnection.lowerAdjoint [Preorder α] [Preorder β] {l : α → β} {u toFun := l gc' := gc #align galois_connection.lower_adjoint GaloisConnection.lowerAdjoint +#align galois_connection.lower_adjoint_to_fun GaloisConnection.lowerAdjoint_toFun /-- Every Galois connection induces a closure operator given by the composition. This is the partial order version of the statement that every adjunction induces a monad. -/ @@ -555,6 +562,7 @@ def GaloisConnection.closureOperator [PartialOrder α] [Preorder β] {l : α → (gc : GaloisConnection l u) : ClosureOperator α := gc.lowerAdjoint.closureOperator #align galois_connection.closure_operator GaloisConnection.closureOperator +#align galois_connection.closure_operator_apply GaloisConnection.closureOperator_apply /-- The set of closed elements has a Galois insertion to the underlying type. -/ def _root_.ClosureOperator.gi [PartialOrder α] (c : ClosureOperator α) : diff --git a/Mathlib/Order/Compare.lean b/Mathlib/Order/Compare.lean index 54e9113436f3d..49d440883c3a8 100644 --- a/Mathlib/Order/Compare.lean +++ b/Mathlib/Order/Compare.lean @@ -56,6 +56,7 @@ def Compares [LT α] : Ordering → α → α → Prop | lt, a, b => a < b | eq, a, b => a = b | gt, a, b => a > b +#align ordering.compares Ordering.Compares @[simp] lemma compares_lt [LT α] (a b : α) : Compares lt a b = (a < b) := rfl @@ -71,48 +72,60 @@ theorem compares_swap [LT α] {a b : α} {o : Ordering} : o.swap.Compares a b · exact Iff.rfl · exact eq_comm · exact Iff.rfl +#align ordering.compares_swap Ordering.compares_swap alias compares_swap ↔ Compares.of_swap Compares.swap +#align ordering.compares.of_swap Ordering.Compares.of_swap +#align ordering.compares.swap Ordering.Compares.swap theorem swap_eq_iff_eq_swap {o o' : Ordering} : o.swap = o' ↔ o = o'.swap := by rw [← swap_inj, swap_swap] +#align ordering.swap_eq_iff_eq_swap Ordering.swap_eq_iff_eq_swap theorem Compares.eq_lt [Preorder α] : ∀ {o} {a b : α}, Compares o a b → (o = lt ↔ a < b) | lt, a, b, h => ⟨fun _ => h, fun _ => rfl⟩ | eq, a, b, h => ⟨fun h => by injection h, fun h' => (ne_of_lt h' h).elim⟩ | gt, a, b, h => ⟨fun h => by injection h, fun h' => (lt_asymm h h').elim⟩ +#align ordering.compares.eq_lt Ordering.Compares.eq_lt theorem Compares.ne_lt [Preorder α] : ∀ {o} {a b : α}, Compares o a b → (o ≠ lt ↔ b ≤ a) | lt, a, b, h => ⟨absurd rfl, fun h' => (not_le_of_lt h h').elim⟩ | eq, a, b, h => ⟨fun _ => ge_of_eq h, fun _ h => by injection h⟩ | gt, a, b, h => ⟨fun _ => le_of_lt h, fun _ h => by injection h⟩ +#align ordering.compares.ne_lt Ordering.Compares.ne_lt theorem Compares.eq_eq [Preorder α] : ∀ {o} {a b : α}, Compares o a b → (o = eq ↔ a = b) | lt, a, b, h => ⟨fun h => by injection h, fun h' => (ne_of_lt h h').elim⟩ | eq, a, b, h => ⟨fun _ => h, fun _ => rfl⟩ | gt, a, b, h => ⟨fun h => by injection h, fun h' => (ne_of_gt h h').elim⟩ +#align ordering.compares.eq_eq Ordering.Compares.eq_eq theorem Compares.eq_gt [Preorder α] {o} {a b : α} (h : Compares o a b) : o = gt ↔ b < a := swap_eq_iff_eq_swap.symm.trans h.swap.eq_lt +#align ordering.compares.eq_gt Ordering.Compares.eq_gt theorem Compares.ne_gt [Preorder α] {o} {a b : α} (h : Compares o a b) : o ≠ gt ↔ a ≤ b := (not_congr swap_eq_iff_eq_swap.symm).trans h.swap.ne_lt +#align ordering.compares.ne_gt Ordering.Compares.ne_gt theorem Compares.le_total [Preorder α] {a b : α} : ∀ {o}, Compares o a b → a ≤ b ∨ b ≤ a | lt, h => Or.inl (le_of_lt h) | eq, h => Or.inl (le_of_eq h) | gt, h => Or.inr (le_of_lt h) +#align ordering.compares.le_total Ordering.Compares.le_total theorem Compares.le_antisymm [Preorder α] {a b : α} : ∀ {o}, Compares o a b → a ≤ b → b ≤ a → a = b | lt, h, _, hba => (not_le_of_lt h hba).elim | eq, h, _, _ => h | gt, h, hab, _ => (not_le_of_lt h hab).elim +#align ordering.compares.le_antisymm Ordering.Compares.le_antisymm theorem Compares.inj [Preorder α] {o₁} : ∀ {o₂} {a b : α}, Compares o₁ a b → Compares o₂ a b → o₁ = o₂ | lt, _, _, h₁, h₂ => h₁.eq_lt.2 h₂ | eq, _, _, h₁, h₂ => h₁.eq_eq.2 h₂ | gt, _, _, h₁, h₂ => h₁.eq_gt.2 h₂ +#align ordering.compares.inj Ordering.Compares.inj -- Porting note: mathlib3 proof uses `change ... at hab` theorem compares_iff_of_compares_impl [LinearOrder α] [Preorder β] {a b : α} {a' b' : β} @@ -126,6 +139,7 @@ theorem compares_iff_of_compares_impl [LinearOrder α] [Preorder β] {a b : α} rwa [ho.inj (h hab)] · have hab : Compares Ordering.gt a b := hab rwa [ho.inj (h hab)] +#align ordering.compares_iff_of_compares_impl Ordering.compares_iff_of_compares_impl theorem swap_orElse (o₁ o₂) : (orElse o₁ o₂).swap = orElse o₁.swap o₂.swap := by cases o₁ <;> rfl @@ -155,16 +169,19 @@ theorem ofDual_compares_ofDual [LT α] {a b : αᵒᵈ} {o : Ordering} : theorem cmp_compares [LinearOrder α] (a b : α) : (cmp a b).Compares a b := by obtain h | h | h := lt_trichotomy a b <;> simp [cmp, cmpUsing, h, h.not_lt] +#align cmp_compares cmp_compares theorem Ordering.Compares.cmp_eq [LinearOrder α] {a b : α} {o : Ordering} (h : o.Compares a b) : cmp a b = o := (cmp_compares a b).inj h +#align ordering.compares.cmp_eq Ordering.Compares.cmp_eq @[simp] theorem cmp_swap [Preorder α] [@DecidableRel α (· < ·)] (a b : α) : (cmp a b).swap = cmp b a := by unfold cmp cmpUsing by_cases a < b <;> by_cases h₂:b < a <;> simp [h, h₂, Ordering.swap] exact lt_asymm h h₂ +#align cmp_swap cmp_swap -- Porting note: Not sure why the simpNF linter doesn't like this. @semorrison @[simp, nolint simpNF] @@ -204,42 +221,51 @@ def linearOrderOfCompares [Preorder α] (cmp : α → α → Ordering) decidable_le := H, decidable_lt := fun a b => decidable_of_iff _ (h a b).eq_lt, decidable_eq := fun a b => decidable_of_iff _ (h a b).eq_eq } +#align linear_order_of_compares linearOrderOfCompares variable [LinearOrder α] (x y : α) @[simp] theorem cmp_eq_lt_iff : cmp x y = Ordering.lt ↔ x < y := Ordering.Compares.eq_lt (cmp_compares x y) +#align cmp_eq_lt_iff cmp_eq_lt_iff @[simp] theorem cmp_eq_eq_iff : cmp x y = Ordering.eq ↔ x = y := Ordering.Compares.eq_eq (cmp_compares x y) +#align cmp_eq_eq_iff cmp_eq_eq_iff @[simp] theorem cmp_eq_gt_iff : cmp x y = Ordering.gt ↔ y < x := Ordering.Compares.eq_gt (cmp_compares x y) +#align cmp_eq_gt_iff cmp_eq_gt_iff @[simp] theorem cmp_self_eq_eq : cmp x x = Ordering.eq := by rw [cmp_eq_eq_iff] +#align cmp_self_eq_eq cmp_self_eq_eq variable {x y} {β : Type _} [LinearOrder β] {x' y' : β} theorem cmp_eq_cmp_symm : cmp x y = cmp x' y' ↔ cmp y x = cmp y' x' := ⟨fun h => by rwa [← cmp_swap x', ← cmp_swap, swap_inj], fun h => by rwa [← cmp_swap y', ← cmp_swap, swap_inj]⟩ +#align cmp_eq_cmp_symm cmp_eq_cmp_symm theorem lt_iff_lt_of_cmp_eq_cmp (h : cmp x y = cmp x' y') : x < y ↔ x' < y' := by rw [← cmp_eq_lt_iff, ← cmp_eq_lt_iff, h] +#align lt_iff_lt_of_cmp_eq_cmp lt_iff_lt_of_cmp_eq_cmp theorem le_iff_le_of_cmp_eq_cmp (h : cmp x y = cmp x' y') : x ≤ y ↔ x' ≤ y' := by rw [← not_lt, ← not_lt] apply not_congr apply lt_iff_lt_of_cmp_eq_cmp rwa [cmp_eq_cmp_symm] +#align le_iff_le_of_cmp_eq_cmp le_iff_le_of_cmp_eq_cmp theorem eq_iff_eq_of_cmp_eq_cmp (h : cmp x y = cmp x' y') : x = y ↔ x' = y' := by rw [le_antisymm_iff, le_antisymm_iff, le_iff_le_of_cmp_eq_cmp h, le_iff_le_of_cmp_eq_cmp (cmp_eq_cmp_symm.1 h)] +#align eq_iff_eq_of_cmp_eq_cmp eq_iff_eq_of_cmp_eq_cmp theorem LT.lt.cmp_eq_lt (h : x < y) : cmp x y = Ordering.lt := (cmp_eq_lt_iff _ _).2 h @@ -249,6 +275,8 @@ theorem LT.lt.cmp_eq_gt (h : x < y) : cmp y x = Ordering.gt := theorem Eq.cmp_eq_eq (h : x = y) : cmp x y = Ordering.eq := (cmp_eq_eq_iff _ _).2 h +#align eq.cmp_eq_eq Eq.cmp_eq_eq theorem Eq.cmp_eq_eq' (h : x = y) : cmp y x = Ordering.eq := h.symm.cmp_eq_eq +#align eq.cmp_eq_eq' Eq.cmp_eq_eq' diff --git a/Mathlib/Order/CompleteLattice.lean b/Mathlib/Order/CompleteLattice.lean index f972938d5bacc..253a488d66d89 100644 --- a/Mathlib/Order/CompleteLattice.lean +++ b/Mathlib/Order/CompleteLattice.lean @@ -1928,6 +1928,7 @@ congr_arg₂ Prod.mk (congr_arg infₛ $ fst_image_prod _ ht) (congr_arg infₛ lemma Sup_prod [SupSet α] [SupSet β] {s : Set α} {t : Set β} (hs : s.Nonempty) (ht : t.Nonempty) : supₛ (s ×ˢ t) = (supₛ s, supₛ t) := congr_arg₂ Prod.mk (congr_arg supₛ $ fst_image_prod _ ht) (congr_arg supₛ $ snd_image_prod hs _) +#align Sup_prod Sup_prod section CompleteLattice diff --git a/Mathlib/Order/Concept.lean b/Mathlib/Order/Concept.lean index ce5d04b8b44b9..78a92e5c2147f 100644 --- a/Mathlib/Order/Concept.lean +++ b/Mathlib/Order/Concept.lean @@ -375,6 +375,7 @@ instance : Inhabited (Concept α β r) := def swap (c : Concept α β r) : Concept β α (swap r) := ⟨c.toProd.swap, c.closure_snd, c.closure_fst⟩ #align concept.swap Concept.swap +#align concept.swap_to_prod Concept.swap_toProd @[simp] theorem swap_swap (c : Concept α β r) : c.swap.swap = c := @@ -400,5 +401,7 @@ def swapEquiv : (Concept α β r)ᵒᵈ ≃o Concept β α (Function.swap r) whe right_inv := swap_swap map_rel_iff' := swap_le_swap_iff #align concept.swap_equiv Concept.swapEquiv +#align concept.swap_equiv_symm_apply Concept.swapEquiv_symmApply +#align concept.swap_equiv_apply Concept.swapEquiv_apply end Concept diff --git a/Mathlib/Order/Cover.lean b/Mathlib/Order/Cover.lean index 05c7480315bdc..10a1f4a262ae0 100644 --- a/Mathlib/Order/Cover.lean +++ b/Mathlib/Order/Cover.lean @@ -145,8 +145,10 @@ theorem ofDual_wcovby_ofDual_iff {a b : αᵒᵈ} : ofDual a ⩿ ofDual b ↔ b #align of_dual_wcovby_of_dual_iff ofDual_wcovby_ofDual_iff alias toDual_wcovby_toDual_iff ↔ _ Wcovby.toDual +#align wcovby.to_dual Wcovby.toDual alias ofDual_wcovby_ofDual_iff ↔ _ Wcovby.ofDual +#align wcovby.of_dual Wcovby.ofDual end Preorder @@ -230,6 +232,7 @@ theorem not_covby_iff (h : a < b) : ¬a ⋖ b ↔ ∃ c, a < c ∧ c < b := by #align not_covby_iff not_covby_iff alias not_covby_iff ↔ exists_lt_lt_of_not_covby _ +#align exists_lt_lt_of_not_covby exists_lt_lt_of_not_covby alias exists_lt_lt_of_not_covby ← LT.lt.exists_lt_lt @@ -255,8 +258,10 @@ theorem ofDual_covby_ofDual_iff {a b : αᵒᵈ} : ofDual a ⋖ ofDual b ↔ b #align of_dual_covby_of_dual_iff ofDual_covby_ofDual_iff alias toDual_covby_toDual_iff ↔ _ Covby.toDual +#align covby.to_dual Covby.toDual alias ofDual_covby_ofDual_iff ↔ _ Covby.ofDual +#align covby.of_dual Covby.ofDual end LT @@ -382,8 +387,10 @@ theorem wcovby_iff_eq_or_covby : a ⩿ b ↔ a = b ∨ a ⋖ b := #align wcovby_iff_eq_or_covby wcovby_iff_eq_or_covby alias wcovby_iff_covby_or_eq ↔ Wcovby.covby_or_eq _ +#align wcovby.covby_or_eq Wcovby.covby_or_eq alias wcovby_iff_eq_or_covby ↔ Wcovby.eq_or_covby _ +#align wcovby.eq_or_covby Wcovby.eq_or_covby theorem Covby.eq_or_eq (h : a ⋖ b) (h2 : a ≤ c) (h3 : c ≤ b) : c = a ∨ c = b := h.wcovby.eq_or_eq h2 h3 diff --git a/Mathlib/Order/Disjoint.lean b/Mathlib/Order/Disjoint.lean index 8df367f02c7b8..02f471d5fb649 100644 --- a/Mathlib/Order/Disjoint.lean +++ b/Mathlib/Order/Disjoint.lean @@ -55,9 +55,11 @@ theorem Disjoint.symm ⦃a b : α⦄ : Disjoint a b → Disjoint b a := theorem symmetric_disjoint : Symmetric (Disjoint : α → α → Prop) := Disjoint.symm +#align symmetric_disjoint symmetric_disjoint @[simp] theorem disjoint_bot_left : Disjoint ⊥ a := fun _ hbot _ ↦ hbot +#align disjoint_bot_left disjoint_bot_left @[simp] theorem disjoint_bot_right : Disjoint a ⊥ := fun _ _ hbot ↦ hbot @@ -78,10 +80,12 @@ theorem Disjoint.mono_right : b ≤ c → Disjoint a c → Disjoint a b := @[simp] theorem disjoint_self : Disjoint a a ↔ a = ⊥ := ⟨fun hd ↦ bot_unique <| hd le_rfl le_rfl, fun h _ ha _ ↦ ha.trans_eq h⟩ +#align disjoint_self disjoint_self /- TODO: Rename `Disjoint.eq_bot` to `Disjoint.inf_eq` and `Disjoint.eq_bot_of_self` to `Disjoint.eq_bot` -/ alias disjoint_self ↔ Disjoint.eq_bot_of_self _ +#align disjoint.eq_bot_of_self Disjoint.eq_bot_of_self theorem Disjoint.ne (ha : a ≠ ⊥) (hab : Disjoint a b) : a ≠ b := fun h ↦ ha <| disjoint_self.1 <| by rwa [← h] at hab @@ -104,10 +108,12 @@ variable [PartialOrder α] [BoundedOrder α] {a : α} @[simp] theorem disjoint_top : Disjoint a ⊤ ↔ a = ⊥ := ⟨fun h ↦ bot_unique <| h le_rfl le_top, fun h _ ha _ ↦ ha.trans_eq h⟩ +#align disjoint_top disjoint_top @[simp] theorem top_disjoint : Disjoint ⊤ a ↔ a = ⊥ := ⟨fun h ↦ bot_unique <| h le_top le_rfl, fun h _ _ ha ↦ ha.trans_eq h⟩ +#align top_disjoint top_disjoint end PartialBoundedOrder @@ -117,9 +123,11 @@ variable [SemilatticeInf α] [OrderBot α] {a b c d : α} theorem disjoint_iff_inf_le : Disjoint a b ↔ a ⊓ b ≤ ⊥ := ⟨fun hd ↦ hd inf_le_left inf_le_right, fun h _ ha hb ↦ (le_inf ha hb).trans h⟩ +#align disjoint_iff_inf_le disjoint_iff_inf_le theorem disjoint_iff : Disjoint a b ↔ a ⊓ b = ⊥ := disjoint_iff_inf_le.trans le_bot_iff +#align disjoint_iff disjoint_iff theorem Disjoint.le_bot : Disjoint a b → a ⊓ b ≤ ⊥ := disjoint_iff_inf_le.mp @@ -131,12 +139,15 @@ theorem Disjoint.eq_bot : Disjoint a b → a ⊓ b = ⊥ := theorem disjoint_assoc : Disjoint (a ⊓ b) c ↔ Disjoint a (b ⊓ c) := by rw [disjoint_iff_inf_le, disjoint_iff_inf_le, inf_assoc] +#align disjoint_assoc disjoint_assoc theorem disjoint_left_comm : Disjoint a (b ⊓ c) ↔ Disjoint b (a ⊓ c) := by simp_rw [disjoint_iff_inf_le, inf_left_comm] +#align disjoint_left_comm disjoint_left_comm theorem disjoint_right_comm : Disjoint (a ⊓ b) c ↔ Disjoint (a ⊓ c) b := by simp_rw [disjoint_iff_inf_le, inf_right_comm] +#align disjoint_right_comm disjoint_right_comm variable (c) @@ -175,10 +186,12 @@ variable [DistribLattice α] [OrderBot α] {a b c : α} @[simp] theorem disjoint_sup_left : Disjoint (a ⊔ b) c ↔ Disjoint a c ∧ Disjoint b c := by simp only [disjoint_iff, inf_sup_right, sup_eq_bot_iff] +#align disjoint_sup_left disjoint_sup_left @[simp] theorem disjoint_sup_right : Disjoint a (b ⊔ c) ↔ Disjoint a b ∧ Disjoint a c := by simp only [disjoint_iff, inf_sup_left, sup_eq_bot_iff] +#align disjoint_sup_right disjoint_sup_right theorem Disjoint.sup_left (ha : Disjoint a c) (hb : Disjoint b c) : Disjoint (a ⊔ b) c := disjoint_sup_left.2 ⟨ha, hb⟩ @@ -226,12 +239,15 @@ theorem Codisjoint.symm ⦃a b : α⦄ : Codisjoint a b → Codisjoint b a := theorem symmetric_codisjoint : Symmetric (Codisjoint : α → α → Prop) := Codisjoint.symm +#align symmetric_codisjoint symmetric_codisjoint @[simp] theorem codisjoint_top_left : Codisjoint ⊤ a := fun _ htop _ ↦ htop +#align codisjoint_top_left codisjoint_top_left @[simp] theorem codisjoint_top_right : Codisjoint a ⊤ := fun _ _ htop ↦ htop +#align codisjoint_top_right codisjoint_top_right theorem Codisjoint.mono (h₁ : a ≤ b) (h₂ : c ≤ d) : Codisjoint a c → Codisjoint b d := fun h _ ha hc ↦ h (h₁.trans ha) (h₂.trans hc) @@ -248,10 +264,12 @@ theorem Codisjoint.mono_right : b ≤ c → Codisjoint a b → Codisjoint a c := @[simp] theorem codisjoint_self : Codisjoint a a ↔ a = ⊤ := ⟨fun hd ↦ top_unique <| hd le_rfl le_rfl, fun h _ ha _ ↦ h.symm.trans_le ha⟩ +#align codisjoint_self codisjoint_self /- TODO: Rename `Codisjoint.eq_top` to `Codisjoint.sup_eq` and `Codisjoint.eq_top_of_self` to `Codisjoint.eq_top` -/ alias codisjoint_self ↔ Codisjoint.eq_top_of_self _ +#align codisjoint.eq_top_of_self Codisjoint.eq_top_of_self theorem Codisjoint.ne (ha : a ≠ ⊤) (hab : Codisjoint a b) : a ≠ b := fun h ↦ ha <| codisjoint_self.1 <| by rwa [← h] at hab @@ -274,10 +292,12 @@ variable [PartialOrder α] [BoundedOrder α] {a : α} @[simp] theorem codisjoint_bot : Codisjoint a ⊥ ↔ a = ⊤ := ⟨fun h ↦ top_unique <| h le_rfl bot_le, fun h _ ha _ ↦ h.symm.trans_le ha⟩ +#align codisjoint_bot codisjoint_bot @[simp] theorem bot_codisjoint : Codisjoint ⊥ a ↔ a = ⊤ := ⟨fun h ↦ top_unique <| h bot_le le_rfl, fun h _ _ ha ↦ h.symm.trans_le ha⟩ +#align bot_codisjoint bot_codisjoint end PartialBoundedOrder @@ -303,12 +323,15 @@ theorem Codisjoint.eq_top : Codisjoint a b → a ⊔ b = ⊤ := theorem codisjoint_assoc : Codisjoint (a ⊔ b) c ↔ Codisjoint a (b ⊔ c) := @disjoint_assoc αᵒᵈ _ _ _ _ _ +#align codisjoint_assoc codisjoint_assoc theorem codisjoint_left_comm : Codisjoint a (b ⊔ c) ↔ Codisjoint b (a ⊔ c) := @disjoint_left_comm αᵒᵈ _ _ _ _ _ +#align codisjoint_left_comm codisjoint_left_comm theorem codisjoint_right_comm : Codisjoint (a ⊔ b) c ↔ Codisjoint (a ⊔ c) b := @disjoint_right_comm αᵒᵈ _ _ _ _ _ +#align codisjoint_right_comm codisjoint_right_comm variable (c) @@ -349,10 +372,12 @@ variable [DistribLattice α] [OrderTop α] {a b c : α} @[simp] theorem codisjoint_inf_left : Codisjoint (a ⊓ b) c ↔ Codisjoint a c ∧ Codisjoint b c := by simp only [codisjoint_iff, sup_inf_right, inf_eq_top_iff] +#align codisjoint_inf_left codisjoint_inf_left @[simp] theorem codisjoint_inf_right : Codisjoint a (b ⊓ c) ↔ Codisjoint a b ∧ Codisjoint a c := by simp only [codisjoint_iff, sup_inf_left, inf_eq_top_iff] +#align codisjoint_inf_right codisjoint_inf_right theorem Codisjoint.inf_left (ha : Codisjoint a c) (hb : Codisjoint b c) : Codisjoint (a ⊓ b) c := codisjoint_inf_left.2 ⟨ha, hb⟩ diff --git a/Mathlib/Order/Heyting/Basic.lean b/Mathlib/Order/Heyting/Basic.lean index 651d633033911..4a18c25189e69 100644 --- a/Mathlib/Order/Heyting/Basic.lean +++ b/Mathlib/Order/Heyting/Basic.lean @@ -558,8 +558,10 @@ theorem sdiff_sup_self (a b : α) : b \ a ⊔ a = b ⊔ a := by rw [sup_comm, su #align sdiff_sup_self sdiff_sup_self alias sdiff_sup_self ← sup_sdiff_self_left +#align sup_sdiff_self_left sup_sdiff_self_left alias sup_sdiff_self ← sup_sdiff_self_right +#align sup_sdiff_self_right sup_sdiff_self_right theorem sup_sdiff_eq_sup (h : c ≤ a) : a ⊔ b \ c = a ⊔ b := sup_congr_left (sdiff_le.trans le_sup_right) <| le_sup_sdiff.trans <| sup_le_sup_right h _ @@ -817,12 +819,16 @@ theorem le_compl_comm : a ≤ bᶜ ↔ b ≤ aᶜ := by #align le_compl_comm le_compl_comm alias le_compl_iff_disjoint_right ↔ _ Disjoint.le_compl_right +#align disjoint.le_compl_right Disjoint.le_compl_right alias le_compl_iff_disjoint_left ↔ _ Disjoint.le_compl_left +#align disjoint.le_compl_left Disjoint.le_compl_left alias le_compl_comm ← le_compl_iff_le_compl +#align le_compl_iff_le_compl le_compl_iff_le_compl alias le_compl_comm ↔ le_compl_of_le_compl _ +#align le_compl_of_le_compl le_compl_of_le_compl theorem disjoint_compl_left : Disjoint (aᶜ) a := disjoint_iff_inf_le.mpr <| le_himp_iff.1 (himp_bot _).ge @@ -1014,8 +1020,10 @@ theorem hnot_le_comm : ¬a ≤ b ↔ ¬b ≤ a := by #align hnot_le_comm hnot_le_comm alias hnot_le_iff_codisjoint_right ↔ _ Codisjoint.hnot_le_right +#align codisjoint.hnot_le_right Codisjoint.hnot_le_right alias hnot_le_iff_codisjoint_left ↔ _ Codisjoint.hnot_le_left +#align codisjoint.hnot_le_left Codisjoint.hnot_le_left theorem codisjoint_hnot_right : Codisjoint a (¬a) := codisjoint_iff_le_sup.2 <| sdiff_le_iff.1 (top_sdiff' _).le diff --git a/Mathlib/Order/Hom/Basic.lean b/Mathlib/Order/Hom/Basic.lean index 53fcb58004da2..378ab0b7a8ba4 100644 --- a/Mathlib/Order/Hom/Basic.lean +++ b/Mathlib/Order/Hom/Basic.lean @@ -287,6 +287,7 @@ theorem copy_eq (f : α →o β) (f' : α → β) (h : f' = f) : f.copy f' h = f def id : α →o α := ⟨_root_.id, monotone_id⟩ #align order_hom.id OrderHom.id +#align order_hom.id_coe OrderHom.id_coe instance : Inhabited (α →o α) := ⟨id⟩ @@ -349,6 +350,7 @@ theorem curry_symm_apply (f : α →o β →o γ) (x : α × β) : curry.symm f def comp (g : β →o γ) (f : α →o β) : α →o γ := ⟨g ∘ f, g.mono.comp f.mono⟩ #align order_hom.comp OrderHom.comp +#align order_hom.comp_coe OrderHom.comp_coe -- Porting note: `mono` tactic not implemented yet. -- @[mono] @@ -361,6 +363,7 @@ theorem comp_mono ⦃g₁ g₂ : β →o γ⦄ (hg : g₁ ≤ g₂) ⦃f₁ f₂ def compₘ : (β →o γ) →o (α →o β) →o α →o γ := curry ⟨fun f : (β →o γ) × (α →o β) => f.1.comp f.2, fun _ _ h => comp_mono h.1 h.2⟩ #align order_hom.compₘ OrderHom.compₘ +#align order_hom.compₘ_coe_coe_coe OrderHom.compₘ_coe_coe_coe @[simp] theorem comp_id (f : α →o β) : comp f id = f := by @@ -380,6 +383,7 @@ def const (α : Type _) [Preorder α] {β : Type _} [Preorder β] : β →o α toFun b := ⟨Function.const α b, fun _ _ _ => le_rfl⟩ monotone' _ _ h _ := h #align order_hom.const OrderHom.const +#align order_hom.const_coe_coe OrderHom.const_coe_coe @[simp] theorem const_comp (f : α →o β) (c : γ) : (const β c).comp f = const α c := @@ -398,6 +402,7 @@ theorem comp_const (γ : Type _) [Preorder γ] (f : α →o β) (c : α) : protected def prod (f : α →o β) (g : α →o γ) : α →o β × γ := ⟨fun x => (f x, g x), fun _ _ h => ⟨f.mono h, g.mono h⟩⟩ #align order_hom.prod OrderHom.prod +#align order_hom.prod_coe OrderHom.prod_coe --@[mono] theorem prod_mono {f₁ f₂ : α →o β} (hf : f₁ ≤ f₂) {g₁ g₂ : α →o γ} (hg : g₁ ≤ g₂) : @@ -415,30 +420,35 @@ theorem comp_prod_comp_same (f₁ f₂ : β →o γ) (g : α →o β) : def prodₘ : (α →o β) →o (α →o γ) →o α →o β × γ := curry ⟨fun f : (α →o β) × (α →o γ) => f.1.prod f.2, fun _ _ h => prod_mono h.1 h.2⟩ #align order_hom.prodₘ OrderHom.prodₘ +#align order_hom.prodₘ_coe_coe_coe OrderHom.prodₘ_coe_coe_coe /-- Diagonal embedding of `α` into `α × α` as a `OrderHom`. -/ @[simps] def diag : α →o α × α := id.prod id #align order_hom.diag OrderHom.diag +#align order_hom.diag_coe OrderHom.diag_coe /-- Restriction of `f : α →o α →o β` to the diagonal. -/ @[simps (config := { simpRhs := true })] def onDiag (f : α →o α →o β) : α →o β := (curry.symm f).comp diag #align order_hom.on_diag OrderHom.onDiag +#align order_hom.on_diag_coe OrderHom.onDiag_coe /-- `Prod.fst` as a `OrderHom`. -/ @[simps] def fst : α × β →o α := ⟨Prod.fst, fun _ _ h => h.1⟩ #align order_hom.fst OrderHom.fst +#align order_hom.fst_coe OrderHom.fst_coe /-- `Prod.snd` as a `OrderHom`. -/ @[simps] def snd : α × β →o β := ⟨Prod.snd, fun _ _ h => h.2⟩ #align order_hom.snd OrderHom.snd +#align order_hom.snd_coe OrderHom.snd_coe @[simp] theorem fst_prod_snd : (fst : α × β →o α).prod snd = id := by @@ -466,12 +476,15 @@ def prodIso : (α →o β × γ) ≃o (α →o β) × (α →o γ) where right_inv f := by ext <;> rfl map_rel_iff' := forall_and.symm #align order_hom.prod_iso OrderHom.prodIso +#align order_hom.prod_iso_apply OrderHom.prodIso_apply +#align order_hom.prod_iso_symm_apply OrderHom.prodIso_symmApply /-- `Prod.map` of two `OrderHom`s as a `OrderHom`. -/ @[simps] def prodMap (f : α →o β) (g : γ →o δ) : α × γ →o β × δ := ⟨Prod.map f g, fun _ _ h => ⟨f.mono h.1, g.mono h.2⟩⟩ #align order_hom.prod_map OrderHom.prodMap +#align order_hom.prod_map_coe OrderHom.prodMap_coe variable {ι : Type _} {π : ι → Type _} [∀ i, Preorder (π i)] @@ -480,6 +493,7 @@ variable {ι : Type _} {π : ι → Type _} [∀ i, Preorder (π i)] def _root_.Pi.evalOrderHom (i : ι) : (∀ j, π j) →o π i := ⟨Function.eval i, Function.monotone_eval i⟩ #align pi.eval_order_hom Pi.evalOrderHom +#align pi.eval_order_hom_coe Pi.evalOrderHom_coe /-- The "forgetful functor" from `α →o β` to `α → β` that takes the underlying function, is monotone. -/ @@ -488,6 +502,7 @@ def coeFnHom : (α →o β) →o α → β where toFun f := f monotone' _ _ h := h #align order_hom.coe_fn_hom OrderHom.coeFnHom +#align order_hom.coe_fn_hom_coe OrderHom.coeFnHom_coe /-- Function application `fun f => f a` (for fixed `a`) is a monotone function from the monotone function space `α →o β` to `β`. See also `Pi.evalOrderHom`. -/ @@ -495,6 +510,7 @@ monotone function space `α →o β` to `β`. See also `Pi.evalOrderHom`. -/ def apply (x : α) : (α →o β) →o β := (Pi.evalOrderHom x).comp coeFnHom #align order_hom.apply OrderHom.apply +#align order_hom.apply_coe OrderHom.apply_coe /-- Construct a bundled monotone map `α →o Π i, π i` from a family of monotone maps `f i : α →o π i`. -/ @@ -502,6 +518,7 @@ def apply (x : α) : (α →o β) →o β := def pi (f : ∀ i, α →o π i) : α →o ∀ i, π i := ⟨fun x i => f i x, fun _ _ h i => (f i).mono h⟩ #align order_hom.pi OrderHom.pi +#align order_hom.pi_coe OrderHom.pi_coe /-- Order isomorphism between bundled monotone maps `α →o Π i, π i` and families of bundled monotone maps `Π i, α →o π i`. -/ @@ -517,12 +534,15 @@ def piIso : (α →o ∀ i, π i) ≃o ∀ i, α →o π i where rfl map_rel_iff' := forall_swap #align order_hom.pi_iso OrderHom.piIso +#align order_hom.pi_iso_apply OrderHom.piIso_apply +#align order_hom.pi_iso_symm_apply OrderHom.piIso_symmApply /-- `Subtype.val` as a bundled monotone function. -/ @[simps (config := { fullyApplied := false })] def Subtype.val (p : α → Prop) : Subtype p →o α := ⟨_root_.Subtype.val, fun _ _ h => h⟩ #align order_hom.subtype.val OrderHom.Subtype.val +#align order_hom.subtype.val_coe OrderHom.Subtype.val_coe /-- There is a unique monotone map from a subsingleton to itself. -/ instance unique [Subsingleton α] : Unique (α →o α) where @@ -543,6 +563,8 @@ protected def dual : (α →o β) ≃ (αᵒᵈ →o βᵒᵈ) where left_inv _ := ext _ _ rfl right_inv _ := ext _ _ rfl #align order_hom.dual OrderHom.dual +#align order_hom.dual_apply_coe OrderHom.dual_apply_coe +#align order_hom.dual_symm_apply_coe OrderHom.dual_symm_apply_coe -- Porting note: We used to be able to write `(OrderHom.id : α →o α).dual` here rather than -- `OrderHom.dual (OrderHom.id : α →o α)`. @@ -580,12 +602,14 @@ def dualIso (α β : Type _) [Preorder α] [Preorder β] : (α →o β) ≃o (α protected def withBotMap (f : α →o β) : WithBot α →o WithBot β := ⟨WithBot.map f, f.mono.withBot_map⟩ #align order_hom.with_bot_map OrderHom.withBotMap +#align order_hom.with_bot_map_coe OrderHom.withBotMap_coe /-- Lift an order homomorphism `f : α →o β` to an order homomorphism `WithTop α →o WithTop β`. -/ @[simps (config := { fullyApplied := false })] protected def withTopMap (f : α →o β) : WithTop α →o WithTop β := ⟨WithTop.map f, f.mono.withTop_map⟩ #align order_hom.with_top_map OrderHom.withTopMap +#align order_hom.with_top_map_coe OrderHom.withTopMap_coe end OrderHom @@ -668,12 +692,14 @@ protected def withBotMap (f : α ↪o β) : WithBot α ↪o WithBot β := toFun := WithBot.map f, map_rel_iff' := @fun a b => WithBot.map_le_iff f f.map_rel_iff a b } #align order_embedding.with_bot_map OrderEmbedding.withBotMap +#align order_embedding.with_bot_map_apply OrderEmbedding.withBotMap_apply /-- A version of `WithTop.map` for order embeddings. -/ @[simps (config := { fullyApplied := false })] protected def withTopMap (f : α ↪o β) : WithTop α ↪o WithTop β := { f.dual.withBotMap.dual with toFun := WithTop.map f } #align order_embedding.with_top_map OrderEmbedding.withTopMap +#align order_embedding.with_top_map_apply OrderEmbedding.withTopMap_apply /-- To define an order embedding from a partial order to a preorder it suffices to give a function together with a proof that it satisfies `f a ≤ f b ↔ a ≤ b`. @@ -705,6 +731,7 @@ theorem coe_ofStrictMono {α β} [LinearOrder α] [Preorder β] {f : α → β} def subtype (p : α → Prop) : Subtype p ↪o α := ⟨Function.Embedding.subtype p, Iff.rfl⟩ #align order_embedding.subtype OrderEmbedding.subtype +#align order_embedding.subtype_apply OrderEmbedding.subtype_apply /-- Convert an `OrderEmbedding` to a `OrderHom`. -/ @[simps (config := { fullyApplied := false })] @@ -712,6 +739,7 @@ def toOrderHom {X Y : Type _} [Preorder X] [Preorder Y] (f : X ↪o Y) : X →o toFun := f monotone' := f.monotone #align order_embedding.to_order_hom OrderEmbedding.toOrderHom +#align order_embedding.to_order_hom_coe OrderEmbedding.toOrderHom_coe end OrderEmbedding @@ -730,6 +758,7 @@ def toOrderHom : α →o β where toFun := f monotone' := StrictMono.monotone fun _ _ => f.map_rel #align rel_hom.to_order_hom RelHom.toOrderHom +#align rel_hom.to_order_hom_coe RelHom.toOrderHom_coe end RelHom @@ -1070,6 +1099,8 @@ def funUnique (α β : Type _) [Unique α] [Preorder β] : (α → β) ≃o β w toEquiv := Equiv.funUnique α β map_rel_iff' := by simp [Pi.le_def, Unique.forall_iff] #align order_iso.fun_unique OrderIso.funUnique +#align order_iso.fun_unique_apply OrderIso.funUnique_apply +#align order_iso.fun_unique_to_equiv OrderIso.funUnique_toEquiv @[simp] theorem funUnique_symm_apply {α β : Type _} [Unique α] [Preorder β] : @@ -1118,6 +1149,8 @@ def orderIsoOfRightInverse (g : β → α) (hg : Function.RightInverse g f) : α left_inv := fun _ => h_mono.injective <| hg _, right_inv := hg } #align strict_mono.order_iso_of_right_inverse StrictMono.orderIsoOfRightInverse +#align strict_mono.order_iso_of_right_inverse_apply StrictMono.orderIsoOfRightInverse_apply +#align strict_mono.order_iso_of_right_inverse_symm_apply StrictMono.orderIsoOfRightInverse_symmApply end StrictMono @@ -1283,6 +1316,7 @@ def withTopCongr (e : α ≃o β) : WithTop α ≃o WithTop β := { e.toOrderEmbedding.withTopMap with toEquiv := e.toEquiv.optionCongr } #align order_iso.with_top_congr OrderIso.withTopCongr +#align order_iso.with_top_congr_apply OrderIso.withTopCongr_apply @[simp] theorem withTopCongr_refl : (OrderIso.refl α).withTopCongr = OrderIso.refl _ := @@ -1305,6 +1339,7 @@ theorem withTopCongr_trans (e₁ : α ≃o β) (e₂ : β ≃o γ) : def withBotCongr (e : α ≃o β) : WithBot α ≃o WithBot β := { e.toOrderEmbedding.withBotMap with toEquiv := e.toEquiv.optionCongr } #align order_iso.with_bot_congr OrderIso.withBotCongr +#align order_iso.with_bot_congr_apply OrderIso.withBotCongr_apply @[simp] theorem withBotCongr_refl : (OrderIso.refl α).withBotCongr = OrderIso.refl _ := diff --git a/Mathlib/Order/Hom/Bounded.lean b/Mathlib/Order/Hom/Bounded.lean index 9ed7a745b0810..cc16520cc9154 100644 --- a/Mathlib/Order/Hom/Bounded.lean +++ b/Mathlib/Order/Hom/Bounded.lean @@ -723,6 +723,8 @@ protected def dual : left_inv _ := TopHom.ext fun _ => rfl right_inv _ := BotHom.ext fun _ => rfl #align top_hom.dual TopHom.dual +#align top_hom.dual_apply_apply TopHom.dual_apply_apply +#align top_hom.dual_symm_apply_apply TopHom.dual_symm_apply_apply @[simp] theorem dual_id : TopHom.dual (TopHom.id α) = BotHom.id _ := @@ -761,6 +763,8 @@ protected def dual : left_inv _ := BotHom.ext fun _ => rfl right_inv _ := TopHom.ext fun _ => rfl #align bot_hom.dual BotHom.dual +#align bot_hom.dual_apply_apply BotHom.dual_apply_apply +#align bot_hom.dual_symm_apply_apply BotHom.dual_symm_apply_apply @[simp] theorem dual_id : BotHom.dual (BotHom.id α) = TopHom.id _ := @@ -802,6 +806,8 @@ protected def dual : left_inv _ := ext fun _ => rfl right_inv _ := ext fun _ => rfl #align bounded_order_hom.dual BoundedOrderHom.dual +#align bounded_order_hom.dual_apply_to_order_hom BoundedOrderHom.dual_apply_toOrderHom +#align bounded_order_hom.dual_symm_apply_to_order_hom BoundedOrderHom.dual_symm_apply_toOrderHom @[simp] theorem dual_id : BoundedOrderHom.dual (BoundedOrderHom.id α) = BoundedOrderHom.id _ := diff --git a/Mathlib/Order/Hom/Set.lean b/Mathlib/Order/Hom/Set.lean index 0954ba14ee515..5ebe0acd8b356 100644 --- a/Mathlib/Order/Hom/Set.lean +++ b/Mathlib/Order/Hom/Set.lean @@ -110,6 +110,7 @@ protected noncomputable def orderIso : toEquiv := Equiv.ofInjective f h_mono.injective map_rel_iff' := h_mono.le_iff_le #align strict_mono.order_iso StrictMono.orderIso +#align strict_mono.order_iso_apply StrictMono.orderIso_apply /-- A strictly monotone surjective function from a linear order is an order isomorphism. -/ noncomputable def orderIsoOfSurjective : α ≃o β := @@ -148,6 +149,8 @@ def OrderIso.compl : α ≃o αᵒᵈ where right_inv := compl_compl (α := αᵒᵈ) map_rel_iff' := compl_le_compl_iff_le #align order_iso.compl OrderIso.compl +#align order_iso.compl_symm_apply OrderIso.compl_symmApply +#align order_iso.compl_apply OrderIso.compl_apply theorem compl_strictAnti : StrictAnti (compl : α → α) := (OrderIso.compl α).strictMono diff --git a/Mathlib/Order/InitialSeg.lean b/Mathlib/Order/InitialSeg.lean index c072e6fdf4734..87b36103df1b9 100644 --- a/Mathlib/Order/InitialSeg.lean +++ b/Mathlib/Order/InitialSeg.lean @@ -76,6 +76,7 @@ instance : EmbeddingLike (r ≼i s) α β := @[ext] lemma ext {f g : r ≼i s} (h : ∀ x, f x = g x) : f = g := FunLike.ext f g h +#align initial_seg.ext InitialSeg.ext @[simp] theorem coe_coe_fn (f : r ≼i s) : ((f : r ↪r s) : α → β) = f := @@ -88,6 +89,7 @@ theorem init' (f : r ≼i s) {a : α} {b : β} : s b (f a) → ∃ a', f a' = b theorem map_rel_iff (f : r ≼i s) : s (f a) (f b) ↔ r a b := f.map_rel_iff' +#align initial_seg.map_rel_iff InitialSeg.map_rel_iff theorem init_iff (f : r ≼i s) {a : α} {b : β} : s b (f a) ↔ ∃ a', f a' = b ∧ r a' a := ⟨fun h => by diff --git a/Mathlib/Order/Lattice.lean b/Mathlib/Order/Lattice.lean index f5839ded19313..2b3c3e95d4f22 100644 --- a/Mathlib/Order/Lattice.lean +++ b/Mathlib/Order/Lattice.lean @@ -197,8 +197,11 @@ theorem right_eq_sup : b = a ⊔ b ↔ a ≤ b := #align right_eq_sup right_eq_sup alias sup_eq_left ↔ _ sup_of_le_left +#align sup_of_le_left sup_of_le_left alias sup_eq_right ↔ le_of_sup_eq sup_of_le_right +#align sup_of_le_right sup_of_le_right +#align le_of_sup_eq le_of_sup_eq attribute [simp] sup_of_le_left sup_of_le_right @@ -451,8 +454,11 @@ theorem right_eq_inf : b = a ⊓ b ↔ b ≤ a := #align right_eq_inf right_eq_inf alias inf_eq_left ↔ le_of_inf_eq inf_of_le_left +#align inf_of_le_left inf_of_le_left +#align le_of_inf_eq le_of_inf_eq alias inf_eq_right ↔ _ inf_of_le_right +#align inf_of_le_right inf_of_le_right attribute [simp] inf_of_le_left inf_of_le_right @@ -681,6 +687,8 @@ theorem sup_le_inf : a ⊔ b ≤ a ⊓ b ↔ a = b := by simp [le_antisymm_iff, @[simp] lemma sup_eq_inf : a ⊔ b = a ⊓ b ↔ a = b := eq_comm.trans inf_eq_sup @[simp] lemma inf_lt_sup : a ⊓ b < a ⊔ b ↔ a ≠ b := by rw [inf_le_sup.lt_iff_ne, Ne.def, inf_eq_sup] #align inf_lt_sup inf_lt_sup +#align sup_eq_inf sup_eq_inf +#align inf_eq_sup inf_eq_sup lemma inf_eq_and_sup_eq_iff : a ⊓ b = c ∧ a ⊔ b = c ↔ a = c ∧ b = c := by refine' ⟨fun h ↦ _, _⟩ @@ -688,6 +696,7 @@ lemma inf_eq_and_sup_eq_iff : a ⊓ b = c ∧ a ⊔ b = c ↔ a = c ∧ b = c := simpa using h } { rintro ⟨rfl, rfl⟩ exact ⟨inf_idem, sup_idem⟩ } +#align inf_eq_and_sup_eq_iff inf_eq_and_sup_eq_iff /-! #### Distributivity laws diff --git a/Mathlib/Order/Max.lean b/Mathlib/Order/Max.lean index 227d406a1110e..305dbefd103ff 100644 --- a/Mathlib/Order/Max.lean +++ b/Mathlib/Order/Max.lean @@ -158,6 +158,7 @@ theorem noBotOrder_iff_noMinOrder (α : Type _) [LinearOrder α] : NoBotOrder α fun h => haveI := h inferInstance⟩ +#align no_bot_order_iff_no_min_order noBotOrder_iff_noMinOrder theorem noTopOrder_iff_noMaxOrder (α : Type _) [LinearOrder α] : NoTopOrder α ↔ NoMaxOrder α := ⟨fun h => @@ -166,12 +167,15 @@ theorem noTopOrder_iff_noMaxOrder (α : Type _) [LinearOrder α] : NoTopOrder α fun h => haveI := h inferInstance⟩ +#align no_top_order_iff_no_max_order noTopOrder_iff_noMaxOrder theorem NoMinOrder.not_acc [LT α] [NoMinOrder α] (a : α) : ¬Acc (· < ·) a := fun h => Acc.recOn h fun x _ => (exists_lt x).recOn +#align no_min_order.not_acc NoMinOrder.not_acc theorem NoMaxOrder.not_acc [LT α] [NoMaxOrder α] (a : α) : ¬Acc (· > ·) a := fun h => Acc.recOn h fun x _ => (exists_gt x).recOn +#align no_max_order.not_acc NoMaxOrder.not_acc section LE @@ -183,6 +187,7 @@ several bottom elements. When `α` is linear, this is useful to make a case disj `NoMinOrder α` within a proof. -/ def IsBot (a : α) : Prop := ∀ b, a ≤ b +#align is_bot IsBot /-- `a : α` is a top element of `α` if it is greater than or equal to any other element of `α`. This predicate is roughly an unbundled version of `OrderBot`, except that a preorder may have @@ -190,18 +195,21 @@ several top elements. When `α` is linear, this is useful to make a case disjunc `NoMaxOrder α` within a proof. -/ def IsTop (a : α) : Prop := ∀ b, b ≤ a +#align is_top IsTop /-- `a` is a minimal element of `α` if no element is strictly less than it. We spell it without `<` to avoid having to convert between `≤` and `<`. Instead, `isMin_iff_forall_not_lt` does the conversion. -/ def IsMin (a : α) : Prop := ∀ ⦃b⦄, b ≤ a → a ≤ b +#align is_min IsMin /-- `a` is a maximal element of `α` if no element is strictly greater than it. We spell it without `<` to avoid having to convert between `≤` and `<`. Instead, `isMax_iff_forall_not_lt` does the conversion. -/ def IsMax (a : α) : Prop := ∀ ⦃b⦄, a ≤ b → b ≤ a +#align is_max IsMax @[simp] theorem not_isBot [NoBotOrder α] (a : α) : ¬IsBot a := fun h => @@ -262,20 +270,28 @@ theorem isMax_ofDual_iff {a : αᵒᵈ} : IsMax (ofDual a) ↔ IsMin a := #align is_max_of_dual_iff isMax_ofDual_iff alias isBot_toDual_iff ↔ _ IsTop.toDual +#align is_top.to_dual IsTop.toDual alias isTop_toDual_iff ↔ _ IsBot.toDual +#align is_bot.to_dual IsBot.toDual alias isMin_toDual_iff ↔ _ IsMax.toDual +#align is_max.to_dual IsMax.toDual alias isMax_toDual_iff ↔ _ IsMin.toDual +#align is_min.to_dual IsMin.toDual alias isBot_ofDual_iff ↔ _ IsTop.ofDual +#align is_top.of_dual IsTop.ofDual alias isTop_ofDual_iff ↔ _ IsBot.ofDual +#align is_bot.of_dual IsBot.ofDual alias isMin_ofDual_iff ↔ _ IsMax.ofDual +#align is_max.of_dual IsMax.ofDual alias isMax_ofDual_iff ↔ _ IsMin.ofDual +#align is_min.of_dual IsMin.ofDual end LE diff --git a/Mathlib/Order/ModularLattice.lean b/Mathlib/Order/ModularLattice.lean index 453d5178783af..95109af3143cd 100644 --- a/Mathlib/Order/ModularLattice.lean +++ b/Mathlib/Order/ModularLattice.lean @@ -108,8 +108,10 @@ theorem covby_sup_of_inf_covby_of_inf_covby_right : a ⊓ b ⋖ a → a ⊓ b #align covby_sup_of_inf_covby_of_inf_covby_right covby_sup_of_inf_covby_of_inf_covby_right alias covby_sup_of_inf_covby_of_inf_covby_left ← Covby.sup_of_inf_of_inf_left +#align covby.sup_of_inf_of_inf_left Covby.sup_of_inf_of_inf_left alias covby_sup_of_inf_covby_of_inf_covby_right ← Covby.sup_of_inf_of_inf_right +#align covby.sup_of_inf_of_inf_right Covby.sup_of_inf_of_inf_right instance : IsWeakLowerModularLattice (OrderDual α) := ⟨fun ha hb => (ha.ofDual.sup_of_inf_of_inf_left hb.ofDual).toDual⟩ @@ -130,8 +132,10 @@ theorem inf_covby_of_covby_sup_of_covby_sup_right : a ⋖ a ⊔ b → b ⋖ a #align inf_covby_of_covby_sup_of_covby_sup_right inf_covby_of_covby_sup_of_covby_sup_right alias inf_covby_of_covby_sup_of_covby_sup_left ← Covby.inf_of_sup_of_sup_left +#align covby.inf_of_sup_of_sup_left Covby.inf_of_sup_of_sup_left alias inf_covby_of_covby_sup_of_covby_sup_right ← Covby.inf_of_sup_of_sup_right +#align covby.inf_of_sup_of_sup_right Covby.inf_of_sup_of_sup_right instance : IsWeakUpperModularLattice (OrderDual α) := ⟨fun ha hb => (ha.ofDual.inf_of_sup_of_sup_left hb.ofDual).toDual⟩ @@ -152,8 +156,10 @@ theorem covby_sup_of_inf_covby_right : a ⊓ b ⋖ b → a ⋖ a ⊔ b := by #align covby_sup_of_inf_covby_right covby_sup_of_inf_covby_right alias covby_sup_of_inf_covby_left ← Covby.sup_of_inf_left +#align covby.sup_of_inf_left Covby.sup_of_inf_left alias covby_sup_of_inf_covby_right ← Covby.sup_of_inf_right +#align covby.sup_of_inf_right Covby.sup_of_inf_right -- See note [lower instance priority] instance (priority := 100) IsUpperModularLattice.to_isWeakUpperModularLattice : @@ -180,8 +186,10 @@ theorem inf_covby_of_covby_sup_right : b ⋖ a ⊔ b → a ⊓ b ⋖ a := by #align inf_covby_of_covby_sup_right inf_covby_of_covby_sup_right alias inf_covby_of_covby_sup_left ← Covby.inf_of_sup_left +#align covby.inf_of_sup_left Covby.inf_of_sup_left alias inf_covby_of_covby_sup_right ← Covby.inf_of_sup_right +#align covby.inf_of_sup_right Covby.inf_of_sup_right -- See note [lower instance priority] instance (priority := 100) IsLowerModularLattice.to_isWeakLowerModularLattice : @@ -301,6 +309,8 @@ def infIccOrderIsoIccSup (a b : α) : Set.Icc (a ⊓ b) a ≃o Set.Icc b (a ⊔ sup_eq_right.2 y.prop.1, inf_sup_assoc_of_le _ y.prop.2, @sup_comm _ _ b] exact inf_le_inf_left _ h #align inf_Icc_order_iso_Icc_sup infIccOrderIsoIccSup +#align inf_Icc_order_iso_Icc_sup_apply_coe infIccOrderIsoIccSup_apply_coe +#align inf_Icc_order_iso_Icc_sup_symm_apply_coe infIccOrderIsoIccSup_symmApply_coe theorem inf_strictMonoOn_Icc_sup {a b : α} : StrictMonoOn (fun c => a ⊓ c) (Icc b (a ⊔ b)) := StrictMono.of_restrict (infIccOrderIsoIccSup a b).symm.strictMono @@ -335,6 +345,8 @@ def infIooOrderIsoIooSup (a b : α) : Ioo (a ⊓ b) a ≃o Ioo b (a ⊔ b) where @OrderIso.le_iff_le _ _ _ _ (infIccOrderIsoIccSup _ _) ⟨c.1, Ioo_subset_Icc_self c.2⟩ ⟨d.1, Ioo_subset_Icc_self d.2⟩ #align inf_Ioo_order_iso_Ioo_sup infIooOrderIsoIooSup +#align inf_Ioo_order_iso_Ioo_sup_apply_coe infIooOrderIsoIooSup_apply_coe +#align inf_Ioo_order_iso_Ioo_sup_symm_apply_coe infIooOrderIsoIooSup_symmApply_coe -- See note [lower instance priority] instance (priority := 100) IsModularLattice.to_isLowerModularLattice : IsLowerModularLattice α := diff --git a/Mathlib/Order/Monotone/Basic.lean b/Mathlib/Order/Monotone/Basic.lean index ac3904c540750..661e1c43e8222 100644 --- a/Mathlib/Order/Monotone/Basic.lean +++ b/Mathlib/Order/Monotone/Basic.lean @@ -79,36 +79,44 @@ variable [Preorder α] [Preorder β] /-- A function `f` is monotone if `a ≤ b` implies `f a ≤ f b`. -/ def Monotone (f : α → β) : Prop := ∀ ⦃a b⦄, a ≤ b → f a ≤ f b +#align monotone Monotone /-- A function `f` is antitone if `a ≤ b` implies `f b ≤ f a`. -/ def Antitone (f : α → β) : Prop := ∀ ⦃a b⦄, a ≤ b → f b ≤ f a +#align antitone Antitone /-- A function `f` is monotone on `s` if, for all `a, b ∈ s`, `a ≤ b` implies `f a ≤ f b`. -/ def MonotoneOn (f : α → β) (s : Set α) : Prop := ∀ ⦃a⦄ (_ : a ∈ s) ⦃b⦄ (_ : b ∈ s), a ≤ b → f a ≤ f b +#align monotone_on MonotoneOn /-- A function `f` is antitone on `s` if, for all `a, b ∈ s`, `a ≤ b` implies `f b ≤ f a`. -/ def AntitoneOn (f : α → β) (s : Set α) : Prop := ∀ ⦃a⦄ (_ : a ∈ s) ⦃b⦄ (_ : b ∈ s), a ≤ b → f b ≤ f a +#align antitone_on AntitoneOn /-- A function `f` is strictly monotone if `a < b` implies `f a < f b`. -/ def StrictMono (f : α → β) : Prop := ∀ ⦃a b⦄, a < b → f a < f b +#align strict_mono StrictMono /-- A function `f` is strictly antitone if `a < b` implies `f b < f a`. -/ def StrictAnti (f : α → β) : Prop := ∀ ⦃a b⦄, a < b → f b < f a +#align strict_anti StrictAnti /-- A function `f` is strictly monotone on `s` if, for all `a, b ∈ s`, `a < b` implies `f a < f b`. -/ def StrictMonoOn (f : α → β) (s : Set α) : Prop := ∀ ⦃a⦄ (_ : a ∈ s) ⦃b⦄ (_ : b ∈ s), a < b → f a < f b +#align strict_mono_on StrictMonoOn /-- A function `f` is strictly antitone on `s` if, for all `a, b ∈ s`, `a < b` implies `f b < f a`. -/ def StrictAntiOn (f : α → β) (s : Set α) : Prop := ∀ ⦃a⦄ (_ : a ∈ s) ⦃b⦄ (_ : b ∈ s), a < b → f b < f a +#align strict_anti_on StrictAntiOn end MonotoneDef @@ -213,65 +221,89 @@ theorem strictAntiOn_toDual_comp_iff : StrictAntiOn (toDual ∘ f : α → βᵒ protected theorem Monotone.dual (hf : Monotone f) : Monotone (toDual ∘ f ∘ ofDual : αᵒᵈ → βᵒᵈ) := swap hf +#align monotone.dual Monotone.dual protected theorem Antitone.dual (hf : Antitone f) : Antitone (toDual ∘ f ∘ ofDual : αᵒᵈ → βᵒᵈ) := swap hf +#align antitone.dual Antitone.dual protected theorem MonotoneOn.dual (hf : MonotoneOn f s) : MonotoneOn (toDual ∘ f ∘ ofDual : αᵒᵈ → βᵒᵈ) s := swap₂ hf +#align monotone_on.dual MonotoneOn.dual protected theorem AntitoneOn.dual (hf : AntitoneOn f s) : AntitoneOn (toDual ∘ f ∘ ofDual : αᵒᵈ → βᵒᵈ) s := swap₂ hf +#align antitone_on.dual AntitoneOn.dual protected theorem StrictMono.dual (hf : StrictMono f) : StrictMono (toDual ∘ f ∘ ofDual : αᵒᵈ → βᵒᵈ) := swap hf +#align strict_mono.dual StrictMono.dual protected theorem StrictAnti.dual (hf : StrictAnti f) : StrictAnti (toDual ∘ f ∘ ofDual : αᵒᵈ → βᵒᵈ) := swap hf +#align strict_anti.dual StrictAnti.dual protected theorem StrictMonoOn.dual (hf : StrictMonoOn f s) : StrictMonoOn (toDual ∘ f ∘ ofDual : αᵒᵈ → βᵒᵈ) s := swap₂ hf +#align strict_mono_on.dual StrictMonoOn.dual protected theorem StrictAntiOn.dual (hf : StrictAntiOn f s) : StrictAntiOn (toDual ∘ f ∘ ofDual : αᵒᵈ → βᵒᵈ) s := swap₂ hf +#align strict_anti_on.dual StrictAntiOn.dual alias antitone_comp_ofDual_iff ↔ _ Monotone.dual_left +#align monotone.dual_left Monotone.dual_left alias monotone_comp_ofDual_iff ↔ _ Antitone.dual_left +#align antitone.dual_left Antitone.dual_left alias antitone_toDual_comp_iff ↔ _ Monotone.dual_right +#align monotone.dual_right Monotone.dual_right alias monotone_toDual_comp_iff ↔ _ Antitone.dual_right +#align antitone.dual_right Antitone.dual_right alias antitoneOn_comp_ofDual_iff ↔ _ MonotoneOn.dual_left +#align monotone_on.dual_left MonotoneOn.dual_left alias monotoneOn_comp_ofDual_iff ↔ _ AntitoneOn.dual_left +#align antitone_on.dual_left AntitoneOn.dual_left alias antitoneOn_toDual_comp_iff ↔ _ MonotoneOn.dual_right +#align monotone_on.dual_right MonotoneOn.dual_right alias monotoneOn_toDual_comp_iff ↔ _ AntitoneOn.dual_right +#align antitone_on.dual_right AntitoneOn.dual_right alias strictAnti_comp_ofDual_iff ↔ _ StrictMono.dual_left +#align strict_mono.dual_left StrictMono.dual_left alias strictMono_comp_ofDual_iff ↔ _ StrictAnti.dual_left +#align strict_anti.dual_left StrictAnti.dual_left alias strictAnti_toDual_comp_iff ↔ _ StrictMono.dual_right +#align strict_mono.dual_right StrictMono.dual_right alias strictMono_toDual_comp_iff ↔ _ StrictAnti.dual_right +#align strict_anti.dual_right StrictAnti.dual_right alias strictAntiOn_comp_ofDual_iff ↔ _ StrictMonoOn.dual_left +#align strict_mono_on.dual_left StrictMonoOn.dual_left alias strictMonoOn_comp_ofDual_iff ↔ _ StrictAntiOn.dual_left +#align strict_anti_on.dual_left StrictAntiOn.dual_left alias strictAntiOn_toDual_comp_iff ↔ _ StrictMonoOn.dual_right +#align strict_mono_on.dual_right StrictMonoOn.dual_right alias strictMonoOn_toDual_comp_iff ↔ _ StrictAntiOn.dual_right +#align strict_anti_on.dual_right StrictAntiOn.dual_right end OrderDual @@ -286,25 +318,31 @@ theorem Monotone.comp_le_comp_left [Preorder β] {f : β → α} {g h : γ → β} (hf : Monotone f) (le_gh : g ≤ h) : LE.le.{max w u} (f ∘ g) (f ∘ h) := fun x ↦ hf (le_gh x) +#align monotone.comp_le_comp_left Monotone.comp_le_comp_left variable [Preorder γ] theorem monotone_lam {f : α → β → γ} (hf : ∀ b, Monotone fun a ↦ f a b) : Monotone f := fun _ _ h b ↦ hf b h +#align monotone_lam monotone_lam theorem monotone_app (f : β → α → γ) (b : β) (hf : Monotone fun a b ↦ f b a) : Monotone (f b) := fun _ _ h ↦ hf h b +#align monotone_app monotone_app theorem antitone_lam {f : α → β → γ} (hf : ∀ b, Antitone fun a ↦ f a b) : Antitone f := fun _ _ h b ↦ hf b h +#align antitone_lam antitone_lam theorem antitone_app (f : β → α → γ) (b : β) (hf : Antitone fun a b ↦ f b a) : Antitone (f b) := fun _ _ h ↦ hf h b +#align antitone_app antitone_app end Preorder theorem Function.monotone_eval {ι : Type u} {α : ι → Type v} [∀ i, Preorder (α i)] (i : ι) : Monotone (Function.eval i : (∀ i, α i) → α i) := fun _ _ H ↦ H i +#align function.monotone_eval Function.monotone_eval /-! ### Monotonicity hierarchy -/ @@ -326,21 +364,27 @@ However if you find yourself writing `hf.imp h`, then you should have written `h theorem Monotone.imp (hf : Monotone f) (h : a ≤ b) : f a ≤ f b := hf h +#align monotone.imp Monotone.imp theorem Antitone.imp (hf : Antitone f) (h : a ≤ b) : f b ≤ f a := hf h +#align antitone.imp Antitone.imp theorem StrictMono.imp (hf : StrictMono f) (h : a < b) : f a < f b := hf h +#align strict_mono.imp StrictMono.imp theorem StrictAnti.imp (hf : StrictAnti f) (h : a < b) : f b < f a := hf h +#align strict_anti.imp StrictAnti.imp protected theorem Monotone.monotoneOn (hf : Monotone f) (s : Set α) : MonotoneOn f s := fun _ _ _ _ ↦ hf.imp +#align monotone.monotone_on Monotone.monotoneOn protected theorem Antitone.antitoneOn (hf : Antitone f) (s : Set α) : AntitoneOn f s := fun _ _ _ _ ↦ hf.imp +#align antitone.antitone_on Antitone.antitoneOn @[simp] theorem monotoneOn_univ : MonotoneOn f Set.univ ↔ Monotone f := ⟨fun h _ _ ↦ h trivial trivial, fun h ↦ h.monotoneOn _⟩ @@ -391,10 +435,12 @@ variable [PartialOrder α] [Preorder β] {f : α → β} {s : Set α} theorem monotone_iff_forall_lt : Monotone f ↔ ∀ ⦃a b⦄, a < b → f a ≤ f b := forall₂_congr fun _ _ ↦ ⟨fun hf h ↦ hf h.le, fun hf h ↦ h.eq_or_lt.elim (fun H ↦ (congr_arg _ H).le) hf⟩ +#align monotone_iff_forall_lt monotone_iff_forall_lt theorem antitone_iff_forall_lt : Antitone f ↔ ∀ ⦃a b⦄, a < b → f b ≤ f a := forall₂_congr fun _ _ ↦ ⟨fun hf h ↦ hf h.le, fun hf h ↦ h.eq_or_lt.elim (fun H ↦ (congr_arg _ H).ge) hf⟩ +#align antitone_iff_forall_lt antitone_iff_forall_lt theorem monotoneOn_iff_forall_lt : MonotoneOn f s ↔ ∀ ⦃a⦄ (_ : a ∈ s) ⦃b⦄ (_ : b ∈ s), a < b → f a ≤ f b := @@ -420,9 +466,11 @@ protected theorem StrictAntiOn.antitoneOn (hf : StrictAntiOn f s) : AntitoneOn f protected theorem StrictMono.monotone (hf : StrictMono f) : Monotone f := monotone_iff_forall_lt.2 fun _ _ h ↦ (hf h).le +#align strict_mono.monotone StrictMono.monotone protected theorem StrictAnti.antitone (hf : StrictAnti f) : Antitone f := antitone_iff_forall_lt.2 fun _ _ h ↦ (hf h).le +#align strict_anti.antitone StrictAnti.antitone end PartialOrder @@ -435,15 +483,19 @@ variable [Preorder α] [Preorder β] protected theorem monotone [Subsingleton α] (f : α → β) : Monotone f := fun _ _ _ ↦ (congr_arg _ <| Subsingleton.elim _ _).le +#align subsingleton.monotone Subsingleton.monotone protected theorem antitone [Subsingleton α] (f : α → β) : Antitone f := fun _ _ _ ↦ (congr_arg _ <| Subsingleton.elim _ _).le +#align subsingleton.antitone Subsingleton.antitone theorem monotone' [Subsingleton β] (f : α → β) : Monotone f := fun _ _ _ ↦ (Subsingleton.elim _ _).le +#align subsingleton.monotone' Subsingleton.monotone' theorem antitone' [Subsingleton β] (f : α → β) : Antitone f := fun _ _ _ ↦ (Subsingleton.elim _ _).le +#align subsingleton.antitone' Subsingleton.antitone' protected theorem strictMono [Subsingleton α] (f : α → β) : StrictMono f := fun _ _ h ↦ (h.ne <| Subsingleton.elim _ _).elim @@ -459,6 +511,7 @@ end Subsingleton theorem monotone_id [Preorder α] : Monotone (id : α → α) := fun _ _ ↦ id +#align monotone_id monotone_id theorem monotoneOn_id [Preorder α] {s : Set α} : MonotoneOn id s := fun _ _ _ _ ↦ id #align monotone_on_id monotoneOn_id @@ -471,6 +524,7 @@ theorem strictMonoOn_id [Preorder α] {s : Set α} : StrictMonoOn id s := fun _ theorem monotone_const [Preorder α] [Preorder β] {c : β} : Monotone fun _ : α ↦ c := fun _ _ _ ↦ le_rfl +#align monotone_const monotone_const theorem monotoneOn_const [Preorder α] [Preorder β] {c : β} {s : Set α} : MonotoneOn (fun _ : α ↦ c) s := @@ -479,6 +533,7 @@ theorem monotoneOn_const [Preorder α] [Preorder β] {c : β} {s : Set α} : theorem antitone_const [Preorder α] [Preorder β] {c : β} : Antitone fun _ : α ↦ c := fun _ _ _ ↦ le_refl c +#align antitone_const antitone_const theorem antitoneOn_const [Preorder α] [Preorder β] {c : β} {s : Set α} : AntitoneOn (fun _ : α ↦ c) s := @@ -508,6 +563,7 @@ theorem injective_of_lt_imp_ne [LinearOrder α] {f : α → β} (h : ∀ x y, x theorem injective_of_le_imp_le [PartialOrder α] [Preorder β] (f : α → β) (h : ∀ {x y}, f x ≤ f y → x ≤ y) : Injective f := fun _ _ hxy ↦ (h hxy.le).antisymm (h hxy.ge) +#align injective_of_le_imp_le injective_of_le_imp_le section Preorder @@ -548,11 +604,13 @@ protected theorem StrictMono.ite' (hf : StrictMono f) (hg : StrictMono g) {p : by_cases hx:p x · simpa [hx, hy] using hfg hx hy h · simpa [hx, hy] using hg h +#align strict_mono.ite' StrictMono.ite' protected theorem StrictMono.ite (hf : StrictMono f) (hg : StrictMono g) {p : α → Prop} [DecidablePred p] (hp : ∀ ⦃x y⦄, x < y → p y → p x) (hfg : ∀ x, f x ≤ g x) : StrictMono fun x ↦ if p x then f x else g x := (hf.ite' hg hp) fun _ y _ _ h ↦ (hf h).trans_le (hfg y) +#align strict_mono.ite StrictMono.ite -- Porting note: `Strict*.dual_right` dot notation is not working here for some reason protected theorem StrictAnti.ite' (hf : StrictAnti f) (hg : StrictAnti g) {p : α → Prop} @@ -560,6 +618,7 @@ protected theorem StrictAnti.ite' (hf : StrictAnti f) (hg : StrictAnti g) {p : (hp : ∀ ⦃x y⦄, x < y → p y → p x) (hfg : ∀ ⦃x y⦄, p x → ¬p y → x < y → g y < f x) : StrictAnti fun x ↦ if p x then f x else g x := StrictMono.ite' (StrictAnti.dual_right hf) (StrictAnti.dual_right hg) hp hfg +#align strict_anti.ite' StrictAnti.ite' protected theorem StrictAnti.ite (hf : StrictAnti f) (hg : StrictAnti g) {p : α → Prop} [DecidablePred p] (hp : ∀ ⦃x y⦄, x < y → p y → p x) (hfg : ∀ x, g x ≤ f x) : @@ -578,15 +637,19 @@ variable [Preorder α] [Preorder β] [Preorder γ] {g : β → γ} {f : α → protected theorem Monotone.comp (hg : Monotone g) (hf : Monotone f) : Monotone (g ∘ f) := fun _ _ h ↦ hg (hf h) +#align monotone.comp Monotone.comp theorem Monotone.comp_antitone (hg : Monotone g) (hf : Antitone f) : Antitone (g ∘ f) := fun _ _ h ↦ hg (hf h) +#align monotone.comp_antitone Monotone.comp_antitone protected theorem Antitone.comp (hg : Antitone g) (hf : Antitone f) : Monotone (g ∘ f) := fun _ _ h ↦ hg (hf h) +#align antitone.comp Antitone.comp theorem Antitone.comp_monotone (hg : Antitone g) (hf : Monotone f) : Antitone (g ∘ f) := fun _ _ h ↦ hg (hf h) +#align antitone.comp_monotone Antitone.comp_monotone protected theorem Monotone.iterate {f : α → α} (hf : Monotone f) (n : ℕ) : Monotone (f^[n]) := Nat.recOn n monotone_id fun _ h ↦ h.comp hf @@ -612,6 +675,7 @@ theorem Antitone.comp_monotoneOn (hg : Antitone g) (hf : MonotoneOn f s) : Antit protected theorem StrictMono.comp (hg : StrictMono g) (hf : StrictMono f) : StrictMono (g ∘ f) := fun _ _ h ↦ hg (hf h) +#align strict_mono.comp StrictMono.comp theorem StrictMono.comp_strictAnti (hg : StrictMono g) (hf : StrictAnti f) : StrictAnti (g ∘ f) := fun _ _ h ↦ hg (hf h) @@ -619,6 +683,7 @@ theorem StrictMono.comp_strictAnti (hg : StrictMono g) (hf : StrictAnti f) : Str protected theorem StrictAnti.comp (hg : StrictAnti g) (hf : StrictAnti f) : StrictMono (g ∘ f) := fun _ _ h ↦ hg (hf h) +#align strict_anti.comp StrictAnti.comp theorem StrictAnti.comp_strictMono (hg : StrictAnti g) (hf : StrictMono f) : StrictAnti (g ∘ f) := fun _ _ h ↦ hg (hf h) @@ -626,6 +691,7 @@ theorem StrictAnti.comp_strictMono (hg : StrictAnti g) (hf : StrictMono f) : Str protected theorem StrictMono.iterate {f : α → α} (hf : StrictMono f) (n : ℕ) : StrictMono (f^[n]) := Nat.recOn n strictMono_id fun _ h ↦ h.comp hf +#align strict_mono.iterate StrictMono.iterate protected theorem StrictMono.comp_strictMonoOn (hg : StrictMono g) (hf : StrictMonoOn f s) : StrictMonoOn (g ∘ f) s := @@ -656,9 +722,11 @@ section Fold theorem foldl_monotone [Preorder α] {f : α → β → α} (H : ∀ b, Monotone fun a ↦ f a b) (l : List β) : Monotone fun a ↦ l.foldl f a := List.recOn l (fun _ _ ↦ id) fun _ _ hl _ _ h ↦ hl (H _ h) +#align list.foldl_monotone List.foldl_monotone theorem foldr_monotone [Preorder β] {f : α → β → β} (H : ∀ a, Monotone (f a)) (l : List α) : Monotone fun b ↦ l.foldr f b := fun _ _ h ↦ List.recOn l h fun i _ hl ↦ H i hl +#align list.foldr_monotone List.foldr_monotone theorem foldl_strictMono [Preorder α] {f : α → β → α} (H : ∀ b, StrictMono fun a ↦ f a b) (l : List β) : StrictMono fun a ↦ l.foldl f a := @@ -688,50 +756,62 @@ open Ordering theorem Monotone.reflect_lt (hf : Monotone f) {a b : α} (h : f a < f b) : a < b := lt_of_not_ge fun h' ↦ h.not_le (hf h') +#align monotone.reflect_lt Monotone.reflect_lt theorem Antitone.reflect_lt (hf : Antitone f) {a b : α} (h : f a < f b) : b < a := lt_of_not_ge fun h' ↦ h.not_le (hf h') +#align antitone.reflect_lt Antitone.reflect_lt theorem MonotoneOn.reflect_lt (hf : MonotoneOn f s) {a b : α} (ha : a ∈ s) (hb : b ∈ s) (h : f a < f b) : a < b := lt_of_not_ge fun h' ↦ h.not_le <| hf hb ha h' +#align monotone_on.reflect_lt MonotoneOn.reflect_lt theorem AntitoneOn.reflect_lt (hf : AntitoneOn f s) {a b : α} (ha : a ∈ s) (hb : b ∈ s) (h : f a < f b) : b < a := lt_of_not_ge fun h' ↦ h.not_le <| hf ha hb h' +#align antitone_on.reflect_lt AntitoneOn.reflect_lt theorem StrictMonoOn.le_iff_le (hf : StrictMonoOn f s) {a b : α} (ha : a ∈ s) (hb : b ∈ s) : f a ≤ f b ↔ a ≤ b := ⟨fun h ↦ le_of_not_gt fun h' ↦ (hf hb ha h').not_le h, fun h ↦ h.lt_or_eq_dec.elim (fun h' ↦ (hf ha hb h').le) fun h' ↦ h' ▸ le_rfl⟩ +#align strict_mono_on.le_iff_le StrictMonoOn.le_iff_le theorem StrictAntiOn.le_iff_le (hf : StrictAntiOn f s) {a b : α} (ha : a ∈ s) (hb : b ∈ s) : f a ≤ f b ↔ b ≤ a := hf.dual_right.le_iff_le hb ha +#align strict_anti_on.le_iff_le StrictAntiOn.le_iff_le theorem StrictMonoOn.eq_iff_eq (hf : StrictMonoOn f s) {a b : α} (ha : a ∈ s) (hb : b ∈ s) : f a = f b ↔ a = b := ⟨fun h ↦ le_antisymm ((hf.le_iff_le ha hb).mp h.le) ((hf.le_iff_le hb ha).mp h.ge), by rintro rfl rfl⟩ +#align strict_mono_on.eq_iff_eq StrictMonoOn.eq_iff_eq theorem StrictAntiOn.eq_iff_eq (hf : StrictAntiOn f s) {a b : α} (ha : a ∈ s) (hb : b ∈ s) : f a = f b ↔ b = a := (hf.dual_right.eq_iff_eq ha hb).trans eq_comm +#align strict_anti_on.eq_iff_eq StrictAntiOn.eq_iff_eq theorem StrictMonoOn.lt_iff_lt (hf : StrictMonoOn f s) {a b : α} (ha : a ∈ s) (hb : b ∈ s) : f a < f b ↔ a < b := by rw [lt_iff_le_not_le, lt_iff_le_not_le, hf.le_iff_le ha hb, hf.le_iff_le hb ha] +#align strict_mono_on.lt_iff_lt StrictMonoOn.lt_iff_lt theorem StrictAntiOn.lt_iff_lt (hf : StrictAntiOn f s) {a b : α} (ha : a ∈ s) (hb : b ∈ s) : f a < f b ↔ b < a := hf.dual_right.lt_iff_lt hb ha +#align strict_anti_on.lt_iff_lt StrictAntiOn.lt_iff_lt theorem StrictMono.le_iff_le (hf : StrictMono f) {a b : α} : f a ≤ f b ↔ a ≤ b := (hf.strictMonoOn Set.univ).le_iff_le trivial trivial +#align strict_mono.le_iff_le StrictMono.le_iff_le theorem StrictAnti.le_iff_le (hf : StrictAnti f) {a b : α} : f a ≤ f b ↔ b ≤ a := (hf.strictAntiOn Set.univ).le_iff_le trivial trivial +#align strict_anti.le_iff_le StrictAnti.le_iff_le theorem StrictMono.lt_iff_lt (hf : StrictMono f) {a b : α} : f a < f b ↔ a < b := (hf.strictMonoOn Set.univ).lt_iff_lt trivial trivial @@ -739,6 +819,7 @@ theorem StrictMono.lt_iff_lt (hf : StrictMono f) {a b : α} : f a < f b ↔ a < theorem StrictAnti.lt_iff_lt (hf : StrictAnti f) {a b : α} : f a < f b ↔ b < a := (hf.strictAntiOn Set.univ).lt_iff_lt trivial trivial +#align strict_anti.lt_iff_lt StrictAnti.lt_iff_lt protected theorem StrictMonoOn.compares (hf : StrictMonoOn f s) {a b : α} (ha : a ∈ s) (hb : b ∈ s) : ∀ {o : Ordering}, o.Compares (f a) (f b) ↔ o.Compares a b @@ -751,14 +832,17 @@ protected theorem StrictMonoOn.compares (hf : StrictMonoOn f s) {a b : α} (ha : protected theorem StrictAntiOn.compares (hf : StrictAntiOn f s) {a b : α} (ha : a ∈ s) (hb : b ∈ s) {o : Ordering} : o.Compares (f a) (f b) ↔ o.Compares b a := toDual_compares_toDual.trans <| hf.dual_right.compares hb ha +#align strict_anti_on.compares StrictAntiOn.compares protected theorem StrictMono.compares (hf : StrictMono f) {a b : α} {o : Ordering} : o.Compares (f a) (f b) ↔ o.Compares a b := (hf.strictMonoOn Set.univ).compares trivial trivial +#align strict_mono.compares StrictMono.compares protected theorem StrictAnti.compares (hf : StrictAnti f) {a b : α} {o : Ordering} : o.Compares (f a) (f b) ↔ o.Compares b a := (hf.strictAntiOn Set.univ).compares trivial trivial +#align strict_anti.compares StrictAnti.compares theorem StrictMono.injective (hf : StrictMono f) : Injective f := fun x y h ↦ show Compares eq x y from hf.compares.1 h @@ -771,18 +855,22 @@ theorem StrictAnti.injective (hf : StrictAnti f) : Injective f := theorem StrictMono.maximal_of_maximal_image (hf : StrictMono f) {a} (hmax : ∀ p, p ≤ f a) (x : α) : x ≤ a := hf.le_iff_le.mp (hmax (f x)) +#align strict_mono.maximal_of_maximal_image StrictMono.maximal_of_maximal_image theorem StrictMono.minimal_of_minimal_image (hf : StrictMono f) {a} (hmin : ∀ p, f a ≤ p) (x : α) : a ≤ x := hf.le_iff_le.mp (hmin (f x)) +#align strict_mono.minimal_of_minimal_image StrictMono.minimal_of_minimal_image theorem StrictAnti.minimal_of_maximal_image (hf : StrictAnti f) {a} (hmax : ∀ p, p ≤ f a) (x : α) : a ≤ x := hf.le_iff_le.mp (hmax (f x)) +#align strict_anti.minimal_of_maximal_image StrictAnti.minimal_of_maximal_image theorem StrictAnti.maximal_of_minimal_image (hf : StrictAnti f) {a} (hmin : ∀ p, f a ≤ p) (x : α) : x ≤ a := hf.le_iff_le.mp (hmin (f x)) +#align strict_anti.maximal_of_minimal_image StrictAnti.maximal_of_minimal_image end Preorder @@ -831,6 +919,7 @@ lemma not_monotone_not_antitone_iff_exists_le_le : obtain hbd | hdb := le_total b d { exact ⟨a, b, d, hab, hbd, Or.inr ⟨hfba, hfbd⟩⟩ } { exact ⟨a, d, b, had, hdb, Or.inl ⟨hfac.trans_lt hfcd, hfbd⟩⟩ } } +#align not_monotone_not_antitone_iff_exists_le_le not_monotone_not_antitone_iff_exists_le_le /-- A function between linear orders which is neither monotone nor antitone makes a dent upright or downright. -/ @@ -841,6 +930,7 @@ lemma not_monotone_not_antitone_iff_exists_lt_lt : refine' exists₃_congr (fun a b c ↦ and_congr_left $ fun h ↦ (Ne.le_iff_lt _).and $ Ne.le_iff_lt _) <;> (rintro rfl; simp at h) +#align not_monotone_not_antitone_iff_exists_lt_lt not_monotone_not_antitone_iff_exists_lt_lt /-! ### Strictly monotone functions and `cmp` @@ -852,16 +942,20 @@ variable [LinearOrder β] {f : α → β} {s : Set α} {x y : α} theorem StrictMonoOn.cmp_map_eq (hf : StrictMonoOn f s) (hx : x ∈ s) (hy : y ∈ s) : cmp (f x) (f y) = cmp x y := ((hf.compares hx hy).2 (cmp_compares x y)).cmp_eq +#align strict_mono_on.cmp_map_eq StrictMonoOn.cmp_map_eq theorem StrictMono.cmp_map_eq (hf : StrictMono f) (x y : α) : cmp (f x) (f y) = cmp x y := (hf.strictMonoOn Set.univ).cmp_map_eq trivial trivial +#align strict_mono.cmp_map_eq StrictMono.cmp_map_eq theorem StrictAntiOn.cmp_map_eq (hf : StrictAntiOn f s) (hx : x ∈ s) (hy : y ∈ s) : cmp (f x) (f y) = cmp y x := hf.dual_right.cmp_map_eq hy hx +#align strict_anti_on.cmp_map_eq StrictAntiOn.cmp_map_eq theorem StrictAnti.cmp_map_eq (hf : StrictAnti f) (x y : α) : cmp (f x) (f y) = cmp y x := (hf.strictAntiOn Set.univ).cmp_map_eq trivial trivial +#align strict_anti.cmp_map_eq StrictAnti.cmp_map_eq end LinearOrder @@ -909,6 +1003,7 @@ theorem strictMono_nat_of_lt_succ {f : ℕ → α} (hf : ∀ n, f n < f (n + 1)) theorem strictAnti_nat_of_succ_lt {f : ℕ → α} (hf : ∀ n, f (n + 1) < f n) : StrictAnti f := @strictMono_nat_of_lt_succ αᵒᵈ _ f hf +#align strict_anti_nat_of_succ_lt strictAnti_nat_of_succ_lt namespace Nat @@ -952,16 +1047,20 @@ theorem Int.rel_of_forall_rel_succ_of_lt (r : β → β → Prop) [IsTrans β r] apply h · rw [Int.ofNat_succ, ← Int.add_assoc] exact _root_.trans ihn (h _) +#align int.rel_of_forall_rel_succ_of_lt Int.rel_of_forall_rel_succ_of_lt theorem Int.rel_of_forall_rel_succ_of_le (r : β → β → Prop) [IsRefl β r] [IsTrans β r] {f : ℤ → β} (h : ∀ n, r (f n) (f (n + 1))) ⦃a b : ℤ⦄ (hab : a ≤ b) : r (f a) (f b) := hab.eq_or_lt.elim (fun h ↦ h ▸ refl _) fun h' ↦ Int.rel_of_forall_rel_succ_of_lt r h h' +#align int.rel_of_forall_rel_succ_of_le Int.rel_of_forall_rel_succ_of_le theorem monotone_int_of_le_succ {f : ℤ → α} (hf : ∀ n, f n ≤ f (n + 1)) : Monotone f := Int.rel_of_forall_rel_succ_of_le (· ≤ ·) hf +#align monotone_int_of_le_succ monotone_int_of_le_succ theorem antitone_int_of_succ_le {f : ℤ → α} (hf : ∀ n, f (n + 1) ≤ f n) : Antitone f := Int.rel_of_forall_rel_succ_of_le (· ≥ ·) hf +#align antitone_int_of_succ_le antitone_int_of_succ_le theorem strictMono_int_of_lt_succ {f : ℤ → α} (hf : ∀ n, f n < f (n + 1)) : StrictMono f := Int.rel_of_forall_rel_succ_of_lt (· < ·) hf @@ -1006,6 +1105,7 @@ theorem Monotone.ne_of_lt_of_lt_nat {f : ℕ → α} (hf : Monotone f) (n : ℕ) (h2 : x < f (n + 1)) (a : ℕ) : f a ≠ x := by rintro rfl exact (hf.reflect_lt h1).not_le (Nat.le_of_lt_succ <| hf.reflect_lt h2) +#align monotone.ne_of_lt_of_lt_nat Monotone.ne_of_lt_of_lt_nat /-- If `f` is an antitone function from `ℕ` to a preorder such that `x` lies between `f (n + 1)` and `f n`, then `x` doesn't lie in the range of `f`. -/ @@ -1013,6 +1113,7 @@ theorem Antitone.ne_of_lt_of_lt_nat {f : ℕ → α} (hf : Antitone f) (n : ℕ) (h1 : f (n + 1) < x) (h2 : x < f n) (a : ℕ) : f a ≠ x := by rintro rfl exact (hf.reflect_lt h2).not_le (Nat.le_of_lt_succ <| hf.reflect_lt h1) +#align antitone.ne_of_lt_of_lt_nat Antitone.ne_of_lt_of_lt_nat /-- If `f` is a monotone function from `ℤ` to a preorder and `x` lies between `f n` and `f (n + 1)`, then `x` doesn't lie in the range of `f`. -/ @@ -1020,6 +1121,7 @@ theorem Monotone.ne_of_lt_of_lt_int {f : ℤ → α} (hf : Monotone f) (n : ℤ) (h2 : x < f (n + 1)) (a : ℤ) : f a ≠ x := by rintro rfl exact (hf.reflect_lt h1).not_le (Int.le_of_lt_add_one <| hf.reflect_lt h2) +#align monotone.ne_of_lt_of_lt_int Monotone.ne_of_lt_of_lt_int /-- If `f` is an antitone function from `ℤ` to a preorder and `x` lies between `f (n + 1)` and `f n`, then `x` doesn't lie in the range of `f`. -/ @@ -1027,14 +1129,17 @@ theorem Antitone.ne_of_lt_of_lt_int {f : ℤ → α} (hf : Antitone f) (n : ℤ) (h1 : f (n + 1) < x) (h2 : x < f n) (a : ℤ) : f a ≠ x := by rintro rfl exact (hf.reflect_lt h2).not_le (Int.le_of_lt_add_one <| hf.reflect_lt h1) +#align antitone.ne_of_lt_of_lt_int Antitone.ne_of_lt_of_lt_int theorem StrictMono.id_le {φ : ℕ → ℕ} (h : StrictMono φ) : ∀ n, n ≤ φ n := fun n ↦ Nat.recOn n (Nat.zero_le _) fun n hn ↦ Nat.succ_le_of_lt (hn.trans_lt <| h <| Nat.lt_succ_self n) +#align strict_mono.id_le StrictMono.id_le end Preorder theorem Subtype.mono_coe [Preorder α] (t : Set α) : Monotone ((↑) : Subtype t → α) := fun _ _ ↦ id +#align subtype.mono_coe Subtype.mono_coe theorem Subtype.strictMono_coe [Preorder α] (t : Set α) : StrictMono ((↑) : Subtype t → α) := @@ -1046,14 +1151,18 @@ section Preorder variable [Preorder α] [Preorder β] [Preorder γ] [Preorder δ] {f : α → γ} {g : β → δ} {a b : α} theorem monotone_fst : Monotone (@Prod.fst α β) := fun _ _ ↦ And.left +#align monotone_fst monotone_fst theorem monotone_snd : Monotone (@Prod.snd α β) := fun _ _ ↦ And.right +#align monotone_snd monotone_snd theorem Monotone.prod_map (hf : Monotone f) (hg : Monotone g) : Monotone (Prod.map f g) := fun _ _ h ↦ ⟨hf h.1, hg h.2⟩ +#align monotone.prod_map Monotone.prod_map theorem Antitone.prod_map (hf : Antitone f) (hg : Antitone g) : Antitone (Prod.map f g) := fun _ _ h ↦ ⟨hf h.1, hg h.2⟩ +#align antitone.prod_map Antitone.prod_map end Preorder @@ -1065,11 +1174,13 @@ theorem StrictMono.prod_map (hf : StrictMono f) (hg : StrictMono g) : StrictMono fun a b ↦ by simp only [Prod.lt_iff] exact Or.imp (And.imp hf.imp hg.monotone.imp) (And.imp hf.monotone.imp hg.imp) +#align strict_mono.prod_map StrictMono.prod_map theorem StrictAnti.prod_map (hf : StrictAnti f) (hg : StrictAnti g) : StrictAnti (Prod.map f g) := fun a b ↦ by simp only [Prod.lt_iff] exact Or.imp (And.imp hf.imp hg.antitone.imp) (And.imp hf.antitone.imp hg.imp) +#align strict_anti.prod_map StrictAnti.prod_map end PartialOrder @@ -1078,6 +1189,7 @@ namespace Function variable [Preorder α] theorem const_mono : Monotone (const β : α → β → α) := fun _ _ h _ ↦ h +#align function.const_mono Function.const_mono theorem const_strictMono [Nonempty β] : StrictMono (const β : α → β → α) := fun _ _ ↦ const_lt_const.2 diff --git a/Mathlib/Order/OmegaCompletePartialOrder.lean b/Mathlib/Order/OmegaCompletePartialOrder.lean index b6f98fba740f4..0f9fc285c143f 100644 --- a/Mathlib/Order/OmegaCompletePartialOrder.lean +++ b/Mathlib/Order/OmegaCompletePartialOrder.lean @@ -79,6 +79,7 @@ def bind {β γ} (f : α →o Part β) (g : α →o β → Part γ) : α →o Pa intro b hb ha refine' ⟨b, f.monotone h _ hb, g.monotone h _ _ ha⟩ #align order_hom.bind OrderHom.bind +#align order_hom.bind_coe OrderHom.bind_coe end OrderHom @@ -115,6 +116,7 @@ instance : LE (Chain α) where le x y := ∀ i, ∃ j, x i ≤ y j def map : Chain β := f.comp c #align omega_complete_partial_order.chain.map OmegaCompletePartialOrder.Chain.map +#align omega_complete_partial_order.chain.map_coe OmegaCompletePartialOrder.Chain.map_coe variable {f} @@ -155,6 +157,7 @@ theorem map_le_map {g : α →o β} (h : f ≤ g) : c.map f ≤ c.map g := def zip (c₀ : Chain α) (c₁ : Chain β) : Chain (α × β) := OrderHom.prod c₀ c₁ #align omega_complete_partial_order.chain.zip OmegaCompletePartialOrder.Chain.zip +#align omega_complete_partial_order.chain.zip_coe OmegaCompletePartialOrder.Chain.zip_coe end Chain @@ -453,6 +456,8 @@ variable [OmegaCompletePartialOrder γ] protected def ωSup (c : Chain (α × β)) : α × β := (ωSup (c.map OrderHom.fst), ωSup (c.map OrderHom.snd)) #align prod.ωSup Prod.ωSup +#align prod.ωSup_snd Prod.ωSup_snd +#align prod.ωSup_fst Prod.ωSup_fst @[simps ωSup_fst ωSup_snd] instance : OmegaCompletePartialOrder (α × β) where @@ -562,6 +567,7 @@ protected def ωSup (c : Chain (α →o β)) : α →o β where toFun a := ωSup (c.map (OrderHom.apply a)) monotone' _ _ h := ωSup_le_ωSup_of_le ((Chain.map_le_map _) fun a => a.monotone h) #align omega_complete_partial_order.order_hom.ωSup OmegaCompletePartialOrder.OrderHom.ωSup +#align omega_complete_partial_order.order_hom.ωSup_coe OmegaCompletePartialOrder.OrderHom.ωSup_coe @[simps ωSup_coe] instance omegaCompletePartialOrder : OmegaCompletePartialOrder (α →o β) := @@ -569,6 +575,7 @@ instance omegaCompletePartialOrder : OmegaCompletePartialOrder (α →o β) := #align omega_complete_partial_order.order_hom.omega_complete_partial_order OmegaCompletePartialOrder.OrderHom.omegaCompletePartialOrder +#align omega_complete_partial_order.order_hom.omega_complete_partial_order_ωSup_coe OmegaCompletePartialOrder.OrderHom.omegaCompletePartialOrder_ωSup_coe end OrderHom @@ -736,12 +743,14 @@ def ofMono (f : α →o β) (h : ∀ c : Chain α, f (ωSup c) = ωSup (c.map f) def id : α →𝒄 α := ofMono OrderHom.id continuous_id #align omega_complete_partial_order.continuous_hom.id OmegaCompletePartialOrder.ContinuousHom.id +#align omega_complete_partial_order.continuous_hom.id_apply OmegaCompletePartialOrder.ContinuousHom.id_apply /-- The composition of continuous functions. -/ @[simps] def comp (f : β →𝒄 γ) (g : α →𝒄 β) : α →𝒄 γ := ofMono (OrderHom.comp ↑f ↑g) (continuous_comp _ _ g.cont f.cont) #align omega_complete_partial_order.continuous_hom.comp OmegaCompletePartialOrder.ContinuousHom.comp +#align omega_complete_partial_order.continuous_hom.comp_apply OmegaCompletePartialOrder.ContinuousHom.comp_apply @[ext] protected theorem ext (f g : α →𝒄 β) (h : ∀ x, f x = g x) : f = g := by @@ -806,6 +815,7 @@ def toMono : (α →𝒄 β) →o α →o β where #align omega_complete_partial_order.continuous_hom.to_mono OmegaCompletePartialOrder.ContinuousHom.toMono +#align omega_complete_partial_order.continuous_hom.to_mono_coe OmegaCompletePartialOrder.ContinuousHom.toMono_coe /-- When proving that a chain of applications is below a bound `z`, it suffices to consider the functions and values being selected from the same index in the chains. @@ -848,6 +858,7 @@ protected def ωSup (c : Chain (α →𝒄 β)) : α →𝒄 β := OrderHom.omegaCompletePartialOrder_ωSup_coe, forall_forall_merge, forall_forall_merge', (· ∘ ·), Function.eval]) #align omega_complete_partial_order.continuous_hom.ωSup OmegaCompletePartialOrder.ContinuousHom.ωSup +#align omega_complete_partial_order.continuous_hom.ωSup_apply OmegaCompletePartialOrder.ContinuousHom.ωSup_apply @[simps ωSup] instance : OmegaCompletePartialOrder (α →𝒄 β) := @@ -886,6 +897,7 @@ def apply : (α →𝒄 β) × α →𝒄 β where #align omega_complete_partial_order.continuous_hom.prod.apply OmegaCompletePartialOrder.ContinuousHom.Prod.apply +#align omega_complete_partial_order.continuous_hom.prod.apply_apply OmegaCompletePartialOrder.ContinuousHom.Prod.apply_apply end Prod @@ -909,6 +921,7 @@ def flip {α : Type _} (f : α → β →𝒄 γ) : monotone' x y h a := (f a).monotone h cont := by intro _ _; ext x; change f _ _ = _; rw [(f _).continuous]; rfl #align omega_complete_partial_order.continuous_hom.flip OmegaCompletePartialOrder.ContinuousHom.flip +#align omega_complete_partial_order.continuous_hom.flip_apply OmegaCompletePartialOrder.ContinuousHom.flip_apply /-- `Part.bind` as a continuous function. -/ @[simps] --Porting note: removed `(config := { rhsMd := reducible })` @@ -926,6 +939,7 @@ noncomputable def map {β γ : Type v} (f : β → γ) (g : α →𝒄 Part β) simp only [map_eq_bind_pure_comp, bind, OrderHom.bind_coe, const_apply, OrderHom.const_coe_coe] #align omega_complete_partial_order.continuous_hom.map OmegaCompletePartialOrder.ContinuousHom.map +#align omega_complete_partial_order.continuous_hom.map_apply OmegaCompletePartialOrder.ContinuousHom.map_apply /-- `Part.seq` as a continuous function. -/ @[simps] --Porting note: removed `(config := { rhsMd := reducible })` @@ -936,6 +950,7 @@ noncomputable def seq {β γ : Type v} (f : α →𝒄 Part (β → γ)) (g : α bind, OrderHom.bind_coe, flip_apply] rfl #align omega_complete_partial_order.continuous_hom.seq OmegaCompletePartialOrder.ContinuousHom.seq +#align omega_complete_partial_order.continuous_hom.seq_apply OmegaCompletePartialOrder.ContinuousHom.seq_apply end ContinuousHom diff --git a/Mathlib/Order/RelClasses.lean b/Mathlib/Order/RelClasses.lean index bb56b2bf629ce..c1d16b5cdcfe6 100644 --- a/Mathlib/Order/RelClasses.lean +++ b/Mathlib/Order/RelClasses.lean @@ -30,16 +30,20 @@ open Function theorem of_eq [IsRefl α r] : ∀ {a b}, a = b → r a b | _, _, .refl _ => refl _ +#align of_eq of_eq theorem comm [IsSymm α r] {a b : α} : r a b ↔ r b a := ⟨symm, symm⟩ +#align comm comm theorem antisymm' [IsAntisymm α r] {a b : α} : r a b → r b a → b = a := fun h h' => antisymm h' h +#align antisymm' antisymm' theorem antisymm_iff [IsRefl α r] [IsAntisymm α r] {a b : α} : r a b ∧ r b a ↔ a = b := ⟨fun h => antisymm h.1 h.2, by rintro rfl exact ⟨refl _, refl _⟩⟩ +#align antisymm_iff antisymm_iff /-- A version of `antisymm` with `r` explicit. @@ -47,6 +51,7 @@ This lemma matches the lemmas from lean core in `Init.Algebra.Classes`, but is m @[elab_without_expected_type] theorem antisymm_of (r : α → α → Prop) [IsAntisymm α r] {a b : α} : r a b → r b a → a = b := antisymm +#align antisymm_of antisymm_of /-- A version of `antisymm'` with `r` explicit. @@ -54,57 +59,74 @@ This lemma matches the lemmas from lean core in `Init.Algebra.Classes`, but is m @[elab_without_expected_type] theorem antisymm_of' (r : α → α → Prop) [IsAntisymm α r] {a b : α} : r a b → r b a → b = a := antisymm' +#align antisymm_of' antisymm_of' /-- A version of `comm` with `r` explicit. This lemma matches the lemmas from lean core in `Init.Algebra.Classes`, but is missing there. -/ theorem comm_of (r : α → α → Prop) [IsSymm α r] {a b : α} : r a b ↔ r b a := comm +#align comm_of comm_of theorem IsRefl.swap (r) [IsRefl α r] : IsRefl α (swap r) := ⟨refl_of r⟩ +#align is_refl.swap IsRefl.swap theorem IsIrrefl.swap (r) [IsIrrefl α r] : IsIrrefl α (swap r) := ⟨irrefl_of r⟩ +#align is_irrefl.swap IsIrrefl.swap theorem IsTrans.swap (r) [IsTrans α r] : IsTrans α (swap r) := ⟨fun _ _ _ h₁ h₂ => trans_of r h₂ h₁⟩ +#align is_trans.swap IsTrans.swap theorem IsAntisymm.swap (r) [IsAntisymm α r] : IsAntisymm α (swap r) := ⟨fun _ _ h₁ h₂ => _root_.antisymm h₂ h₁⟩ +#align is_antisymm.swap IsAntisymm.swap theorem IsAsymm.swap (r) [IsAsymm α r] : IsAsymm α (swap r) := ⟨fun _ _ h₁ h₂ => asymm_of r h₂ h₁⟩ +#align is_asymm.swap IsAsymm.swap theorem IsTotal.swap (r) [IsTotal α r] : IsTotal α (swap r) := ⟨fun a b => (total_of r a b).symm⟩ +#align is_total.swap IsTotal.swap theorem IsTrichotomous.swap (r) [IsTrichotomous α r] : IsTrichotomous α (swap r) := ⟨fun a b => by simpa [Function.swap, or_comm, or_left_comm] using trichotomous_of r a b⟩ +#align is_trichotomous.swap IsTrichotomous.swap theorem IsPreorder.swap (r) [IsPreorder α r] : IsPreorder α (swap r) := { @IsRefl.swap α r _, @IsTrans.swap α r _ with } +#align is_preorder.swap IsPreorder.swap theorem IsStrictOrder.swap (r) [IsStrictOrder α r] : IsStrictOrder α (swap r) := { @IsIrrefl.swap α r _, @IsTrans.swap α r _ with } +#align is_strict_order.swap IsStrictOrder.swap theorem IsPartialOrder.swap (r) [IsPartialOrder α r] : IsPartialOrder α (swap r) := { @IsPreorder.swap α r _, @IsAntisymm.swap α r _ with } +#align is_partial_order.swap IsPartialOrder.swap theorem IsTotalPreorder.swap (r) [IsTotalPreorder α r] : IsTotalPreorder α (swap r) := { @IsPreorder.swap α r _, @IsTotal.swap α r _ with } +#align is_total_preorder.swap IsTotalPreorder.swap theorem IsLinearOrder.swap (r) [IsLinearOrder α r] : IsLinearOrder α (swap r) := { @IsPartialOrder.swap α r _, @IsTotal.swap α r _ with } +#align is_linear_order.swap IsLinearOrder.swap protected theorem IsAsymm.isAntisymm (r) [IsAsymm α r] : IsAntisymm α r := ⟨fun _ _ h₁ h₂ => (_root_.asymm h₁ h₂).elim⟩ +#align is_asymm.is_antisymm IsAsymm.isAntisymm protected theorem IsAsymm.isIrrefl [IsAsymm α r] : IsIrrefl α r := ⟨fun _ h => _root_.asymm h h⟩ +#align is_asymm.is_irrefl IsAsymm.isIrrefl protected theorem IsTotal.isTrichotomous (r) [IsTotal α r] : IsTrichotomous α r := ⟨fun a b => or_left_comm.1 (Or.inr <| total_of r a b)⟩ +#align is_total.is_trichotomous IsTotal.isTrichotomous -- see Note [lower instance priority] instance (priority := 100) IsTotal.to_isRefl (r) [IsTotal α r] : IsRefl α r := @@ -112,22 +134,28 @@ instance (priority := 100) IsTotal.to_isRefl (r) [IsTotal α r] : IsRefl α r := theorem ne_of_irrefl {r} [IsIrrefl α r] : ∀ {x y : α}, r x y → x ≠ y | _, _, h, rfl => irrefl _ h +#align ne_of_irrefl ne_of_irrefl theorem ne_of_irrefl' {r} [IsIrrefl α r] : ∀ {x y : α}, r x y → y ≠ x | _, _, h, rfl => irrefl _ h +#align ne_of_irrefl' ne_of_irrefl' theorem not_rel_of_subsingleton (r) [IsIrrefl α r] [Subsingleton α] (x y) : ¬r x y := Subsingleton.elim x y ▸ irrefl x +#align not_rel_of_subsingleton not_rel_of_subsingleton theorem rel_of_subsingleton (r) [IsRefl α r] [Subsingleton α] (x y) : r x y := Subsingleton.elim x y ▸ refl x +#align rel_of_subsingleton rel_of_subsingleton @[simp] theorem empty_relation_apply (a b : α) : EmptyRelation a b ↔ False := Iff.rfl +#align empty_relation_apply empty_relation_apply theorem eq_empty_relation (r) [IsIrrefl α r] [Subsingleton α] : r = EmptyRelation := funext₂ <| by simpa using not_rel_of_subsingleton r +#align eq_empty_relation eq_empty_relation instance : IsIrrefl α EmptyRelation := ⟨fun _ => id⟩ @@ -137,20 +165,24 @@ theorem trans_trichotomous_left [IsTrans α r] [IsTrichotomous α r] {a b c : α intro h₁ h₂ rcases trichotomous_of r a b with (h₃ | rfl | h₃) exacts [_root_.trans h₃ h₂, h₂, absurd h₃ h₁] +#align trans_trichotomous_left trans_trichotomous_left theorem trans_trichotomous_right [IsTrans α r] [IsTrichotomous α r] {a b c : α} : r a b → ¬r c b → r a c := by intro h₁ h₂ rcases trichotomous_of r b c with (h₃ | rfl | h₃) exacts [_root_.trans h₁ h₃, h₁, absurd h₃ h₂] +#align trans_trichotomous_right trans_trichotomous_right theorem transitive_of_trans (r : α → α → Prop) [IsTrans α r] : Transitive r := IsTrans.trans +#align transitive_of_trans transitive_of_trans /-- In a trichotomous irreflexive order, every element is determined by the set of predecessors. -/ theorem extensional_of_trichotomous_of_irrefl (r : α → α → Prop) [IsTrichotomous α r] [IsIrrefl α r] {a b : α} (H : ∀ x, r x a ↔ r x b) : a = b := ((@trichotomous _ r _ a b).resolve_left <| mt (H _).2 <| irrefl a).resolve_right <| mt (H _).1 <| irrefl b +#align extensional_of_trichotomous_of_irrefl extensional_of_trichotomous_of_irrefl /-- Construct a partial order from a `isStrictOrder` relation. @@ -173,6 +205,8 @@ def partialOrderOfSO (r) [IsStrictOrder α r] : PartialOrder α where lt_iff_le_not_le x y := ⟨fun h => ⟨Or.inr h, not_or_of_not (fun e => by rw [e] at h; exact irrefl _ h) (asymm h)⟩, fun ⟨h₁, h₂⟩ => h₁.resolve_left fun e => h₂ <| e ▸ Or.inl rfl⟩ +set_option linter.uppercaseLean3 false in +#align partial_order_of_SO partialOrderOfSO /-- Construct a linear order from an `IsStrictTotalOrder` relation. @@ -192,9 +226,12 @@ def linearOrderOfSTO (r) [IsStrictTotalOrder α r] [∀ x y, Decidable ¬r x y] toMin := minOfLe, toMax := maxOfLe, decidable_le := hD } +set_option linter.uppercaseLean3 false in +#align linear_order_of_STO linearOrderOfSTO theorem IsStrictTotalOrder.swap (r) [IsStrictTotalOrder α r] : IsStrictTotalOrder α (swap r) := { IsTrichotomous.swap r, IsStrictOrder.swap r with } +#align is_strict_total_order.swap IsStrictTotalOrder.swap /-! ### Order connection -/ @@ -205,10 +242,12 @@ theorem IsStrictTotalOrder.swap (r) [IsStrictTotalOrder α r] : IsStrictTotalOrd class IsOrderConnected (α : Type u) (lt : α → α → Prop) : Prop where /-- A connected order is one satisfying the condition `a < c → a < b ∨ b < c`. -/ conn : ∀ a b c, lt a c → lt a b ∨ lt b c +#align is_order_connected IsOrderConnected theorem IsOrderConnected.neg_trans {r : α → α → Prop} [IsOrderConnected α r] {a b c} (h₁ : ¬r a b) (h₂ : ¬r b c) : ¬r a c := mt (IsOrderConnected.conn a b c) <| by simp [h₁, h₂] +#align is_order_connected.neg_trans IsOrderConnected.neg_trans theorem isStrictWeakOrder_of_isOrderConnected [IsAsymm α r] [IsOrderConnected α r] : IsStrictWeakOrder α r := @@ -238,6 +277,8 @@ instance (priority := 100) isStrictTotalOrder_of_isStrictTotalOrder [IsStrictTot @[mk_iff] class IsWellFounded (α : Type u) (r : α → α → Prop) : Prop where /-- The relation is `WellFounded`, as a proposition. -/ wf : WellFounded r +#align is_well_founded IsWellFounded +#align is_well_founded_iff IsWellFounded_iff #align has_well_founded WellFoundedRelation set_option linter.uppercaseLean3 false in @@ -257,10 +298,12 @@ variable (r) [IsWellFounded α r] /-- Induction on a well-founded relation. -/ theorem induction {C : α → Prop} : ∀ a, (∀ x, (∀ y, r y x → C y) → C x) → C a := wf.induction +#align is_well_founded.induction IsWellFounded.induction /-- All values are accessible under the well-founded relation. -/ theorem apply : ∀ a, Acc r a := wf.apply +#align is_well_founded.apply IsWellFounded.apply -- Porting note: WellFounded.fix is noncomputable, at 2022-10-29 lean /-- Creates data, given a way to generate a value from all that compare as less under a well-founded @@ -268,11 +311,13 @@ relation. See also `IsWellFounded.fix_eq`. -/ noncomputable def fix {C : α → Sort _} : (∀ x : α, (∀ y : α, r y x → C y) → C x) → ∀ x : α, C x := wf.fix +#align is_well_founded.fix IsWellFounded.fix /-- The value from `IsWellFounded.fix` is built from the previous ones as specified. -/ theorem fix_eq {C : α → Sort _} (F : ∀ x : α, (∀ y : α, r y x → C y) → C x) : ∀ x, fix r F x = F x fun y _ => fix r F y := wf.fix_eq F +#align is_well_founded.fix_eq IsWellFounded.fix_eq /-- Derive a `WellFoundedRelation` instance from an `isWellFounded` instance. -/ def toWellFoundedRelation : WellFoundedRelation α := @@ -283,6 +328,7 @@ end IsWellFounded theorem WellFounded.asymmetric {α : Sort _} {r : α → α → Prop} (h : WellFounded r) (a b) : r a b → ¬r b a := @WellFoundedRelation.asymmetric _ ⟨_, h⟩ _ _ +#align well_founded.asymmetric WellFounded.asymmetric -- see Note [lower instance priority] instance (priority := 100) (r : α → α → Prop) [IsWellFounded α r] : IsAsymm α r := @@ -296,11 +342,13 @@ instance (priority := 100) (r : α → α → Prop) [IsWellFounded α r] : IsIrr @[reducible] def WellFoundedLT (α : Type _) [LT α] : Prop := IsWellFounded α (· < ·) +#align well_founded_lt WellFoundedLT /-- A class for a well founded relation `>`. -/ @[reducible] def WellFoundedGT (α : Type _) [LT α] : Prop := IsWellFounded α (· > ·) +#align well_founded_gt WellFoundedGT -- See note [lower instance priority] instance (priority := 100) (α : Type _) [LT α] [h : WellFoundedLT α] : WellFoundedGT αᵒᵈ := @@ -312,13 +360,16 @@ instance (priority := 100) (α : Type _) [LT α] [h : WellFoundedGT α] : WellFo theorem wellFoundedGT_dual_iff (α : Type _) [LT α] : WellFoundedGT αᵒᵈ ↔ WellFoundedLT α := ⟨fun h => ⟨h.wf⟩, fun h => ⟨h.wf⟩⟩ +#align well_founded_gt_dual_iff wellFoundedGT_dual_iff theorem wellFoundedLT_dual_iff (α : Type _) [LT α] : WellFoundedLT αᵒᵈ ↔ WellFoundedGT α := ⟨fun h => ⟨h.wf⟩, fun h => ⟨h.wf⟩⟩ +#align well_founded_lt_dual_iff wellFoundedLT_dual_iff /-- A well order is a well-founded linear order. -/ class IsWellOrder (α : Type u) (r : α → α → Prop) extends IsTrichotomous α r, IsTrans α r, IsWellFounded α r : Prop +#align is_well_order IsWellOrder -- see Note [lower instance priority] instance (priority := 100) {α} (r : α → α → Prop) [IsWellOrder α r] : @@ -347,10 +398,12 @@ variable [LT α] [WellFoundedLT α] /-- Inducts on a well-founded `<` relation. -/ theorem induction {C : α → Prop} : ∀ a, (∀ x, (∀ y, y < x → C y) → C x) → C a := IsWellFounded.induction _ +#align well_founded_lt.induction WellFoundedLT.induction /-- All values are accessible under the well-founded `<`. -/ theorem apply : ∀ a : α, Acc (· < ·) a := IsWellFounded.apply _ +#align well_founded_lt.apply WellFoundedLT.apply -- Porting note: WellFounded.fix is noncomputable, at 2022-10-29 lean /-- Creates data, given a way to generate a value from all that compare as lesser. See also @@ -358,11 +411,13 @@ theorem apply : ∀ a : α, Acc (· < ·) a := noncomputable def fix {C : α → Sort _} : (∀ x : α, (∀ y : α, y < x → C y) → C x) → ∀ x : α, C x := IsWellFounded.fix (· < ·) +#align well_founded_lt.fix WellFoundedLT.fix /-- The value from `WellFoundedLT.fix` is built from the previous ones as specified. -/ theorem fix_eq {C : α → Sort _} (F : ∀ x : α, (∀ y : α, y < x → C y) → C x) : ∀ x, fix F x = F x fun y _ => fix F y := IsWellFounded.fix_eq _ F +#align well_founded_lt.fix_eq WellFoundedLT.fix_eq /-- Derive a `WellFoundedRelation` instance from a `WellFoundedLT` instance. -/ def toWellFoundedRelation : WellFoundedRelation α := @@ -377,10 +432,12 @@ variable [LT α] [WellFoundedGT α] /-- Inducts on a well-founded `>` relation. -/ theorem induction {C : α → Prop} : ∀ a, (∀ x, (∀ y, x < y → C y) → C x) → C a := IsWellFounded.induction _ +#align well_founded_gt.induction WellFoundedGT.induction /-- All values are accessible under the well-founded `>`. -/ theorem apply : ∀ a : α, Acc (· > ·) a := IsWellFounded.apply _ +#align well_founded_gt.apply WellFoundedGT.apply -- Porting note: WellFounded.fix is noncomputable, at 2022-10-29 lean /-- Creates data, given a way to generate a value from all that compare as greater. See also @@ -388,11 +445,13 @@ theorem apply : ∀ a : α, Acc (· > ·) a := noncomputable def fix {C : α → Sort _} : (∀ x : α, (∀ y : α, x < y → C y) → C x) → ∀ x : α, C x := IsWellFounded.fix (· > ·) +#align well_founded_gt.fix WellFoundedGT.fix /-- The value from `WellFoundedGT.fix` is built from the successive ones as specified. -/ theorem fix_eq {C : α → Sort _} (F : ∀ x : α, (∀ y : α, x < y → C y) → C x) : ∀ x, fix F x = F x fun y _ => fix F y := IsWellFounded.fix_eq _ F +#align well_founded_gt.fix_eq WellFoundedGT.fix_eq /-- Derive a `WellFoundedRelation` instance from a `WellFoundedGT` instance. -/ def toWellFoundedRelation : WellFoundedRelation α := @@ -404,11 +463,13 @@ end WellFoundedGT noncomputable def IsWellOrder.linearOrder (r : α → α → Prop) [IsWellOrder α r] : LinearOrder α := letI := fun x y => Classical.dec ¬r x y linearOrderOfSTO r +#align is_well_order.linear_order IsWellOrder.linearOrder /-- Derive a `WellFoundedRelation` instance from a `IsWellOrder` instance. -/ def IsWellOrder.toHasWellFounded [LT α] [hwo : IsWellOrder α (· < ·)] : WellFoundedRelation α where rel := (· < ·) wf := hwo.wf +#align is_well_order.to_has_well_founded IsWellOrder.toHasWellFounded -- This isn't made into an instance as it loops with `IsIrrefl α r`. theorem Subsingleton.isWellOrder [Subsingleton α] (r : α → α → Prop) [hr : IsIrrefl α r] : @@ -417,6 +478,7 @@ theorem Subsingleton.isWellOrder [Subsingleton α] (r : α → α → Prop) [hr trichotomous := fun a b => Or.inr <| Or.inl <| Subsingleton.elim a b, trans := fun a b _ h => (not_rel_of_subsingleton r a b h).elim, wf := ⟨fun a => ⟨_, fun y h => (not_rel_of_subsingleton r y a h).elim⟩⟩ } +#align subsingleton.is_well_order Subsingleton.isWellOrder instance [Subsingleton α] : IsWellOrder α EmptyRelation := Subsingleton.isWellOrder _ @@ -455,27 +517,33 @@ instance (f : α → ℕ) : IsWellFounded _ (Measure f) := theorem Subrelation.isWellFounded (r : α → α → Prop) [IsWellFounded α r] {s : α → α → Prop} (h : Subrelation s r) : IsWellFounded α s := ⟨h.wf IsWellFounded.wf⟩ +#align subrelation.is_well_founded Subrelation.isWellFounded namespace Set /-- An unbounded or cofinal set. -/ def Unbounded (r : α → α → Prop) (s : Set α) : Prop := ∀ a, ∃ b ∈ s, ¬r b a +#align set.unbounded Set.Unbounded /-- A bounded or final set. Not to be confused with `Metric.bounded`. -/ def Bounded (r : α → α → Prop) (s : Set α) : Prop := ∃ a, ∀ b ∈ s, r b a +#align set.bounded Set.Bounded @[simp] theorem not_bounded_iff {r : α → α → Prop} (s : Set α) : ¬Bounded r s ↔ Unbounded r s := by simp only [Bounded, Unbounded, not_forall, not_exists, exists_prop, not_and, not_not] +#align set.not_bounded_iff Set.not_bounded_iff @[simp] theorem not_unbounded_iff {r : α → α → Prop} (s : Set α) : ¬Unbounded r s ↔ Bounded r s := by rw [not_iff_comm, not_bounded_iff] +#align set.not_unbounded_iff Set.not_unbounded_iff theorem unbounded_of_isEmpty [IsEmpty α] {r : α → α → Prop} (s : Set α) : Unbounded r s := isEmptyElim +#align set.unbounded_of_is_empty Set.unbounded_of_isEmpty end Set @@ -506,15 +574,18 @@ notation on `(⊆)` and `(⊂)`. -/ class IsNonstrictStrictOrder (α : Type _) (r s : α → α → Prop) where /-- The relation `r` is the nonstrict relation corresponding to the strict relation `s`. -/ right_iff_left_not_left (a b : α) : s a b ↔ r a b ∧ ¬r b a +#align is_nonstrict_strict_order IsNonstrictStrictOrder theorem right_iff_left_not_left {r s : α → α → Prop} [IsNonstrictStrictOrder α r s] {a b : α} : s a b ↔ r a b ∧ ¬r b a := IsNonstrictStrictOrder.right_iff_left_not_left _ _ +#align right_iff_left_not_left right_iff_left_not_left /-- A version of `right_iff_left_not_left` with explicit `r` and `s`. -/ theorem right_iff_left_not_left_of (r s : α → α → Prop) [IsNonstrictStrictOrder α r s] {a b : α} : s a b ↔ r a b ∧ ¬r b a := right_iff_left_not_left +#align right_iff_left_not_left_of right_iff_left_not_left_of -- The free parameter `r` is strictly speaking not uniquely determined by `s`, but in practice it -- always has a unique instance, so this is not dangerous. @@ -563,12 +634,15 @@ alias subset_trans ← HasSubset.Subset.trans alias subset_antisymm ← HasSubset.Subset.antisymm #align has_subset.subset.antisymm HasSubset.Subset.antisymm alias superset_antisymm ← HasSubset.Subset.antisymm' +#align has_subset.subset.antisymm' HasSubset.Subset.antisymm' theorem subset_antisymm_iff [IsRefl α (· ⊆ ·)] [IsAntisymm α (· ⊆ ·)] : a = b ↔ a ⊆ b ∧ b ⊆ a := ⟨fun h => ⟨h.subset', h.superset⟩, fun h => h.1.antisymm h.2⟩ +#align subset_antisymm_iff subset_antisymm_iff theorem superset_antisymm_iff [IsRefl α (· ⊆ ·)] [IsAntisymm α (· ⊆ ·)] : a = b ↔ b ⊆ a ∧ a ⊆ b := ⟨fun h => ⟨h.superset, h.subset'⟩, fun h => h.1.antisymm' h.2⟩ +#align superset_antisymm_iff superset_antisymm_iff end Subset @@ -615,63 +689,85 @@ variable [HasSubset α] [HasSSubset α] [IsNonstrictStrictOrder α (· ⊆ ·) ( theorem ssubset_iff_subset_not_subset : a ⊂ b ↔ a ⊆ b ∧ ¬b ⊆ a := right_iff_left_not_left +#align ssubset_iff_subset_not_subset ssubset_iff_subset_not_subset theorem subset_of_ssubset (h : a ⊂ b) : a ⊆ b := (ssubset_iff_subset_not_subset.1 h).1 +#align subset_of_ssubset subset_of_ssubset theorem not_subset_of_ssubset (h : a ⊂ b) : ¬b ⊆ a := (ssubset_iff_subset_not_subset.1 h).2 +#align not_subset_of_ssubset not_subset_of_ssubset theorem not_ssubset_of_subset (h : a ⊆ b) : ¬b ⊂ a := fun h' => not_subset_of_ssubset h' h +#align not_ssubset_of_subset not_ssubset_of_subset theorem ssubset_of_subset_not_subset (h₁ : a ⊆ b) (h₂ : ¬b ⊆ a) : a ⊂ b := ssubset_iff_subset_not_subset.2 ⟨h₁, h₂⟩ +#align ssubset_of_subset_not_subset ssubset_of_subset_not_subset alias subset_of_ssubset ← HasSSubset.SSubset.subset #align has_ssubset.ssubset.subset HasSSubset.SSubset.subset alias not_subset_of_ssubset ← HasSSubset.SSubset.not_subset +#align has_ssubset.ssubset.not_subset HasSSubset.SSubset.not_subset alias not_ssubset_of_subset ← HasSubset.Subset.not_ssubset +#align has_subset.subset.not_ssubset HasSubset.Subset.not_ssubset alias ssubset_of_subset_not_subset ← HasSubset.Subset.ssubset_of_not_subset +#align has_subset.subset.ssubset_of_not_subset HasSubset.Subset.ssubset_of_not_subset theorem ssubset_of_subset_of_ssubset [IsTrans α (· ⊆ ·)] (h₁ : a ⊆ b) (h₂ : b ⊂ c) : a ⊂ c := (h₁.trans h₂.subset).ssubset_of_not_subset fun h => h₂.not_subset <| h.trans h₁ +#align ssubset_of_subset_of_ssubset ssubset_of_subset_of_ssubset theorem ssubset_of_ssubset_of_subset [IsTrans α (· ⊆ ·)] (h₁ : a ⊂ b) (h₂ : b ⊆ c) : a ⊂ c := (h₁.subset.trans h₂).ssubset_of_not_subset fun h => h₁.not_subset <| h₂.trans h +#align ssubset_of_ssubset_of_subset ssubset_of_ssubset_of_subset theorem ssubset_of_subset_of_ne [IsAntisymm α (· ⊆ ·)] (h₁ : a ⊆ b) (h₂ : a ≠ b) : a ⊂ b := h₁.ssubset_of_not_subset <| mt h₁.antisymm h₂ +#align ssubset_of_subset_of_ne ssubset_of_subset_of_ne theorem ssubset_of_ne_of_subset [IsAntisymm α (· ⊆ ·)] (h₁ : a ≠ b) (h₂ : a ⊆ b) : a ⊂ b := ssubset_of_subset_of_ne h₂ h₁ +#align ssubset_of_ne_of_subset ssubset_of_ne_of_subset theorem eq_or_ssubset_of_subset [IsAntisymm α (· ⊆ ·)] (h : a ⊆ b) : a = b ∨ a ⊂ b := (em (b ⊆ a)).imp h.antisymm h.ssubset_of_not_subset +#align eq_or_ssubset_of_subset eq_or_ssubset_of_subset theorem ssubset_or_eq_of_subset [IsAntisymm α (· ⊆ ·)] (h : a ⊆ b) : a ⊂ b ∨ a = b := (eq_or_ssubset_of_subset h).symm +#align ssubset_or_eq_of_subset ssubset_or_eq_of_subset alias ssubset_of_subset_of_ssubset ← HasSubset.Subset.trans_ssubset +#align has_subset.subset.trans_ssubset HasSubset.Subset.trans_ssubset alias ssubset_of_ssubset_of_subset ← HasSSubset.SSubset.trans_subset +#align has_ssubset.ssubset.trans_subset HasSSubset.SSubset.trans_subset alias ssubset_of_subset_of_ne ← HasSubset.Subset.ssubset_of_ne +#align has_subset.subset.ssubset_of_ne HasSubset.Subset.ssubset_of_ne alias ssubset_of_ne_of_subset ← Ne.ssubset_of_subset +#align ne.ssubset_of_subset Ne.ssubset_of_subset alias eq_or_ssubset_of_subset ← HasSubset.Subset.eq_or_ssubset +#align has_subset.subset.eq_or_ssubset HasSubset.Subset.eq_or_ssubset alias ssubset_or_eq_of_subset ← HasSubset.Subset.ssubset_or_eq +#align has_subset.subset.ssubset_or_eq HasSubset.Subset.ssubset_or_eq theorem ssubset_iff_subset_ne [IsAntisymm α (· ⊆ ·)] : a ⊂ b ↔ a ⊆ b ∧ a ≠ b := ⟨fun h => ⟨h.subset, h.ne⟩, fun h => h.1.ssubset_of_ne h.2⟩ +#align ssubset_iff_subset_ne ssubset_iff_subset_ne theorem subset_iff_ssubset_or_eq [IsRefl α (· ⊆ ·)] [IsAntisymm α (· ⊆ ·)] : a ⊆ b ↔ a ⊂ b ∨ a = b := ⟨fun h => h.ssubset_or_eq, fun h => h.elim subset_of_ssubset subset_of_eq⟩ +#align subset_iff_ssubset_or_eq subset_iff_ssubset_or_eq end SubsetSsubset @@ -772,15 +868,19 @@ instance [LinearOrder α] : IsStrictWeakOrder α (· < ·) := by infer_instance theorem transitive_le [Preorder α] : Transitive (@LE.le α _) := transitive_of_trans _ +#align transitive_le transitive_le theorem transitive_lt [Preorder α] : Transitive (@LT.lt α _) := transitive_of_trans _ +#align transitive_lt transitive_lt theorem transitive_ge [Preorder α] : Transitive (@GE.ge α _) := transitive_of_trans _ +#align transitive_ge transitive_ge theorem transitive_gt [Preorder α] : Transitive (@GT.gt α _) := transitive_of_trans _ +#align transitive_gt transitive_gt instance OrderDual.isTotal_le [LE α] [h : IsTotal α (· ≤ ·)] : IsTotal αᵒᵈ (· ≤ ·) := @IsTotal.swap α _ h diff --git a/Mathlib/Order/RelIso/Basic.lean b/Mathlib/Order/RelIso/Basic.lean index d476292cc2df8..a03dca860356f 100644 --- a/Mathlib/Order/RelIso/Basic.lean +++ b/Mathlib/Order/RelIso/Basic.lean @@ -144,12 +144,14 @@ theorem ext_iff {f g : r →r s} : f = g ↔ ∀ x, f x = g x := protected def id (r : α → α → Prop) : r →r r := ⟨fun x => x, fun x => x⟩ #align rel_hom.id RelHom.id +#align rel_hom.id_apply RelHom.id_apply /-- Composition of two relation homomorphisms is a relation homomorphism. -/ @[simps] protected def comp (g : s →r t) (f : r →r s) : r →r t := ⟨fun x => g (f x), fun h => g.2 (f.2 h)⟩ #align rel_hom.comp RelHom.comp +#align rel_hom.comp_apply RelHom.comp_apply /-- A relation homomorphism is also a relation homomorphism between dual relations. -/ protected def swap (f : r →r s) : swap r →r swap s := @@ -286,6 +288,7 @@ theorem ext_iff {f g : r ↪r s} : f = g ↔ ∀ x, f x = g x := protected def refl (r : α → α → Prop) : r ↪r r := ⟨Embedding.refl _, Iff.rfl⟩ #align rel_embedding.refl RelEmbedding.refl +#align rel_embedding.refl_apply RelEmbedding.refl_apply /-- Composition of two relation embeddings is a relation embedding. -/ protected def trans (f : r ↪r s) (g : s ↪r t) : r ↪r t := @@ -397,6 +400,7 @@ noncomputable def _root_.Quotient.outRelEmbedding [s : Setoid α] {r : α → α refine' @fun x y => Quotient.inductionOn₂ x y fun a b => _ apply iff_iff_eq.2 (H _ _ _ _ _ _) <;> apply Quotient.mk_out⟩ #align quotient.out_rel_embedding Quotient.outRelEmbedding +#align quotient.out_rel_embedding_apply Quotient.outRelEmbedding_apply /-- A relation is well founded iff its lift to a quotient is. -/ @[simp] @@ -468,6 +472,7 @@ def sumLiftRelInl (r : α → α → Prop) (s : β → β → Prop) : r ↪r Sum inj' := Sum.inl_injective map_rel_iff' := Sum.liftRel_inl_inl #align rel_embedding.sum_lift_rel_inl RelEmbedding.sumLiftRelInl +#align rel_embedding.sum_lift_rel_inl_apply RelEmbedding.sumLiftRelInl_apply /-- `Sum.inr` as a relation embedding into `Sum.LiftRel r s`. -/ @[simps] @@ -476,6 +481,7 @@ def sumLiftRelInr (r : α → α → Prop) (s : β → β → Prop) : s ↪r Sum inj' := Sum.inr_injective map_rel_iff' := Sum.liftRel_inr_inr #align rel_embedding.sum_lift_rel_inr RelEmbedding.sumLiftRelInr +#align rel_embedding.sum_lift_rel_inr_apply RelEmbedding.sumLiftRelInr_apply /-- `Sum.map` as a relation embedding between `Sum.LiftRel` relations. -/ @[simps] @@ -484,6 +490,7 @@ def sumLiftRelMap (f : r ↪r s) (g : t ↪r u) : Sum.LiftRel r t ↪r Sum.LiftR inj' := f.injective.sum_map g.injective map_rel_iff' := by rintro (a | b) (c | d) <;> simp [f.map_rel_iff, g.map_rel_iff] #align rel_embedding.sum_lift_rel_map RelEmbedding.sumLiftRelMap +#align rel_embedding.sum_lift_rel_map_apply RelEmbedding.sumLiftRelMap_apply /-- `Sum.inl` as a relation embedding into `Sum.Lex r s`. -/ @[simps] @@ -492,6 +499,7 @@ def sumLexInl (r : α → α → Prop) (s : β → β → Prop) : r ↪r Sum.Lex inj' := Sum.inl_injective map_rel_iff' := Sum.lex_inl_inl #align rel_embedding.sum_lex_inl RelEmbedding.sumLexInl +#align rel_embedding.sum_lex_inl_apply RelEmbedding.sumLexInl_apply /-- `Sum.inr` as a relation embedding into `Sum.Lex r s`. -/ @[simps] @@ -500,6 +508,7 @@ def sumLexInr (r : α → α → Prop) (s : β → β → Prop) : s ↪r Sum.Lex inj' := Sum.inr_injective map_rel_iff' := Sum.lex_inr_inr #align rel_embedding.sum_lex_inr RelEmbedding.sumLexInr +#align rel_embedding.sum_lex_inr_apply RelEmbedding.sumLexInr_apply /-- `Sum.map` as a relation embedding between `Sum.Lex` relations. -/ @[simps] @@ -508,6 +517,7 @@ def sumLexMap (f : r ↪r s) (g : t ↪r u) : Sum.Lex r t ↪r Sum.Lex s u where inj' := f.injective.sum_map g.injective map_rel_iff' := by rintro (a | b) (c | d) <;> simp [f.map_rel_iff, g.map_rel_iff] #align rel_embedding.sum_lex_map RelEmbedding.sumLexMap +#align rel_embedding.sum_lex_map_apply RelEmbedding.sumLexMap_apply /-- `λ b, Prod.mk a b` as a relation embedding. -/ @[simps] @@ -516,6 +526,7 @@ def prodLexMkLeft (s : β → β → Prop) {a : α} (h : ¬r a a) : s ↪r Prod. inj' := Prod.mk.inj_left a map_rel_iff' := by simp [Prod.lex_def, h] #align rel_embedding.prod_lex_mk_left RelEmbedding.prodLexMkLeft +#align rel_embedding.prod_lex_mk_left_apply RelEmbedding.prodLexMkLeft_apply /-- `λ a, Prod.mk a b` as a relation embedding. -/ @[simps] @@ -524,6 +535,7 @@ def prodLexMkRight (r : α → α → Prop) {b : β} (h : ¬s b b) : r ↪r Prod inj' := Prod.mk.inj_right b map_rel_iff' := by simp [Prod.lex_def, h] #align rel_embedding.prod_lex_mk_right RelEmbedding.prodLexMkRight +#align rel_embedding.prod_lex_mk_right_apply RelEmbedding.prodLexMkRight_apply /-- `Prod.map` as a relation embedding. -/ @[simps] @@ -532,6 +544,7 @@ def prodLexMap (f : r ↪r s) (g : t ↪r u) : Prod.Lex r t ↪r Prod.Lex s u wh inj' := f.injective.Prod_map g.injective map_rel_iff' := by simp [Prod.lex_def, f.map_rel_iff, g.map_rel_iff] #align rel_embedding.prod_lex_map RelEmbedding.prodLexMap +#align rel_embedding.prod_lex_map_apply RelEmbedding.prodLexMap_apply end RelEmbedding @@ -624,12 +637,14 @@ initialize_simps_projections RelIso (toEquiv_toFun → apply, toEquiv_invFun → protected def refl (r : α → α → Prop) : r ≃r r := ⟨Equiv.refl _, Iff.rfl⟩ #align rel_iso.refl RelIso.refl +#align rel_iso.refl_apply RelIso.refl_apply /-- Composition of two relation isomorphisms is a relation isomorphism. -/ @[simps apply] protected def trans (f₁ : r ≃r s) (f₂ : s ≃r t) : r ≃r t := ⟨f₁.toEquiv.trans f₂.toEquiv, f₂.map_rel_iff.trans f₁.map_rel_iff⟩ #align rel_iso.trans RelIso.trans +#align rel_iso.trans_apply RelIso.trans_apply instance (r : α → α → Prop) : Inhabited (r ≃r r) := ⟨RelIso.refl _⟩ @@ -648,6 +663,8 @@ protected def cast {α β : Type u} {r : α → α → Prop} {s : β → β → rw [eq_of_heq h₂] rfl⟩ #align rel_iso.cast RelIso.cast +#align rel_iso.cast_apply RelIso.cast_apply +#align rel_iso.cast_to_equiv RelIso.cast_toEquiv @[simp] protected theorem cast_symm {α β : Type u} {r : α → α → Prop} {s : β → β → Prop} (h₁ : α = β) @@ -735,6 +752,7 @@ instance IsWellOrder.ulift {α : Type u} (r : α → α → Prop) [IsWellOrder noncomputable def ofSurjective (f : r ↪r s) (H : Surjective f) : r ≃r s := ⟨Equiv.ofBijective f ⟨f.injective, H⟩, f.map_rel_iff⟩ #align rel_iso.of_surjective RelIso.ofSurjective +#align rel_iso.of_surjective_apply RelIso.ofSurjective_apply /-- Given relation isomorphisms `r₁ ≃r s₁` and `r₂ ≃r s₂`, construct a relation isomorphism for the lexicographic orders on the sum. diff --git a/Mathlib/Order/SuccPred/Basic.lean b/Mathlib/Order/SuccPred/Basic.lean index 9c35dd0cf617d..4a6b765cb45b8 100644 --- a/Mathlib/Order/SuccPred/Basic.lean +++ b/Mathlib/Order/SuccPred/Basic.lean @@ -75,6 +75,8 @@ class SuccOrder (α : Type _) [Preorder α] where /--Proof that `succ` satifies ordering invariants betweeen `LE` and `LT`-/ le_of_lt_succ {a b} : a < succ b → a ≤ b #align succ_order SuccOrder +#align succ_order.ext_iff SuccOrder.ext_iff +#align succ_order.ext SuccOrder.ext /-- Order equipped with a sensible predecessor function. -/ @[ext] @@ -90,6 +92,8 @@ class PredOrder (α : Type _) [Preorder α] where /--Proof that `pred` satifies ordering invariants betweeen `LE` and `LT`-/ le_of_pred_lt {a b} : pred a < b → a ≤ b #align pred_order PredOrder +#align pred_order.ext PredOrder.ext +#align pred_order.ext_iff PredOrder.ext_iff instance [Preorder α] [SuccOrder α] : PredOrder αᵒᵈ where @@ -154,6 +158,7 @@ def SuccOrder.ofCore (succ : α → α) (hn : ∀ {a}, ¬IsMax a → ∀ b, a < by_cases (fun h => hm b h ▸ hab.le) fun h => by simpa [hab] using (hn h a).not max_of_succ_le := fun {a} => not_imp_not.mp fun h => by simpa using (hn h a).not } #align succ_order.of_core SuccOrder.ofCore +#align succ_order.of_core_succ SuccOrder.ofCore_succ /-- A constructor for `PredOrder α` for `α` a linear order. -/ @[simps] @@ -169,6 +174,7 @@ def PredOrder.ofCore {α} [LinearOrder α] (pred : α → α) by_cases (fun h => hm a h ▸ hab.le) fun h => by simpa [hab] using (hn h b).not min_of_le_pred := fun {a} => not_imp_not.mp fun h => by simpa using (hn h a).not } #align pred_order.of_core PredOrder.ofCore +#align pred_order.of_core_pred PredOrder.ofCore_pred /-- A constructor for `SuccOrder α` usable when `α` is a linear order with no maximal element. -/ def SuccOrder.ofSuccLeIff (succ : α → α) (hsucc_le_iff : ∀ {a b}, succ a ≤ b ↔ a < b) : diff --git a/Mathlib/Order/SuccPred/Limit.lean b/Mathlib/Order/SuccPred/Limit.lean index 6ac97e1b42e72..26ced77f209ae 100644 --- a/Mathlib/Order/SuccPred/Limit.lean +++ b/Mathlib/Order/SuccPred/Limit.lean @@ -266,8 +266,10 @@ theorem isPredLimit_toDual_iff : IsPredLimit (toDual a) ↔ IsSuccLimit a := by #align order.is_pred_limit_to_dual_iff Order.isPredLimit_toDual_iff alias isSuccLimit_toDual_iff ↔ _ isPredLimit.dual +#align order.is_pred_limit.dual Order.isPredLimit.dual alias isPredLimit_toDual_iff ↔ _ isSuccLimit.dual +#align order.is_succ_limit.dual Order.isSuccLimit.dual end LT diff --git a/Mathlib/Order/Synonym.lean b/Mathlib/Order/Synonym.lean index 7186b2a89857a..f2582b6de1746 100644 --- a/Mathlib/Order/Synonym.lean +++ b/Mathlib/Order/Synonym.lean @@ -129,14 +129,17 @@ theorem toDual_lt [LT α] {a : α} {b : αᵒᵈ} : toDual a < b ↔ ofDual b < @[elab_as_elim] protected def rec {C : αᵒᵈ → Sort _} (h₂ : ∀ a : α, C (toDual a)) : ∀ a : αᵒᵈ, C a := h₂ +#align order_dual.rec OrderDual.rec @[simp] protected theorem «forall» {p : αᵒᵈ → Prop} : (∀ a, p a) ↔ ∀ a, p (toDual a) := Iff.rfl +#align order_dual.forall OrderDual.forall @[simp] protected theorem «exists» {p : αᵒᵈ → Prop} : (∃ a, p a) ↔ ∃ a, p (toDual a) := Iff.rfl +#align order_dual.exists OrderDual.exists alias toDual_le_toDual ↔ _ _root_.LE.le.dual @@ -156,16 +159,19 @@ end OrderDual /-- A type synonym to equip a type with its lexicographic order. -/ def Lex (α : Type _) := α +#align lex Lex /-- `toLex` is the identity function to the `Lex` of a type. -/ @[match_pattern] def toLex : α ≃ Lex α := Equiv.refl _ +#align to_lex toLex /-- `ofLex` is the identity function from the `lex` of a type. -/ @[match_pattern] def ofLex : Lex α ≃ α := Equiv.refl _ +#align of_lex ofLex @[simp] theorem toLex_symm_eq : (@toLex α).symm = ofLex := @@ -201,3 +207,4 @@ theorem ofLex_inj {a b : Lex α} : ofLex a = ofLex b ↔ a = b := /-- A recursor for `Lex`. Use as `induction x using Lex.rec`. -/ protected def Lex.rec {β : Lex α → Sort _} (h : ∀ a, β (toLex a)) : ∀ a, β a := fun a => h (ofLex a) +#align lex.rec Lex.rec diff --git a/Mathlib/RingTheory/RingInvo.lean b/Mathlib/RingTheory/RingInvo.lean index a857e4b3f48dd..487fd302f65cd 100644 --- a/Mathlib/RingTheory/RingInvo.lean +++ b/Mathlib/RingTheory/RingInvo.lean @@ -41,6 +41,7 @@ structure RingInvo [Semiring R] extends R ≃+* Rᵐᵒᵖ where /-- The equivalence of rings underlying a ring involution. -/ add_decl_doc RingInvo.toRingEquiv +#align ring_invo.to_ring_equiv RingInvo.toRingEquiv /-- `RingInvoClass F R` states that `F` is a type of ring involutions. You should extend this class when you extend `RingInvo`. -/