Skip to content

Commit 1d5ae61

Browse files
committed
chore: Remove Init.Propext (#10709)
These lemmas can easily go to `Logic.Basic`
1 parent c06d0ad commit 1d5ae61

File tree

7 files changed

+23
-51
lines changed

7 files changed

+23
-51
lines changed

Mathlib.lean

Lines changed: 0 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -2485,7 +2485,6 @@ import Mathlib.Init.Logic
24852485
import Mathlib.Init.Meta.WellFoundedTactics
24862486
import Mathlib.Init.Order.Defs
24872487
import Mathlib.Init.Order.LinearOrder
2488-
import Mathlib.Init.Propext
24892488
import Mathlib.Init.Quot
24902489
import Mathlib.Init.Set
24912490
import Mathlib.Init.ZeroOne

Mathlib/CategoryTheory/IsConnected.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -185,7 +185,7 @@ The converse is given in `IsConnected.of_induct`.
185185
-/
186186
theorem induct_on_objects [IsPreconnected J] (p : Set J) {j₀ : J} (h0 : j₀ ∈ p)
187187
(h1 : ∀ {j₁ j₂ : J} (_ : j₁ ⟶ j₂), j₁ ∈ p ↔ j₂ ∈ p) (j : J) : j ∈ p := by
188-
let aux (j₁ j₂ : J) (f : j₁ ⟶ j₂) := congrArg ULift.up <| (h1 f).to_eq
188+
let aux (j₁ j₂ : J) (f : j₁ ⟶ j₂) := congrArg ULift.up <| (h1 f).eq
189189
injection constant_of_preserves_morphisms (fun k => ULift.up.{u₁} (k ∈ p)) aux j j₀ with i
190190
rwa [i]
191191
#align category_theory.induct_on_objects CategoryTheory.induct_on_objects

Mathlib/GroupTheory/GroupAction/Quotient.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -427,7 +427,7 @@ theorem card_comm_eq_card_conjClasses_mul_card (G : Type*) [Group G] :
427427
rw [card_congr (Equiv.subtypeProdEquivSigmaSubtype Commute), card_sigma,
428428
sum_equiv ConjAct.toConjAct.toEquiv (fun a ↦ card { b // Commute a b })
429429
(fun g ↦ card (MulAction.fixedBy G g))
430-
fun g ↦ card_congr' <| congr_arg _ <| funext fun h ↦ mul_inv_eq_iff_eq_mul.symm.to_eq,
430+
fun g ↦ card_congr' <| congr_arg _ <| funext fun h ↦ mul_inv_eq_iff_eq_mul.symm.eq,
431431
MulAction.sum_card_fixedBy_eq_card_orbits_mul_card_group]
432432
congr 1; apply card_congr'; congr; ext;
433433
exact (Setoid.comm' _).trans isConj_iff.symm

Mathlib/Init/Propext.lean

Lines changed: 0 additions & 46 deletions
This file was deleted.

Mathlib/Logic/Basic.lean

Lines changed: 20 additions & 0 deletions
Original file line numberDiff line numberDiff line change
@@ -923,6 +923,26 @@ theorem forall_prop_congr' {p p' : Prop} {q q' : p → Prop} (hq : ∀ h, q h
923923
propext (forall_prop_congr hq hp)
924924
#align forall_prop_congr' forall_prop_congr'
925925

926+
#align forall_congr_eq forall_congr
927+
928+
lemma imp_congr_eq {a b c d : Prop} (h₁ : a = c) (h₂ : b = d) : (a → b) = (c → d) :=
929+
propext (imp_congr h₁.to_iff h₂.to_iff)
930+
931+
lemma imp_congr_ctx_eq {a b c d : Prop} (h₁ : a = c) (h₂ : c → b = d) : (a → b) = (c → d) :=
932+
propext (imp_congr_ctx h₁.to_iff fun hc ↦ (h₂ hc).to_iff)
933+
934+
lemma eq_true_intro (h : a) : a = True := propext (iff_true_intro h)
935+
lemma eq_false_intro (h : ¬a) : a = False := propext (iff_false_intro h)
936+
937+
-- FIXME: `alias` creates `def Iff.eq := propext` instead of `lemma Iff.eq := propext`
938+
@[nolint defLemma] alias Iff.eq := propext
939+
940+
lemma iff_eq_eq : (a ↔ b) = (a = b) := propext ⟨propext, Eq.to_iff⟩
941+
942+
-- They were not used in Lean 3 and there are already lemmas with those names in Lean 4
943+
#noalign eq_false
944+
#noalign eq_true
945+
926946
/-- See `IsEmpty.forall_iff` for the `False` version. -/
927947
@[simp] theorem forall_true_left (p : True → Prop) : (∀ x, p x) ↔ p True.intro :=
928948
forall_prop_of_true _

Mathlib/Logic/Relation.lean

Lines changed: 0 additions & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -5,7 +5,6 @@ Authors: Johannes Hölzl
55
-/
66
import Mathlib.Logic.Function.Basic
77
import Mathlib.Logic.Relator
8-
import Mathlib.Init.Propext
98
import Mathlib.Init.Data.Quot
109
import Mathlib.Tactic.Cases
1110
import Mathlib.Tactic.Use

Mathlib/Order/Filter/Basic.lean

Lines changed: 1 addition & 1 deletion
Original file line numberDiff line numberDiff line change
@@ -1477,7 +1477,7 @@ theorem EventuallyEq.rw {l : Filter α} {f g : α → β} (h : f =ᶠ[l] g) (p :
14771477
#align filter.eventually_eq.rw Filter.EventuallyEq.rw
14781478

14791479
theorem eventuallyEq_set {s t : Set α} {l : Filter α} : s =ᶠ[l] t ↔ ∀ᶠ x in l, x ∈ s ↔ x ∈ t :=
1480-
eventually_congr <| eventually_of_forall fun _ => ⟨Eq.to_iff, Iff.to_eq⟩
1480+
eventually_congr <| eventually_of_forall fun _ ↦ eq_iff_iff
14811481
#align filter.eventually_eq_set Filter.eventuallyEq_set
14821482

14831483
alias ⟨EventuallyEq.mem_iff, Eventually.set_eq⟩ := eventuallyEq_set

0 commit comments

Comments
 (0)