logic.basic
⟷
Mathlib.Logic.Basic
The following section lists changes to this file in mathlib3 and mathlib4 that occured after the initial port. Most recent changes are shown first. Hovering over a commit will show all commits associated with the same mathlib3 commit.
(last sync)
More lemmas about is_clique
, is_n_clique
, edge_set
. Also define clique_free_on
, a local version of clique_free
.
@@ -423,6 +423,15 @@ lemma iff.not (h : a ↔ b) : ¬ a ↔ ¬ b := not_congr h
lemma iff.not_left (h : a ↔ ¬ b) : ¬ a ↔ b := h.not.trans not_not
lemma iff.not_right (h : ¬ a ↔ b) : a ↔ ¬ b := not_not.symm.trans h.not
+protected lemma iff.ne {α β : Sort*} {a b : α} {c d : β} : (a = b ↔ c = d) → (a ≠ b ↔ c ≠ d) :=
+iff.not
+
+lemma iff.ne_left {α β : Sort*} {a b : α} {c d : β} : (a = b ↔ c ≠ d) → (a ≠ b ↔ c = d) :=
+iff.not_left
+
+lemma iff.ne_right {α β : Sort*} {a b : α} {c d : β} : (a ≠ b ↔ c = d) → (a = b ↔ c ≠ d) :=
+iff.not_right
+
/-! ### Declarations about `xor` -/
@[simp] theorem xor_true : xor true = not := funext $ λ a, by simp [xor]
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
@@ -3,7 +3,7 @@ Copyright (c) 2016 Jeremy Avigad. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Jeremy Avigad, Leonardo de Moura
-/
-import tactic.doc_commands
+import tactic.mk_simp_attribute
import tactic.reserved_notation
/-!
@@ -63,7 +63,7 @@ instance psum.inhabited_right {α β} [inhabited β] : inhabited (psum α β) :=
{α} [subsingleton α] : decidable_eq α
| a b := is_true (subsingleton.elim a b)
-@[simp] lemma eq_iff_true_of_subsingleton {α : Sort*} [subsingleton α] (x y : α) :
+@[simp, nontriviality] lemma eq_iff_true_of_subsingleton {α : Sort*} [subsingleton α] (x y : α) :
x = y ↔ true :=
by cc
@@ -274,7 +274,7 @@ theorem imp_and_distrib {α} : (α → b ∧ c) ↔ (α → b) ∧ (α → c) :=
⟨λ h, ⟨λ ha, (h ha).left, λ ha, (h ha).right⟩,
λ h ha, ⟨h.left ha, h.right ha⟩⟩
-@[simp] theorem and_imp : (a ∧ b → c) ↔ (a → b → c) :=
+@[simp, mfld_simps] theorem and_imp : (a ∧ b → c) ↔ (a → b → c) :=
iff.intro (λ h ha hb, h ⟨ha, hb⟩) (λ h ⟨ha, hb⟩, h ha hb)
theorem iff_def : (a ↔ b) ↔ (a → b) ∧ (b → a) :=
@@ -842,7 +842,7 @@ end mem
section equality
variables {α : Sort*} {a b : α}
-@[simp] theorem heq_iff_eq : a == b ↔ a = b :=
+@[simp, mfld_simps] theorem heq_iff_eq : a == b ↔ a = b :=
⟨eq_of_heq, heq_of_eq⟩
theorem proof_irrel_heq {p q : Prop} (hp : p) (hq : q) : hp == hq :=
@@ -865,12 +865,12 @@ theorem eq_equivalence : equivalence (@eq α) :=
⟨eq.refl, @eq.symm _, @eq.trans _⟩
/-- Transport through trivial families is the identity. -/
-@[simp]
+@[simp, transport_simps]
lemma eq_rec_constant {α : Sort*} {a a' : α} {β : Sort*} (y : β) (h : a = a') :
(@eq.rec α a (λ a, β) y a' h) = y :=
by { cases h, refl, }
-@[simp]
+@[simp, transport_simps]
lemma eq_mp_eq_cast {α β : Sort*} (h : α = β) : eq.mp h = cast h := rfl
@[simp]
@@ -1096,6 +1096,7 @@ let ⟨a⟩ := ha in
(λ hb, hb $ h $ λ x, (not_imp.1 (h' x)).1), λ ⟨x, hx⟩ h, hx (h x)⟩
-- TODO: duplicate of a lemma in core
+@[mfld_simps]
theorem forall_true_iff : (α → true) ↔ true :=
implies_true_iff α
@@ -1118,7 +1119,7 @@ exists.elim h (λ x hx, ⟨x, and.left hx⟩)
(∃! x, p x) ↔ ∃ x, p x :=
⟨λ h, h.exists, Exists.imp $ λ x hx, ⟨hx, λ y _, subsingleton.elim y x⟩⟩
-@[simp] theorem forall_const (α : Sort*) [i : nonempty α] : (α → b) ↔ b :=
+@[simp, mfld_simps] theorem forall_const (α : Sort*) [i : nonempty α] : (α → b) ↔ b :=
⟨i.elim, λ hb x, hb⟩
/-- For some reason simp doesn't use `forall_const` to simplify in this case. -/
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
This backports a change that was already made in mathlib4.
Indeed, mathport complains that the change was made, see the comments in https://github.com/leanprover-community/mathlib3port/blob/e3a205b1f51e409563e9e4294f41dd4df61f578a/Mathbin/Logic/Equiv/Basic.lean#L1729-L1735
By backporting this, we can deal with the fallout up-front rather than during porting.
@@ -177,11 +177,11 @@ attribute [symm] ne.symm
lemma ne_comm {α} {a b : α} : a ≠ b ↔ b ≠ a := ⟨ne.symm, ne.symm⟩
-@[simp] lemma eq_iff_eq_cancel_left {b c : α} :
+@[simp] lemma eq_iff_eq_cancel_left {α : Sort*} {b c : α} :
(∀ {a}, a = b ↔ a = c) ↔ (b = c) :=
⟨λ h, by rw [← h], λ h a, by rw h⟩
-@[simp] lemma eq_iff_eq_cancel_right {a b : α} :
+@[simp] lemma eq_iff_eq_cancel_right {α : Sort*} {a b : α} :
(∀ {c}, a = c ↔ b = c) ↔ (a = b) :=
⟨λ h, by rw h, λ h a, by rw h⟩
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
@@ -908,9 +908,17 @@ lemma heq_of_cast_eq :
lemma cast_eq_iff_heq {α β : Sort*} {a : α} {a' : β} {e : α = β} : cast e a = a' ↔ a == a' :=
⟨heq_of_cast_eq _, λ h, by cases h; refl⟩
-lemma rec_heq_of_heq {β} {C : α → Sort*} {x : C a} {y : β} (eq : a = b) (h : x == y) :
- @eq.rec α a C x b eq == y :=
-by subst eq; exact h
+lemma rec_heq_of_heq {β} {C : α → Sort*} {x : C a} {y : β} (e : a = b) (h : x == y) :
+ @eq.rec α a C x b e == y :=
+by subst e; exact h
+
+lemma rec_heq_iff_heq {β} {C : α → Sort*} {x : C a} {y : β} {e : a = b} :
+ @eq.rec α a C x b e == y ↔ x == y :=
+by subst e
+
+lemma heq_rec_iff_heq {β} {C : α → Sort*} {x : β} {y : C a} {e : a = b} :
+ x == @eq.rec α a C y b e ↔ x == y :=
+by subst e
protected lemma eq.congr {x₁ x₂ y₁ y₂ : α} (h₁ : x₁ = y₁) (h₂ : x₂ = y₂) :
(x₁ = x₂) ↔ (y₁ = y₂) :=
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
out_param
classes should not repeat the parents' out_param
(#18291)
This PR is a backport for a mathlib4 fix for a large amount of issues encountered in the port of group_theory.subgroup.basic
, leanprover-community/mathlib4#1797. Subobject classes such as mul_mem_class
and submonoid_class
take a set_like
parameter, and we were used to just copy over the M : out_param Type
declaration from set_like
. Since Lean 3 synthesizes from left to right, we'd fill in the out_param
from set_like
, then the has_mul M
instance would be available for synthesis and we'd be in business. In Lean 4 however, we will often go from right to left, so that MulMemClass ?M S [?inst : Mul ?M]
is handled first, which can't be solved before finding a Mul ?M
instance on the metavariable ?M
, causing search through all Mul
instances.
The solution is: whenever a class is declared that takes another class as parameter (i.e. unbundled inheritance), the out_param
s of the parent class should be unmarked and become in-params in the child class. Then Lean will try to find the parent class instance first, fill in the out_param
s and we'll be able to synthesize the child class instance without problems.
One consequence is that sometimes we have to give slightly more type ascription when talking about membership and the types don't quite align: if M
and M'
are semireducibly defeq, then before zero_mem_class S M
would work to prove (0 : M') ∈ (s : S)
, since the out_param
became a metavariable, was filled in, and then checked (up to semireducibility apparently) for equality. Now M
is checked to equal M'
before applying the instance, with instance-reducible transparency. I don't think this is a huge issue since it feels Lean 4 is stricter about these kinds of equalities anyway.
Mathlib4 pair: leanprover-community/mathlib4#1832
@@ -90,7 +90,7 @@ theorem coe_fn_coe_trans
/-- Non-dependent version of `coe_fn_coe_trans`, helps `rw` figure out the argument. -/
theorem coe_fn_coe_trans'
- {α β γ} {δ : out_param $ _} [has_coe α β] [has_coe_t_aux β γ] [has_coe_to_fun γ (λ _, δ)]
+ {α β γ} {δ : _} [has_coe α β] [has_coe_t_aux β γ] [has_coe_to_fun γ (λ _, δ)]
(x : α) : @coe_fn α _ _ x = @coe_fn β _ _ x := rfl
@[simp] theorem coe_fn_coe_base
@@ -99,7 +99,7 @@ theorem coe_fn_coe_trans'
/-- Non-dependent version of `coe_fn_coe_base`, helps `rw` figure out the argument. -/
theorem coe_fn_coe_base'
- {α β} {γ : out_param $ _} [has_coe α β] [has_coe_to_fun β (λ _, γ)]
+ {α β} {γ : _} [has_coe α β] [has_coe_to_fun β (λ _, γ)]
(x : α) : @coe_fn α _ _ x = @coe_fn β _ _ x := rfl
-- This instance should have low priority, to ensure we follow the chain
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
Also add some supporting lemmas.
@@ -1146,10 +1146,13 @@ by simp [and_comm]
@[simp] theorem forall_eq' {a' : α} : (∀a, a' = a → p a) ↔ p a' :=
by simp [@eq_comm _ a']
-theorem and_forall_ne (a : α) : (p a ∧ ∀ b ≠ a, p b) ↔ ∀ b, p b :=
-by simp only [← @forall_eq _ p a, ← forall_and_distrib, ← or_imp_distrib, classical.em,
+theorem decidable.and_forall_ne [decidable_eq α] (a : α) : (p a ∧ ∀ b ≠ a, p b) ↔ ∀ b, p b :=
+by simp only [← @forall_eq _ p a, ← forall_and_distrib, ← or_imp_distrib, decidable.em,
forall_const]
+theorem and_forall_ne (a : α) : (p a ∧ ∀ b ≠ a, p b) ↔ ∀ b, p b :=
+decidable.and_forall_ne a
+
-- this lemma is needed to simplify the output of `list.mem_cons_iff`
@[simp] theorem forall_eq_or_imp {a' : α} : (∀ a, a = a' ∨ q a → p a) ↔ p a' ∧ ∀ a, q a → p a :=
by simp only [or_imp_distrib, forall_and_distrib, forall_eq]
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(no changes)
(first ported)
mathlib commit https://github.com/leanprover-community/mathlib/commit/65a1391a0106c9204fe45bc73a039f056558cb83
@@ -787,7 +787,7 @@ theorem xor_false : Xor' False = id :=
#print xor_comm /-
theorem xor_comm (a b) : Xor' a b ↔ Xor' b a :=
- or_comm' _ _
+ or_comm _ _
#align xor_comm xor_comm
-/
@@ -813,7 +813,7 @@ theorem xor_not_right : Xor' a ¬b ↔ (a ↔ b) := by by_cases a <;> simp [*]
-/
#print xor_not_not /-
-theorem xor_not_not : Xor' (¬a) ¬b ↔ Xor' a b := by simp [Xor', or_comm', and_comm']
+theorem xor_not_not : Xor' (¬a) ¬b ↔ Xor' a b := by simp [Xor', or_comm, and_comm]
#align xor_not_not xor_not_not
-/
@@ -877,7 +877,7 @@ theorem and_right_comm : (a ∧ b) ∧ c ↔ (a ∧ c) ∧ b := by simp only [an
#print and_and_and_comm /-
theorem and_and_and_comm (a b c d : Prop) : (a ∧ b) ∧ c ∧ d ↔ (a ∧ c) ∧ b ∧ d := by
- rw [← and_assoc', @and_right_comm a, and_assoc']
+ rw [← and_assoc, @and_right_comm a, and_assoc]
#align and_and_and_comm and_and_and_comm
-/
@@ -960,7 +960,7 @@ theorem iff_self_and {p q : Prop} : (p ↔ p ∧ q) ↔ p → q := by rw [@Iff.c
#print iff_and_self /-
@[simp]
-theorem iff_and_self {p q : Prop} : (p ↔ q ∧ p) ↔ p → q := by rw [and_comm', iff_self_and]
+theorem iff_and_self {p q : Prop} : (p ↔ q ∧ p) ↔ p → q := by rw [and_comm, iff_self_and]
#align iff_and_self iff_and_self
-/
@@ -1010,13 +1010,13 @@ theorem or_congr_right (h : b ↔ c) : a ∨ b ↔ a ∨ c :=
#align or_congr_right' or_congr_rightₓ
#print or_right_comm /-
-theorem or_right_comm : (a ∨ b) ∨ c ↔ (a ∨ c) ∨ b := by rw [or_assoc', or_assoc', or_comm' b]
+theorem or_right_comm : (a ∨ b) ∨ c ↔ (a ∨ c) ∨ b := by rw [or_assoc, or_assoc, or_comm b]
#align or.right_comm or_right_comm
-/
#print or_or_or_comm /-
theorem or_or_or_comm (a b c d : Prop) : (a ∨ b) ∨ c ∨ d ↔ (a ∨ c) ∨ b ∨ d := by
- rw [← or_assoc', @or_right_comm a, or_assoc']
+ rw [← or_assoc, @or_right_comm a, or_assoc]
#align or_or_or_comm or_or_or_comm
-/
@@ -1206,7 +1206,7 @@ theorem or_iff_left_iff_imp : (a ∨ b ↔ a) ↔ b → a :=
#print or_iff_right_iff_imp /-
@[simp]
-theorem or_iff_right_iff_imp : (a ∨ b ↔ b) ↔ a → b := by rw [or_comm', or_iff_left_iff_imp]
+theorem or_iff_right_iff_imp : (a ∨ b ↔ b) ↔ a → b := by rw [or_comm, or_iff_left_iff_imp]
#align or_iff_right_iff_imp or_iff_right_iff_imp
-/
@@ -2213,7 +2213,7 @@ theorem exists_and_left {q : Prop} {p : α → Prop} : (∃ x, q ∧ p x) ↔ q
#print exists_and_right /-
@[simp]
theorem exists_and_right {q : Prop} {p : α → Prop} : (∃ x, p x ∧ q) ↔ (∃ x, p x) ∧ q := by
- simp [and_comm']
+ simp [and_comm]
#align exists_and_distrib_right exists_and_right
-/
@@ -2465,7 +2465,7 @@ theorem forall_or_left {q : Prop} {p : α → Prop} : (∀ x, q ∨ p x) ↔ q
#print Decidable.forall_or_right /-
-- See Note [decidable namespace]
protected theorem Decidable.forall_or_right {q : Prop} {p : α → Prop} [Decidable q] :
- (∀ x, p x ∨ q) ↔ (∀ x, p x) ∨ q := by simp [or_comm', Decidable.forall_or_left]
+ (∀ x, p x ∨ q) ↔ (∀ x, p x) ∨ q := by simp [or_comm, Decidable.forall_or_left]
#align decidable.forall_or_distrib_right Decidable.forall_or_right
-/
mathlib commit https://github.com/leanprover-community/mathlib/commit/65a1391a0106c9204fe45bc73a039f056558cb83
@@ -1676,20 +1676,20 @@ theorem proof_irrel_heq {p q : Prop} (hp : p) (hq : q) : HEq hp hq :=
#align proof_irrel_heq proof_irrel_heq
-/
-#print ball_cond_comm /-
+#print forall_cond_comm /-
-- todo: change name
-theorem ball_cond_comm {α} {s : α → Prop} {p : α → α → Prop} :
+theorem forall_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
+#align ball_cond_comm forall_cond_comm
-/
/- ./././Mathport/Syntax/Translate/Basic.lean:642:2: warning: expanding binder collection (a b «expr ∈ » s) -/
-#print ball_mem_comm /-
-theorem ball_mem_comm {α β} [Membership α β] {s : β} {p : α → α → Prop} :
+#print forall_mem_comm /-
+theorem forall_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
+ forall_cond_comm
+#align ball_mem_comm forall_mem_comm
-/
#print ne_of_apply_ne /-
@@ -2756,16 +2756,20 @@ theorem BEx.intro (a : α) (h₁ : p a) (h₂ : P a h₁) : ∃ (x : _) (h : p x
#align bex.intro BEx.intro
-/
-#print ball_congr /-
-theorem ball_congr (H : ∀ x h, P x h ↔ Q x h) : (∀ x h, P x h) ↔ ∀ x h, Q x h :=
+/- warning: ball_congr clashes with forall₂_congr -> forall₂_congr
+Case conversion may be inaccurate. Consider using '#align ball_congr forall₂_congrₓ'. -/
+#print forall₂_congr /-
+theorem forall₂_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
+#align ball_congr forall₂_congr
-/
-#print bex_congr /-
-theorem bex_congr (H : ∀ x h, P x h ↔ Q x h) : (∃ x h, P x h) ↔ ∃ x h, Q x h :=
+/- warning: bex_congr clashes with exists₂_congr -> exists₂_congr
+Case conversion may be inaccurate. Consider using '#align bex_congr exists₂_congrₓ'. -/
+#print exists₂_congr /-
+theorem exists₂_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
+#align bex_congr exists₂_congr
-/
#print bex_eq_left /-
@@ -2810,80 +2814,82 @@ theorem forall_of_ball (H : ∀ x, p x) (h : ∀ x, p x → q x) (x) : q x :=
#align forall_of_ball forall_of_ball
-/
-#print bex_of_exists /-
-theorem bex_of_exists (H : ∀ x, p x) : (∃ x, q x) → ∃ (x : _) (_ : p x), q x
+#print exists_mem_of_exists /-
+theorem exists_mem_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
+#align bex_of_exists exists_mem_of_exists
-/
-#print exists_of_bex /-
-theorem exists_of_bex : (∃ (x : _) (_ : p x), q x) → ∃ x, q x
+#print exists_of_exists_mem /-
+theorem exists_of_exists_mem : (∃ (x : _) (_ : p x), q x) → ∃ x, q x
| ⟨x, _, hq⟩ => ⟨x, hq⟩
-#align exists_of_bex exists_of_bex
+#align exists_of_bex exists_of_exists_mem
-/
-#print bex_imp /-
+#print exists₂_imp /-
@[simp]
-theorem bex_imp : (∃ x h, P x h) → b ↔ ∀ x h, P x h → b := by simp
-#align bex_imp_distrib bex_imp
+theorem exists₂_imp : (∃ x h, P x h) → b ↔ ∀ x h, P x h → b := by simp
+#align bex_imp_distrib exists₂_imp
-/
-#print not_bex /-
-theorem not_bex : (¬∃ x h, P x h) ↔ ∀ x h, ¬P x h :=
- bex_imp
-#align not_bex not_bex
+#print not_exists_mem /-
+theorem not_exists_mem : (¬∃ x h, P x h) ↔ ∀ x h, ¬P x h :=
+ exists₂_imp
+#align not_bex not_exists_mem
-/
-#print not_ball_of_bex_not /-
-theorem not_ball_of_bex_not : (∃ x h, ¬P x h) → ¬∀ x h, P x h
+#print not_forall₂_of_exists₂_not /-
+theorem not_forall₂_of_exists₂_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
+#align not_ball_of_bex_not not_forall₂_of_exists₂_not
-/
-#print Decidable.not_ball /-
+#print Decidable.not_forall₂ /-
-- See Note [decidable namespace]
-protected theorem Decidable.not_ball [Decidable (∃ x h, ¬P x h)] [∀ x h, Decidable (P x h)] :
+protected theorem Decidable.not_forall₂ [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
+ not_forall₂_of_exists₂_not⟩
+#align decidable.not_ball Decidable.not_forall₂
-/
-#print not_ball /-
-theorem not_ball : (¬∀ x h, P x h) ↔ ∃ x h, ¬P x h :=
- Decidable.not_ball
-#align not_ball not_ball
+#print not_forall₂ /-
+theorem not_forall₂ : (¬∀ x h, P x h) ↔ ∃ x h, ¬P x h :=
+ Decidable.not_forall₂
+#align not_ball not_forall₂
-/
-#print ball_true_iff /-
-theorem ball_true_iff (p : α → Prop) : (∀ x, p x → True) ↔ True :=
+/- warning: ball_true_iff clashes with forall_2_true_iff -> forall₂_true_iff
+Case conversion may be inaccurate. Consider using '#align ball_true_iff forall₂_true_iffₓ'. -/
+#print forall₂_true_iff /-
+theorem forall₂_true_iff (p : α → Prop) : (∀ x, p x → True) ↔ True :=
iff_true_intro fun h hrx => trivial
-#align ball_true_iff ball_true_iff
+#align ball_true_iff forall₂_true_iff
-/
-#print ball_and /-
-theorem ball_and : (∀ x h, P x h ∧ Q x h) ↔ (∀ x h, P x h) ∧ ∀ x h, Q x h :=
+#print forall₂_and /-
+theorem forall₂_and : (∀ x h, P x h ∧ Q x h) ↔ (∀ x h, P x h) ∧ ∀ x h, Q x h :=
Iff.trans (forall_congr' fun x => forall_and) forall_and
-#align ball_and_distrib ball_and
+#align ball_and_distrib forall₂_and
-/
-#print bex_or /-
-theorem bex_or : (∃ x h, P x h ∨ Q x h) ↔ (∃ x h, P x h) ∨ ∃ x h, Q x h :=
+#print exists_mem_or /-
+theorem exists_mem_or : (∃ x h, P x h ∨ Q x h) ↔ (∃ x h, P x h) ∨ ∃ x h, Q x h :=
Iff.trans (exists_congr fun x => exists_or) exists_or
-#align bex_or_distrib bex_or
+#align bex_or_distrib exists_mem_or
-/
-#print ball_or_left /-
-theorem ball_or_left : (∀ x, p x ∨ q x → r x) ↔ (∀ x, p x → r x) ∧ ∀ x, q x → r x :=
+#print forall₂_or_left /-
+theorem forall₂_or_left : (∀ x, p x ∨ q x → r x) ↔ (∀ x, p x → r x) ∧ ∀ x, q x → r x :=
Iff.trans (forall_congr' fun x => or_imp) forall_and
-#align ball_or_left_distrib ball_or_left
+#align ball_or_left_distrib forall₂_or_left
-/
-#print bex_or_left /-
-theorem bex_or_left :
+#print exists_mem_or_left /-
+theorem exists_mem_or_left :
(∃ (x : _) (_ : p x ∨ q x), r x) ↔ (∃ (x : _) (_ : p x), r x) ∨ ∃ (x : _) (_ : q x), r x := by
simp only [exists_prop] <;> exact Iff.trans (exists_congr fun x => or_and_right) exists_or
-#align bex_or_left_distrib bex_or_left
+#align bex_or_left_distrib exists_mem_or_left
-/
end BoundedQuantifiers
@@ -2892,12 +2898,10 @@ namespace Classical
attribute [local instance] prop_decidable
-/- warning: classical.not_ball clashes with not_ball -> not_ball
-Case conversion may be inaccurate. Consider using '#align classical.not_ball not_ballₓ'. -/
#print not_ball /-
theorem not_ball {α : Sort _} {p : α → Prop} {P : ∀ x : α, p x → Prop} :
(¬∀ x h, P x h) ↔ ∃ x h, ¬P x h :=
- not_ball
+ not_forall₂
#align classical.not_ball not_ball
-/
mathlib commit https://github.com/leanprover-community/mathlib/commit/65a1391a0106c9204fe45bc73a039f056558cb83
@@ -1684,7 +1684,7 @@ theorem ball_cond_comm {α} {s : α → Prop} {p : α → α → Prop} :
#align ball_cond_comm ball_cond_comm
-/
-/- ./././Mathport/Syntax/Translate/Basic.lean:641:2: warning: expanding binder collection (a b «expr ∈ » s) -/
+/- ./././Mathport/Syntax/Translate/Basic.lean:642:2: warning: expanding binder collection (a b «expr ∈ » s) -/
#print ball_mem_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 :=
@@ -2230,7 +2230,7 @@ theorem forall_eq' {a' : α} : (∀ a, a' = a → p a) ↔ p a' := by simp [@eq_
#align forall_eq' forall_eq'
-/
-/- ./././Mathport/Syntax/Translate/Basic.lean:641:2: warning: expanding binder collection (b «expr ≠ » a) -/
+/- ./././Mathport/Syntax/Translate/Basic.lean:642:2: warning: expanding binder collection (b «expr ≠ » a) -/
#print Decidable.and_forall_ne /-
theorem Decidable.and_forall_ne [DecidableEq α] (a : α) :
(p a ∧ ∀ (b) (_ : b ≠ a), p b) ↔ ∀ b, p b := by
@@ -2238,7 +2238,7 @@ theorem Decidable.and_forall_ne [DecidableEq α] (a : α) :
#align decidable.and_forall_ne Decidable.and_forall_ne
-/
-/- ./././Mathport/Syntax/Translate/Basic.lean:641:2: warning: expanding binder collection (b «expr ≠ » a) -/
+/- ./././Mathport/Syntax/Translate/Basic.lean:642:2: warning: expanding binder collection (b «expr ≠ » a) -/
#print and_forall_ne /-
theorem and_forall_ne (a : α) : (p a ∧ ∀ (b) (_ : b ≠ a), p b) ↔ ∀ b, p b :=
Decidable.and_forall_ne a
mathlib commit https://github.com/leanprover-community/mathlib/commit/65a1391a0106c9204fe45bc73a039f056558cb83
@@ -709,7 +709,7 @@ theorem imp_not_self : a → ¬a ↔ ¬a :=
#print Decidable.not_imp_self /-
theorem Decidable.not_imp_self [Decidable a] : ¬a → a ↔ a := by have := @imp_not_self ¬a;
- rwa [Decidable.not_not] at this
+ rwa [Decidable.not_not] at this
#align decidable.not_imp_self Decidable.not_imp_self
-/
@@ -967,7 +967,7 @@ theorem iff_and_self {p q : Prop} : (p ↔ q ∧ p) ↔ p → q := by rw [and_co
#print and_congr_right_iff /-
@[simp]
theorem and_congr_right_iff : (a ∧ b ↔ a ∧ c) ↔ a → (b ↔ c) :=
- ⟨fun h ha => by simp [ha] at h <;> exact h, and_congr_right⟩
+ ⟨fun h ha => by simp [ha] at h <;> exact h, and_congr_right⟩
#align and.congr_right_iff and_congr_right_iff
-/
@@ -2609,7 +2609,7 @@ theorem ExistsUnique.elim₂ {α : Sort _} {p : α → Sort _} [∀ x, Subsingle
{q : ∀ (x) (h : p x), Prop} {b : Prop} (h₂ : ∃! (x : _) (h : p x), q x h)
(h₁ : ∀ (x) (h : p x), q x h → (∀ (y) (hy : p y), q y hy → y = x) → b) : b :=
by
- simp only [exists_unique_iff_exists] at h₂
+ simp only [exists_unique_iff_exists] at h₂
apply h₂.elim
exact fun x ⟨hxp, hxq⟩ H => h₁ x hxp hxq fun y hyp hyq => H y ⟨hyp, hyq⟩
#align exists_unique.elim2 ExistsUnique.elim₂
@@ -2637,7 +2637,7 @@ theorem ExistsUnique.unique₂ {α : Sort _} {p : α → Sort _} [∀ x, Subsing
{q : ∀ (x : α) (hx : p x), Prop} (h : ∃! (x : _) (hx : p x), q x hx) {y₁ y₂ : α} (hpy₁ : p y₁)
(hqy₁ : q y₁ hpy₁) (hpy₂ : p y₂) (hqy₂ : q y₂ hpy₂) : y₁ = y₂ :=
by
- simp only [exists_unique_iff_exists] at h
+ simp only [exists_unique_iff_exists] at h
exact h.unique ⟨hpy₁, hqy₁⟩ ⟨hpy₂, hqy₂⟩
#align exists_unique.unique2 ExistsUnique.unique₂
-/
mathlib commit https://github.com/leanprover-community/mathlib/commit/65a1391a0106c9204fe45bc73a039f056558cb83
@@ -791,7 +791,7 @@ theorem xor_comm (a b) : Xor' a b ↔ Xor' b a :=
#align xor_comm xor_comm
-/
-instance : IsCommutative Prop Xor' :=
+instance : Std.Commutative Prop Xor' :=
⟨fun a b => propext <| xor_comm a b⟩
#print xor_self /-
mathlib commit https://github.com/leanprover-community/mathlib/commit/65a1391a0106c9204fe45bc73a039f056558cb83
@@ -1349,16 +1349,16 @@ theorem not_imp_of_and_not : a ∧ ¬b → ¬(a → b)
#align not_imp_of_and_not not_imp_of_and_not
-/
-#print Decidable.not_imp /-
+#print Decidable.not_imp_iff_and_not /-
-- See Note [decidable namespace]
-protected theorem Decidable.not_imp [Decidable a] : ¬(a → b) ↔ a ∧ ¬b :=
+protected theorem Decidable.not_imp_iff_and_not [Decidable a] : ¬(a → b) ↔ a ∧ ¬b :=
⟨fun h => ⟨Decidable.of_not_imp h, not_of_not_imp h⟩, not_imp_of_and_not⟩
-#align decidable.not_imp Decidable.not_imp
+#align decidable.not_imp Decidable.not_imp_iff_and_not
-/
#print not_imp /-
theorem not_imp : ¬(a → b) ↔ a ∧ ¬b :=
- Decidable.not_imp
+ Decidable.not_imp_iff_and_not
#align not_imp not_imp
-/
@@ -1526,25 +1526,25 @@ theorem not_and_of_not_or_not (h : ¬a ∨ ¬b) : ¬(a ∧ b)
#align not_and_of_not_or_not not_and_of_not_or_not
-/
-#print Decidable.not_and /-
+#print Decidable.not_and_iff_or_not_not /-
-- See Note [decidable namespace]
-protected theorem Decidable.not_and [Decidable a] : ¬(a ∧ b) ↔ ¬a ∨ ¬b :=
+protected theorem Decidable.not_and_iff_or_not_not [Decidable a] : ¬(a ∧ b) ↔ ¬a ∨ ¬b :=
⟨fun h => if ha : a then Or.inr fun hb => h ⟨ha, hb⟩ else Or.inl ha, not_and_of_not_or_not⟩
-#align decidable.not_and_distrib Decidable.not_and
+#align decidable.not_and_distrib Decidable.not_and_iff_or_not_not
-/
-#print Decidable.not_and' /-
+#print Decidable.not_and_iff_or_not_not' /-
-- See Note [decidable namespace]
-protected theorem Decidable.not_and' [Decidable b] : ¬(a ∧ b) ↔ ¬a ∨ ¬b :=
+protected theorem Decidable.not_and_iff_or_not_not' [Decidable b] : ¬(a ∧ b) ↔ ¬a ∨ ¬b :=
⟨fun h => if hb : b then Or.inl fun ha => h ⟨ha, hb⟩ else Or.inr hb, not_and_of_not_or_not⟩
-#align decidable.not_and_distrib' Decidable.not_and'
+#align decidable.not_and_distrib' Decidable.not_and_iff_or_not_not'
-/
#print not_and_or /-
/-- One of de Morgan's laws: the negation of a conjunction is logically equivalent to the
disjunction of the negations. -/
theorem not_and_or : ¬(a ∧ b) ↔ ¬a ∨ ¬b :=
- Decidable.not_and
+ Decidable.not_and_iff_or_not_not
#align not_and_distrib not_and_or
-/
@@ -1585,7 +1585,7 @@ theorem or_iff_not_and_not : a ∨ b ↔ ¬(¬a ∧ ¬b) :=
#print Decidable.and_iff_not_or_not /-
-- See Note [decidable namespace]
protected theorem Decidable.and_iff_not_or_not [Decidable a] [Decidable b] : a ∧ b ↔ ¬(¬a ∨ ¬b) :=
- by rw [← Decidable.not_and, Decidable.not_not]
+ by rw [← Decidable.not_and_iff_or_not_not, Decidable.not_not]
#align decidable.and_iff_not_or_not Decidable.and_iff_not_or_not
-/
mathlib commit https://github.com/leanprover-community/mathlib/commit/65a1391a0106c9204fe45bc73a039f056558cb83
@@ -276,7 +276,7 @@ theorem eq_iff_eq_cancel_right {α : Sort _} {a b : α} : (∀ {c}, a = c ↔ b
-/
#print Fact /-
-/- ./././Mathport/Syntax/Translate/Command.lean:404:30: infer kinds are unsupported in Lean 4: #[`out] [] -/
+/- ./././Mathport/Syntax/Translate/Command.lean:400:30: infer kinds are unsupported in Lean 4: #[`out] [] -/
/-- Wrapper for adding elementary propositions to the type class systems.
Warning: this can easily be abused. See the rest of this docstring for details.
mathlib commit https://github.com/leanprover-community/mathlib/commit/65a1391a0106c9204fe45bc73a039f056558cb83
@@ -276,7 +276,7 @@ theorem eq_iff_eq_cancel_right {α : Sort _} {a b : α} : (∀ {c}, a = c ↔ b
-/
#print Fact /-
-/- ./././Mathport/Syntax/Translate/Command.lean:394:30: infer kinds are unsupported in Lean 4: #[`out] [] -/
+/- ./././Mathport/Syntax/Translate/Command.lean:404:30: infer kinds are unsupported in Lean 4: #[`out] [] -/
/-- Wrapper for adding elementary propositions to the type class systems.
Warning: this can easily be abused. See the rest of this docstring for details.
mathlib commit https://github.com/leanprover-community/mathlib/commit/65a1391a0106c9204fe45bc73a039f056558cb83
@@ -750,17 +750,23 @@ theorem Iff.not_right (h : ¬a ↔ b) : a ↔ ¬b :=
#align iff.not_right Iff.not_right
-/
+#print Iff.ne /-
protected theorem Iff.ne {α β : Sort _} {a b : α} {c d : β} : (a = b ↔ c = d) → (a ≠ b ↔ c ≠ d) :=
Iff.not
#align iff.ne Iff.ne
+-/
+#print Iff.ne_left /-
theorem Iff.ne_left {α β : Sort _} {a b : α} {c d : β} : (a = b ↔ c ≠ d) → (a ≠ b ↔ c = d) :=
Iff.not_left
#align iff.ne_left Iff.ne_left
+-/
+#print Iff.ne_right /-
theorem Iff.ne_right {α β : Sort _} {a b : α} {c d : β} : (a ≠ b ↔ c = d) → (a = b ↔ c ≠ d) :=
Iff.not_right
#align iff.ne_right Iff.ne_right
+-/
/-! ### Declarations about `xor` -/
mathlib commit https://github.com/leanprover-community/mathlib/commit/65a1391a0106c9204fe45bc73a039f056558cb83
@@ -1080,10 +1080,10 @@ protected theorem Decidable.or_iff_not_imp_left [Decidable a] : a ∨ b ↔ ¬a
#align decidable.or_iff_not_imp_left Decidable.or_iff_not_imp_left
-/
-#print or_iff_not_imp_left /-
-theorem or_iff_not_imp_left : a ∨ b ↔ ¬a → b :=
+#print Classical.or_iff_not_imp_left /-
+theorem Classical.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
+#align or_iff_not_imp_left Classical.or_iff_not_imp_left
-/
-- See Note [decidable namespace]
@@ -1091,10 +1091,10 @@ protected theorem Decidable.or_iff_not_imp_right [Decidable b] : a ∨ b ↔ ¬b
or_comm.trans Decidable.or_iff_not_imp_left
#align decidable.or_iff_not_imp_right Decidable.or_iff_not_imp_rightₓ
-#print or_iff_not_imp_right /-
-theorem or_iff_not_imp_right : a ∨ b ↔ ¬b → a :=
+#print Classical.or_iff_not_imp_right /-
+theorem Classical.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
+#align or_iff_not_imp_right Classical.or_iff_not_imp_right
-/
#print Decidable.not_or_of_imp /-
mathlib commit https://github.com/leanprover-community/mathlib/commit/3365b20c2ffa7c35e47e5209b89ba9abdddf3ffe
@@ -6,7 +6,7 @@ Authors: Jeremy Avigad, Leonardo de Moura
import Tactic.MkSimpAttribute
import Tactic.ReservedNotation
-#align_import logic.basic from "leanprover-community/mathlib"@"48fb5b5280e7c81672afc9524185ae994553ebf4"
+#align_import logic.basic from "leanprover-community/mathlib"@"3365b20c2ffa7c35e47e5209b89ba9abdddf3ffe"
/-!
# Basic logic properties
@@ -750,6 +750,18 @@ theorem Iff.not_right (h : ¬a ↔ b) : a ↔ ¬b :=
#align iff.not_right Iff.not_right
-/
+protected theorem Iff.ne {α β : Sort _} {a b : α} {c d : β} : (a = b ↔ c = d) → (a ≠ b ↔ c ≠ d) :=
+ Iff.not
+#align iff.ne Iff.ne
+
+theorem Iff.ne_left {α β : Sort _} {a b : α} {c d : β} : (a = b ↔ c ≠ d) → (a ≠ b ↔ c = d) :=
+ Iff.not_left
+#align iff.ne_left Iff.ne_left
+
+theorem Iff.ne_right {α β : Sort _} {a b : α} {c d : β} : (a ≠ b ↔ c = d) → (a = b ↔ c ≠ d) :=
+ Iff.not_right
+#align iff.ne_right Iff.ne_right
+
/-! ### Declarations about `xor` -/
mathlib commit https://github.com/leanprover-community/mathlib/commit/b1abe23ae96fef89ad30d9f4362c307f72a55010
@@ -2054,11 +2054,11 @@ protected theorem Decidable.not_forall {p : α → Prop} [Decidable (∃ x, ¬p
#align decidable.not_forall Decidable.not_forall
-/
-#print not_forall /-
+#print Classical.not_forall /-
@[simp]
-theorem not_forall {p : α → Prop} : (¬∀ x, p x) ↔ ∃ x, ¬p x :=
+theorem Classical.not_forall {p : α → Prop} : (¬∀ x, p x) ↔ ∃ x, ¬p x :=
Decidable.not_forall
-#align not_forall not_forall
+#align not_forall Classical.not_forall
-/
#print Decidable.not_forall_not /-
@@ -2081,11 +2081,11 @@ protected theorem Decidable.not_exists_not [∀ x, Decidable (p x)] : (¬∃ x,
#align decidable.not_exists_not Decidable.not_exists_not
-/
-#print not_exists_not /-
+#print Classical.not_exists_not /-
@[simp]
-theorem not_exists_not : (¬∃ x, ¬p x) ↔ ∀ x, p x :=
+theorem Classical.not_exists_not : (¬∃ x, ¬p x) ↔ ∀ x, p x :=
Decidable.not_exists_not
-#align not_exists_not not_exists_not
+#align not_exists_not Classical.not_exists_not
-/
#print forall_imp_iff_exists_imp /-
@@ -2945,13 +2945,13 @@ theorem ite_eq_right_iff : ite P a b = b ↔ P → a = b :=
#print dite_ne_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]
+ rw [Ne.def, dite_eq_left_iff, Classical.not_forall]; exact exists_congr fun h => by rw [ne_comm]
#align dite_ne_left_iff dite_ne_left_iff
-/
#print dite_ne_right_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]
+ simp only [Ne.def, dite_eq_right_iff, Classical.not_forall]
#align dite_ne_right_iff dite_ne_right_iff
-/
mathlib commit https://github.com/leanprover-community/mathlib/commit/b1abe23ae96fef89ad30d9f4362c307f72a55010
@@ -2354,33 +2354,33 @@ theorem exists_or_eq_right' (y : α) (p : α → Prop) : ∃ x : α, p x ∨ y =
#align exists_or_eq_right' exists_or_eq_right'
-/
-#print forall_apply_eq_imp_iff /-
+#print forall_apply_eq_imp_iff' /-
@[simp]
-theorem forall_apply_eq_imp_iff {f : α → β} {p : β → Prop} :
+theorem forall_apply_eq_imp_iff' {f : α → β} {p : β → Prop} :
(∀ a, ∀ b, f a = b → p b) ↔ ∀ a, p (f a) :=
⟨fun h a => h a (f a) rfl, fun h a b hab => hab ▸ h a⟩
-#align forall_apply_eq_imp_iff forall_apply_eq_imp_iff
+#align forall_apply_eq_imp_iff forall_apply_eq_imp_iff'
-/
-#print forall_apply_eq_imp_iff' /-
+#print forall_apply_eq_imp_iff /-
@[simp]
-theorem forall_apply_eq_imp_iff' {f : α → β} {p : β → Prop} :
+theorem forall_apply_eq_imp_iff {f : α → β} {p : β → Prop} :
(∀ b, ∀ a, f a = b → p b) ↔ ∀ a, p (f a) := by rw [forall_swap]; simp
-#align forall_apply_eq_imp_iff' forall_apply_eq_imp_iff'
+#align forall_apply_eq_imp_iff' forall_apply_eq_imp_iff
-/
-#print forall_eq_apply_imp_iff /-
+#print forall_eq_apply_imp_iff' /-
@[simp]
-theorem forall_eq_apply_imp_iff {f : α → β} {p : β → Prop} :
+theorem forall_eq_apply_imp_iff' {f : α → β} {p : β → Prop} :
(∀ a, ∀ b, b = f a → p b) ↔ ∀ a, p (f a) := by simp [@eq_comm _ _ (f _)]
-#align forall_eq_apply_imp_iff forall_eq_apply_imp_iff
+#align forall_eq_apply_imp_iff forall_eq_apply_imp_iff'
-/
-#print forall_eq_apply_imp_iff' /-
+#print forall_eq_apply_imp_iff /-
@[simp]
-theorem forall_eq_apply_imp_iff' {f : α → β} {p : β → Prop} :
+theorem forall_eq_apply_imp_iff {f : α → β} {p : β → Prop} :
(∀ b, ∀ a, b = f a → p b) ↔ ∀ a, p (f a) := by rw [forall_swap]; simp
-#align forall_eq_apply_imp_iff' forall_eq_apply_imp_iff'
+#align forall_eq_apply_imp_iff' forall_eq_apply_imp_iff
-/
#print forall_apply_eq_imp_iff₂ /-
mathlib commit https://github.com/leanprover-community/mathlib/commit/ce64cd319bb6b3e82f31c2d38e79080d377be451
@@ -3,8 +3,8 @@ Copyright (c) 2016 Jeremy Avigad. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Jeremy Avigad, Leonardo de Moura
-/
-import Mathbin.Tactic.MkSimpAttribute
-import Mathbin.Tactic.ReservedNotation
+import Tactic.MkSimpAttribute
+import Tactic.ReservedNotation
#align_import logic.basic from "leanprover-community/mathlib"@"48fb5b5280e7c81672afc9524185ae994553ebf4"
@@ -276,7 +276,7 @@ theorem eq_iff_eq_cancel_right {α : Sort _} {a b : α} : (∀ {c}, a = c ↔ b
-/
#print Fact /-
-/- ./././Mathport/Syntax/Translate/Command.lean:393:30: infer kinds are unsupported in Lean 4: #[`out] [] -/
+/- ./././Mathport/Syntax/Translate/Command.lean:394:30: infer kinds are unsupported in Lean 4: #[`out] [] -/
/-- Wrapper for adding elementary propositions to the type class systems.
Warning: this can easily be abused. See the rest of this docstring for details.
@@ -1666,7 +1666,7 @@ theorem ball_cond_comm {α} {s : α → Prop} {p : α → α → Prop} :
#align ball_cond_comm ball_cond_comm
-/
-/- ./././Mathport/Syntax/Translate/Basic.lean:635:2: warning: expanding binder collection (a b «expr ∈ » s) -/
+/- ./././Mathport/Syntax/Translate/Basic.lean:641:2: warning: expanding binder collection (a b «expr ∈ » s) -/
#print ball_mem_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 :=
@@ -2212,7 +2212,7 @@ theorem forall_eq' {a' : α} : (∀ a, a' = a → p a) ↔ p a' := by simp [@eq_
#align forall_eq' forall_eq'
-/
-/- ./././Mathport/Syntax/Translate/Basic.lean:635:2: warning: expanding binder collection (b «expr ≠ » a) -/
+/- ./././Mathport/Syntax/Translate/Basic.lean:641:2: warning: expanding binder collection (b «expr ≠ » a) -/
#print Decidable.and_forall_ne /-
theorem Decidable.and_forall_ne [DecidableEq α] (a : α) :
(p a ∧ ∀ (b) (_ : b ≠ a), p b) ↔ ∀ b, p b := by
@@ -2220,7 +2220,7 @@ theorem Decidable.and_forall_ne [DecidableEq α] (a : α) :
#align decidable.and_forall_ne Decidable.and_forall_ne
-/
-/- ./././Mathport/Syntax/Translate/Basic.lean:635:2: warning: expanding binder collection (b «expr ≠ » a) -/
+/- ./././Mathport/Syntax/Translate/Basic.lean:641:2: warning: expanding binder collection (b «expr ≠ » a) -/
#print and_forall_ne /-
theorem and_forall_ne (a : α) : (p a ∧ ∀ (b) (_ : b ≠ a), p b) ↔ ∀ b, p b :=
Decidable.and_forall_ne a
mathlib commit https://github.com/leanprover-community/mathlib/commit/8ea5598db6caeddde6cb734aa179cc2408dbd345
@@ -2,15 +2,12 @@
Copyright (c) 2016 Jeremy Avigad. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Jeremy Avigad, Leonardo de Moura
-
-! This file was ported from Lean 3 source module logic.basic
-! leanprover-community/mathlib commit 48fb5b5280e7c81672afc9524185ae994553ebf4
-! Please do not edit these lines, except to modify the commit id
-! if you have ported upstream changes.
-/
import Mathbin.Tactic.MkSimpAttribute
import Mathbin.Tactic.ReservedNotation
+#align_import logic.basic from "leanprover-community/mathlib"@"48fb5b5280e7c81672afc9524185ae994553ebf4"
+
/-!
# Basic logic properties
@@ -1669,7 +1666,7 @@ theorem ball_cond_comm {α} {s : α → Prop} {p : α → α → Prop} :
#align ball_cond_comm ball_cond_comm
-/
-/- ./././Mathport/Syntax/Translate/Basic.lean:638:2: warning: expanding binder collection (a b «expr ∈ » s) -/
+/- ./././Mathport/Syntax/Translate/Basic.lean:635:2: warning: expanding binder collection (a b «expr ∈ » s) -/
#print ball_mem_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 :=
@@ -2215,7 +2212,7 @@ theorem forall_eq' {a' : α} : (∀ a, a' = a → p a) ↔ p a' := by simp [@eq_
#align forall_eq' forall_eq'
-/
-/- ./././Mathport/Syntax/Translate/Basic.lean:638:2: warning: expanding binder collection (b «expr ≠ » a) -/
+/- ./././Mathport/Syntax/Translate/Basic.lean:635:2: warning: expanding binder collection (b «expr ≠ » a) -/
#print Decidable.and_forall_ne /-
theorem Decidable.and_forall_ne [DecidableEq α] (a : α) :
(p a ∧ ∀ (b) (_ : b ≠ a), p b) ↔ ∀ b, p b := by
@@ -2223,7 +2220,7 @@ theorem Decidable.and_forall_ne [DecidableEq α] (a : α) :
#align decidable.and_forall_ne Decidable.and_forall_ne
-/
-/- ./././Mathport/Syntax/Translate/Basic.lean:638:2: warning: expanding binder collection (b «expr ≠ » a) -/
+/- ./././Mathport/Syntax/Translate/Basic.lean:635:2: warning: expanding binder collection (b «expr ≠ » a) -/
#print and_forall_ne /-
theorem and_forall_ne (a : α) : (p a ∧ ∀ (b) (_ : b ≠ a), p b) ↔ ∀ b, p b :=
Decidable.and_forall_ne a
mathlib commit https://github.com/leanprover-community/mathlib/commit/728ef9dbb281241906f25cbeb30f90d83e0bb451
@@ -4,11 +4,11 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Jeremy Avigad, Leonardo de Moura
! This file was ported from Lean 3 source module logic.basic
-! leanprover-community/mathlib commit d2d8742b0c21426362a9dacebc6005db895ca963
+! leanprover-community/mathlib commit 48fb5b5280e7c81672afc9524185ae994553ebf4
! Please do not edit these lines, except to modify the commit id
! if you have ported upstream changes.
-/
-import Mathbin.Tactic.DocCommands
+import Mathbin.Tactic.MkSimpAttribute
import Mathbin.Tactic.ReservedNotation
/-!
@@ -94,7 +94,7 @@ instance (priority := 10) decidableEq_of_subsingleton {α} [Subsingleton α] : D
-/
#print eq_iff_true_of_subsingleton /-
-@[simp]
+@[simp, nontriviality]
theorem eq_iff_true_of_subsingleton {α : Sort _} [Subsingleton α] (x y : α) : x = y ↔ True := by cc
#align eq_iff_true_of_subsingleton eq_iff_true_of_subsingleton
-/
@@ -434,7 +434,7 @@ theorem imp_and {α} : α → b ∧ c ↔ (α → b) ∧ (α → c) :=
-/
#print and_imp /-
-@[simp]
+@[simp, mfld_simps]
theorem and_imp : a ∧ b → c ↔ a → b → c :=
Iff.intro (fun h ha hb => h ⟨ha, hb⟩) fun h ⟨ha, hb⟩ => h ha hb
#align and_imp and_imp
@@ -1647,7 +1647,7 @@ section Equality
variable {α : Sort _} {a b : α}
#print heq_iff_eq /-
-@[simp]
+@[simp, mfld_simps]
theorem heq_iff_eq : HEq a b ↔ a = b :=
⟨eq_of_hEq, hEq_of_eq⟩
#align heq_iff_eq heq_iff_eq
@@ -1691,14 +1691,14 @@ theorem eq_equivalence : Equivalence (@Eq α) :=
#print eq_rec_constant /-
/-- Transport through trivial families is the identity. -/
-@[simp]
+@[simp, transport_simps]
theorem eq_rec_constant {α : Sort _} {a a' : α} {β : Sort _} (y : β) (h : a = a') :
@Eq.ndrec α a (fun a => β) y a' h = y := by cases h; rfl
#align eq_rec_constant eq_rec_constant
-/
#print eq_mp_eq_cast /-
-@[simp]
+@[simp, transport_simps]
theorem eq_mp_eq_cast {α β : Sort _} (h : α = β) : Eq.mp h = cast h :=
rfl
#align eq_mp_eq_cast eq_mp_eq_cast
@@ -2103,6 +2103,7 @@ theorem forall_imp_iff_exists_imp [ha : Nonempty α] : (∀ x, p x) → b ↔
#print forall_true_iff /-
-- TODO: duplicate of a lemma in core
+@[mfld_simps]
theorem forall_true_iff : α → True ↔ True :=
imp_true_iff α
#align forall_true_iff forall_true_iff
@@ -2148,7 +2149,7 @@ theorem exists_unique_iff_exists {α : Sort _} [Subsingleton α] {p : α → Pro
-/
#print forall_const /-
-@[simp]
+@[simp, mfld_simps]
theorem forall_const (α : Sort _) [i : Nonempty α] : α → b ↔ b :=
⟨i.elim, fun hb x => hb⟩
#align forall_const forall_const
mathlib commit https://github.com/leanprover-community/mathlib/commit/9fb8964792b4237dac6200193a0d533f1b3f7423
@@ -216,23 +216,31 @@ theorem not_nonempty_pempty : ¬Nonempty PEmpty := fun ⟨h⟩ => h.elim
#align not_nonempty_pempty not_nonempty_pempty
-/
+#print congr_heq /-
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
+-/
+#print congr_arg_heq /-
theorem congr_arg_heq {α} {β : α → Sort _} (f : ∀ a, β a) :
∀ {a₁ a₂ : α}, a₁ = a₂ → HEq (f a₁) (f a₂)
| a, _, rfl => HEq.rfl
#align congr_arg_heq congr_arg_heq
+-/
+#print ULift.down_injective /-
theorem ULift.down_injective {α : Sort _} : Function.Injective (@ULift.down α)
| ⟨a⟩, ⟨b⟩, rfl => rfl
#align ulift.down_injective ULift.down_injective
+-/
+#print ULift.down_inj /-
@[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
+-/
#print PLift.down_injective /-
theorem PLift.down_injective {α : Sort _} : Function.Injective (@PLift.down α)
@@ -271,7 +279,7 @@ theorem eq_iff_eq_cancel_right {α : Sort _} {a b : α} : (∀ {c}, a = c ↔ b
-/
#print Fact /-
-/- ./././Mathport/Syntax/Translate/Command.lean:394:30: infer kinds are unsupported in Lean 4: #[`out] [] -/
+/- ./././Mathport/Syntax/Translate/Command.lean:393:30: infer kinds are unsupported in Lean 4: #[`out] [] -/
/-- Wrapper for adding elementary propositions to the type class systems.
Warning: this can easily be abused. See the rest of this docstring for details.
@@ -715,9 +723,11 @@ theorem not_imp_self : ¬a → a ↔ a :=
#align not_imp_self not_imp_self
-/
+#print Imp.swap /-
theorem Imp.swap : a → b → c ↔ b → a → c :=
⟨swap, swap⟩
#align imp.swap Imp.swap
+-/
#print imp_not_comm /-
theorem imp_not_comm : a → ¬b ↔ b → ¬a :=
@@ -760,9 +770,11 @@ theorem xor_false : Xor' False = id :=
#align xor_false xor_false
-/
+#print xor_comm /-
theorem xor_comm (a b) : Xor' a b ↔ Xor' b a :=
or_comm' _ _
#align xor_comm xor_comm
+-/
instance : IsCommutative Prop Xor' :=
⟨fun a b => propext <| xor_comm a b⟩
@@ -899,9 +911,11 @@ theorem and_iff_left_of_imp {a b : Prop} (h : a → b) : a ∧ b ↔ a :=
#align and_iff_left_of_imp and_iff_left_of_imp
-/
+#print and_iff_right_of_imp /-
theorem and_iff_right_of_imp {a b : Prop} (h : b → a) : a ∧ b ↔ b :=
Iff.intro And.right fun hb => ⟨h hb, hb⟩
#align and_iff_right_of_imp and_iff_right_of_imp
+-/
#print ne_and_eq_iff_right /-
theorem ne_and_eq_iff_right {α : Sort _} {a b c : α} (h : b ≠ c) : a ≠ b ∧ a = c ↔ a = c :=
@@ -1020,13 +1034,17 @@ theorem or_of_or_of_imp_of_imp (h₁ : a ∨ b) (h₂ : a → c) (h₃ : b → d
#align or_of_or_of_imp_of_imp or_of_or_of_imp_of_imp
-/
+#print or_of_or_of_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
+-/
+#print or_of_or_of_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
+-/
#print Or.elim3 /-
theorem Or.elim3 (h : a ∨ b ∨ c) (ha : a → d) (hb : b → d) (hc : c → d) : d :=
@@ -1034,9 +1052,11 @@ theorem Or.elim3 (h : a ∨ b ∨ c) (ha : a → d) (hb : b → d) (hc : c → d
#align or.elim3 Or.elim3
-/
+#print Or.imp3 /-
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
+-/
#print or_imp /-
theorem or_imp : a ∨ b → c ↔ (a → c) ∧ (b → c) :=
@@ -1112,9 +1132,11 @@ protected theorem Decidable.imp_iff_or_not [Decidable b] : b → a ↔ a ∨ ¬b
Decidable.imp_iff_not_or.trans or_comm
#align decidable.imp_iff_or_not Decidable.imp_iff_or_notₓ
+#print imp_iff_or_not /-
theorem imp_iff_or_not : b → a ↔ a ∨ ¬b :=
Decidable.imp_iff_or_not
#align imp_iff_or_not imp_iff_or_not
+-/
#print Decidable.not_imp_not /-
-- See Note [decidable namespace]
@@ -1136,14 +1158,18 @@ protected theorem Function.mtr : (¬a → ¬b) → b → a :=
#align function.mtr Function.mtr
-/
+#print Decidable.or_congr_left' /-
-- See Note [decidable namespace]
protected theorem Decidable.or_congr_left' [Decidable c] (h : ¬c → (a ↔ b)) : a ∨ c ↔ b ∨ c := by
rw [Decidable.or_iff_not_imp_right, Decidable.or_iff_not_imp_right]; exact imp_congr_right h
#align decidable.or_congr_left Decidable.or_congr_left'
+-/
+#print or_congr_left' /-
theorem or_congr_left' (h : ¬c → (a ↔ b)) : a ∨ c ↔ b ∨ c :=
Decidable.or_congr_left' h
#align or_congr_left or_congr_left'
+-/
#print Decidable.or_congr_right' /-
-- See Note [decidable namespace]
@@ -1291,10 +1317,12 @@ theorem imp_or : a → b ∨ c ↔ (a → b) ∨ (a → c) :=
#align imp_or_distrib imp_or
-/
+#print Decidable.imp_or' /-
-- See Note [decidable namespace]
protected theorem Decidable.imp_or' [Decidable b] : a → b ∨ c ↔ (a → b) ∨ (a → c) := by
by_cases b <;> simp [h, or_iff_right_of_imp ((· ∘ ·) False.elim)]
#align decidable.imp_or_distrib' Decidable.imp_or'
+-/
theorem imp_or' : a → b ∨ c ↔ (a → b) ∨ (a → c) :=
Decidable.imp_or'
@@ -1370,10 +1398,12 @@ theorem not_iff_comm : (¬a ↔ b) ↔ (¬b ↔ a) :=
#align not_iff_comm not_iff_comm
-/
+#print Decidable.not_iff /-
-- See Note [decidable namespace]
protected theorem Decidable.not_iff : ∀ [Decidable b], ¬(a ↔ b) ↔ (¬a ↔ b) := by
intro h <;> cases h <;> simp only [h, iff_true_iff, iff_false_iff]
#align decidable.not_iff Decidable.not_iff
+-/
#print not_iff /-
theorem not_iff : ¬(a ↔ b) ↔ (¬a ↔ b) :=
@@ -1394,6 +1424,7 @@ theorem iff_not_comm : (a ↔ ¬b) ↔ (b ↔ ¬a) :=
#align iff_not_comm iff_not_comm
-/
+#print Decidable.iff_iff_and_or_not_and_not /-
-- See Note [decidable namespace]
protected theorem Decidable.iff_iff_and_or_not_and_not [Decidable b] : (a ↔ b) ↔ a ∧ b ∨ ¬a ∧ ¬b :=
by
@@ -1406,6 +1437,7 @@ protected theorem Decidable.iff_iff_and_or_not_and_not [Decidable b] : (a ↔ b)
| contradiction
| assumption
#align decidable.iff_iff_and_or_not_and_not Decidable.iff_iff_and_or_not_and_not
+-/
#print iff_iff_and_or_not_and_not /-
theorem iff_iff_and_or_not_and_not : (a ↔ b) ↔ a ∧ b ∨ ¬a ∧ ¬b :=
@@ -1428,10 +1460,12 @@ theorem iff_iff_not_or_and_or_not : (a ↔ b) ↔ (¬a ∨ b) ∧ (a ∨ ¬b) :=
#align iff_iff_not_or_and_or_not iff_iff_not_or_and_or_not
-/
+#print Decidable.not_and_not_right /-
-- See Note [decidable namespace]
protected theorem Decidable.not_and_not_right [Decidable b] : ¬(a ∧ ¬b) ↔ a → b :=
⟨fun h ha => h.decidable_imp_symm <| And.intro ha, fun h ⟨ha, hb⟩ => hb <| h ha⟩
#align decidable.not_and_not_right Decidable.not_and_not_right
+-/
#print not_and_not_right /-
theorem not_and_not_right : ¬(a ∧ ¬b) ↔ a → b :=
@@ -1484,10 +1518,12 @@ protected theorem Decidable.not_and [Decidable a] : ¬(a ∧ b) ↔ ¬a ∨ ¬b
#align decidable.not_and_distrib Decidable.not_and
-/
+#print Decidable.not_and' /-
-- See Note [decidable namespace]
protected theorem Decidable.not_and' [Decidable b] : ¬(a ∧ b) ↔ ¬a ∨ ¬b :=
⟨fun h => if hb : b then Or.inl fun ha => h ⟨ha, hb⟩ else Or.inr hb, not_and_of_not_or_not⟩
#align decidable.not_and_distrib' Decidable.not_and'
+-/
#print not_and_or /-
/-- One of de Morgan's laws: the negation of a conjunction is logically equivalent to the
@@ -1578,23 +1614,31 @@ section Mem
variable {α β : Type _} [Membership α β] {s t : β} {a b : α}
+#print ne_of_mem_of_not_mem /-
theorem ne_of_mem_of_not_mem (h : a ∈ s) : b ∉ s → a ≠ b :=
mt fun e => e ▸ h
#align ne_of_mem_of_not_mem ne_of_mem_of_not_mem
+-/
+#print ne_of_mem_of_not_mem' /-
theorem ne_of_mem_of_not_mem' (h : a ∈ s) : a ∉ t → s ≠ t :=
mt fun e => e ▸ h
#align ne_of_mem_of_not_mem' ne_of_mem_of_not_mem'
+-/
+#print Membership.mem.ne_of_not_mem /-
/-- **Alias** of `ne_of_mem_of_not_mem`. -/
theorem Membership.mem.ne_of_not_mem : a ∈ s → b ∉ s → a ≠ b :=
ne_of_mem_of_not_mem
#align has_mem.mem.ne_of_not_mem Membership.mem.ne_of_not_mem
+-/
+#print Membership.mem.ne_of_not_mem' /-
/-- **Alias** of `ne_of_mem_of_not_mem'`. -/
theorem Membership.mem.ne_of_not_mem' : a ∈ s → a ∉ t → s ≠ t :=
ne_of_mem_of_not_mem'
#align has_mem.mem.ne_of_not_mem' Membership.mem.ne_of_not_mem'
+-/
end Mem
@@ -1626,14 +1670,18 @@ theorem ball_cond_comm {α} {s : α → Prop} {p : α → α → Prop} :
-/
/- ./././Mathport/Syntax/Translate/Basic.lean:638:2: warning: expanding binder collection (a b «expr ∈ » s) -/
+#print ball_mem_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
+-/
+#print ne_of_apply_ne /-
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
+-/
#print eq_equivalence /-
theorem eq_equivalence : Equivalence (@Eq α) :=
@@ -1641,11 +1689,13 @@ theorem eq_equivalence : Equivalence (@Eq α) :=
#align eq_equivalence eq_equivalence
-/
+#print eq_rec_constant /-
/-- Transport through trivial families is the identity. -/
@[simp]
theorem eq_rec_constant {α : Sort _} {a a' : α} {β : Sort _} (y : β) (h : a = a') :
@Eq.ndrec α a (fun a => β) y a' h = y := by cases h; rfl
#align eq_rec_constant eq_rec_constant
+-/
#print eq_mp_eq_cast /-
@[simp]
@@ -1669,35 +1719,45 @@ theorem cast_cast :
#align cast_cast cast_cast
-/
+#print congr_refl_left /-
@[simp]
theorem congr_refl_left {α β : Sort _} (f : α → β) {a b : α} (h : a = b) :
congr (Eq.refl f) h = congr_arg f h :=
rfl
#align congr_refl_left congr_refl_left
+-/
+#print congr_refl_right /-
@[simp]
theorem congr_refl_right {α β : Sort _} {f g : α → β} (h : f = g) (a : α) :
congr h (Eq.refl a) = congr_fun h a :=
rfl
#align congr_refl_right congr_refl_right
+-/
+#print congr_arg_refl /-
@[simp]
theorem congr_arg_refl {α β : Sort _} (f : α → β) (a : α) :
congr_arg f (Eq.refl a) = Eq.refl (f a) :=
rfl
#align congr_arg_refl congr_arg_refl
+-/
+#print congr_fun_rfl /-
@[simp]
theorem congr_fun_rfl {α β : Sort _} (f : α → β) (a : α) :
congr_fun (Eq.refl f) a = Eq.refl (f a) :=
rfl
#align congr_fun_rfl congr_fun_rfl
+-/
+#print congr_fun_congr_arg /-
@[simp]
theorem congr_fun_congr_arg {α β γ : Sort _} (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
+-/
#print heq_of_cast_eq /-
theorem heq_of_cast_eq :
@@ -1706,25 +1766,35 @@ theorem heq_of_cast_eq :
#align heq_of_cast_eq heq_of_cast_eq
-/
+#print cast_eq_iff_heq /-
theorem cast_eq_iff_heq {α β : Sort _} {a : α} {a' : β} {e : α = β} : 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
+-/
+#print rec_heq_of_heq /-
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
+-/
+#print rec_heq_iff_heq /-
theorem rec_heq_iff_heq {β} {C : α → Sort _} {x : C a} {y : β} {e : a = b} :
HEq (@Eq.ndrec α a C x b e) y ↔ HEq x y := by subst e
#align rec_heq_iff_heq rec_heq_iff_heq
+-/
+#print heq_rec_iff_heq /-
theorem heq_rec_iff_heq {β} {C : α → Sort _} {x : β} {y : C a} {e : a = b} :
HEq x (@Eq.ndrec α a C y b e) ↔ HEq x y := by subst e
#align heq_rec_iff_heq heq_rec_iff_heq
+-/
+#print Eq.congr /-
protected theorem Eq.congr {x₁ x₂ y₁ y₂ : α} (h₁ : x₁ = y₁) (h₂ : x₂ = y₂) : x₁ = x₂ ↔ y₁ = y₂ := by
subst h₁; subst h₂
#align eq.congr Eq.congr
+-/
#print Eq.congr_left /-
theorem Eq.congr_left {x y z : α} (h : x = y) : x = z ↔ y = z := by rw [h]
@@ -1736,28 +1806,38 @@ theorem Eq.congr_right {x y z : α} (h : x = y) : z = x ↔ z = y := by rw [h]
#align eq.congr_right Eq.congr_right
-/
+#print congr_arg₂ /-
theorem congr_arg₂ {α β γ : Sort _} (f : α → β → γ) {x x' : α} {y y' : β} (hx : x = x')
(hy : y = y') : f x y = f x' y' := by subst hx; subst hy
#align congr_arg2 congr_arg₂
+-/
variable {β : α → Sort _} {γ : ∀ a, β a → Sort _} {δ : ∀ a b, γ a b → Sort _}
+#print congr_fun₂ /-
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₂
+-/
+#print 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₃
+-/
+#print funext₂ /-
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₂
+-/
+#print 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
@@ -1808,26 +1888,34 @@ theorem forall₅_congr {p q : ∀ a b c d, ε a b c d → Prop}
#align forall₅_congr forall₅_congr
-/
+#print exists₂_congr /-
theorem exists₂_congr {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_congr fun a => exists_congr <| h a
#align exists₂_congr exists₂_congr
+-/
+#print exists₃_congr /-
theorem exists₃_congr {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_congr fun a => exists₂_congr <| h a
#align exists₃_congr exists₃_congr
+-/
+#print exists₄_congr /-
theorem exists₄_congr {p q : ∀ a b c, δ a b c → Prop} (h : ∀ a b c d, p a b c d ↔ q a b c d) :
(∃ a b c d, p a b c d) ↔ ∃ a b c d, q a b c d :=
exists_congr fun a => exists₃_congr <| h a
#align exists₄_congr exists₄_congr
+-/
+#print exists₅_congr /-
theorem exists₅_congr {p q : ∀ a b c d, ε a b c d → Prop}
(h : ∀ a b c d e, p a b c d e ↔ q a b c d e) :
(∃ a b c d e, p a b c d e) ↔ ∃ a b c d e, q a b c d e :=
exists_congr fun a => exists₄_congr <| h a
#align exists₅_congr exists₅_congr
+-/
#print forall_imp /-
theorem forall_imp {p q : α → Prop} (h : ∀ a, p a → q a) : (∀ a, p a) → ∀ a, q a := fun h' a =>
@@ -1855,34 +1943,44 @@ theorem Exists.imp {p q : α → Prop} (h : ∀ a, p a → q a) : (∃ a, p a)
#align Exists.imp Exists.imp
-/
+#print Exists₂.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
+-/
+#print 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
variable {ι β : Sort _} {κ : ι → Sort _} {p q : α → Prop} {b : Prop}
+#print Exists.imp' /-
theorem Exists.imp' {p : α → Prop} {q : β → Prop} (f : α → β) (hpq : ∀ a, p a → q (f a))
(hp : ∃ a, p a) : ∃ b, q b :=
Exists.elim hp fun a hp' => ⟨_, hpq _ hp'⟩
#align exists_imp_exists' Exists.imp'
+-/
+#print forall_swap /-
theorem forall_swap {p : α → β → Prop} : (∀ x y, p x y) ↔ ∀ y x, p x y :=
⟨swap, swap⟩
#align forall_swap forall_swap
+-/
+#print forall₂_swap /-
theorem forall₂_swap {ι₁ ι₂ : Sort _} {κ₁ : ι₁ → 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
+-/
#print imp_forall_iff /-
/-- We intentionally restrict the type of `α` in this lemma so that this is a safer to use in simp
@@ -1929,11 +2027,13 @@ theorem Exists.choose_spec {p : α → Prop} (P : ∃ a, p a) : p P.some :=
#align Exists.some_spec Exists.choose_spec
-/
+#print not_exists_of_forall_not /-
--theorem forall_not_of_not_exists (h : ¬ ∃ x, p x) : ∀ x, ¬ p x :=
--forall_imp_of_exists_imp h
theorem not_exists_of_forall_not (h : ∀ x, ¬p x) : ¬∃ x, p x :=
exists_imp.2 h
#align not_exists_of_forall_not not_exists_of_forall_not
+-/
#print not_exists /-
@[simp]
@@ -2016,16 +2116,20 @@ theorem forall_true_iff' (h : ∀ a, p a ↔ True) : (∀ a, p a) ↔ True :=
#align forall_true_iff' forall_true_iff'
-/
+#print forall₂_true_iff /-
@[simp]
theorem forall₂_true_iff {β : α → Sort _} : (∀ a, β a → True) ↔ True :=
forall_true_iff' fun _ => forall_true_iff
#align forall_2_true_iff forall₂_true_iff
+-/
+#print forall₃_true_iff /-
@[simp]
theorem forall₃_true_iff {β : α → Sort _} {γ : ∀ a, β a → Sort _} :
(∀ (a) (b : β a), γ a b → True) ↔ True :=
forall_true_iff' fun _ => forall₂_true_iff
#align forall_3_true_iff forall₃_true_iff
+-/
/- warning: exists_unique.exists clashes with exists_of_exists_unique -> ExistsUnique.exists
Case conversion may be inaccurate. Consider using '#align exists_unique.exists ExistsUnique.existsₓ'. -/
@@ -2083,15 +2187,19 @@ theorem exists_or : (∃ x, p x ∨ q x) ↔ (∃ x, p x) ∨ ∃ x, q x :=
#align exists_or_distrib exists_or
-/
+#print exists_and_left /-
@[simp]
theorem exists_and_left {q : Prop} {p : α → Prop} : (∃ x, q ∧ p x) ↔ q ∧ ∃ x, p x :=
⟨fun ⟨x, hq, hp⟩ => ⟨hq, x, hp⟩, fun ⟨hq, x, hp⟩ => ⟨x, hq, hp⟩⟩
#align exists_and_distrib_left exists_and_left
+-/
+#print exists_and_right /-
@[simp]
theorem exists_and_right {q : Prop} {p : α → Prop} : (∃ x, p x ∧ q) ↔ (∃ x, p x) ∧ q := by
simp [and_comm']
#align exists_and_distrib_right exists_and_right
+-/
#print forall_eq /-
@[simp]
@@ -2107,10 +2215,12 @@ theorem forall_eq' {a' : α} : (∀ a, a' = a → p a) ↔ p a' := by simp [@eq_
-/
/- ./././Mathport/Syntax/Translate/Basic.lean:638:2: warning: expanding binder collection (b «expr ≠ » a) -/
+#print Decidable.and_forall_ne /-
theorem Decidable.and_forall_ne [DecidableEq α] (a : α) :
(p a ∧ ∀ (b) (_ : b ≠ a), p b) ↔ ∀ b, p b := by
simp only [← @forall_eq _ p a, ← forall_and, ← or_imp, Decidable.em, forall_const]
#align decidable.and_forall_ne Decidable.and_forall_ne
+-/
/- ./././Mathport/Syntax/Translate/Basic.lean:638:2: warning: expanding binder collection (b «expr ≠ » a) -/
#print and_forall_ne /-
@@ -2174,15 +2284,19 @@ theorem exists_eq_right {a' : α} : (∃ a, p a ∧ a = a') ↔ p a' :=
#align exists_eq_right exists_eq_right
-/
+#print exists_eq_right_right /-
@[simp]
theorem exists_eq_right_right {a' : α} : (∃ a : α, p a ∧ q a ∧ a = a') ↔ p a' ∧ q a' :=
⟨fun ⟨_, hp, hq, rfl⟩ => ⟨hp, hq⟩, fun ⟨hp, hq⟩ => ⟨a', hp, hq, rfl⟩⟩
#align exists_eq_right_right exists_eq_right_right
+-/
+#print exists_eq_right_right' /-
@[simp]
theorem exists_eq_right_right' {a' : α} : (∃ a : α, p a ∧ q a ∧ a' = a) ↔ p a' ∧ q a' :=
⟨fun ⟨_, hp, hq, rfl⟩ => ⟨hp, hq⟩, fun ⟨hp, hq⟩ => ⟨a', hp, hq, rfl⟩⟩
#align exists_eq_right_right' exists_eq_right_right'
+-/
#print exists_apply_eq_apply /-
@[simp]
@@ -2242,32 +2356,42 @@ theorem exists_or_eq_right' (y : α) (p : α → Prop) : ∃ x : α, p x ∨ y =
#align exists_or_eq_right' exists_or_eq_right'
-/
+#print forall_apply_eq_imp_iff /-
@[simp]
theorem forall_apply_eq_imp_iff {f : α → β} {p : β → Prop} :
(∀ a, ∀ b, f a = b → p b) ↔ ∀ a, p (f a) :=
⟨fun h a => h a (f a) rfl, fun h a b hab => hab ▸ h a⟩
#align forall_apply_eq_imp_iff forall_apply_eq_imp_iff
+-/
+#print 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 rw [forall_swap]; simp
#align forall_apply_eq_imp_iff' forall_apply_eq_imp_iff'
+-/
+#print forall_eq_apply_imp_iff /-
@[simp]
theorem forall_eq_apply_imp_iff {f : α → β} {p : β → Prop} :
(∀ a, ∀ b, b = f a → p b) ↔ ∀ a, p (f a) := by simp [@eq_comm _ _ (f _)]
#align forall_eq_apply_imp_iff forall_eq_apply_imp_iff
+-/
+#print 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 rw [forall_swap]; simp
#align forall_eq_apply_imp_iff' forall_eq_apply_imp_iff'
+-/
+#print forall_apply_eq_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 b a ha hb => hb ▸ h a ha⟩
#align forall_apply_eq_imp_iff₂ forall_apply_eq_imp_iff₂
+-/
#print exists_eq_left' /-
@[simp]
@@ -2287,11 +2411,13 @@ theorem exists_comm {p : α → β → Prop} : (∃ a b, p a b) ↔ ∃ b a, p a
#align exists_comm exists_comm
-/
+#print exists₂_comm /-
theorem exists₂_comm {ι₁ ι₂ : Sort _} {κ₁ : ι₁ → 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
+-/
#print And.exists /-
theorem And.exists {p q : Prop} {f : p ∧ q → Prop} : (∃ h, f h) ↔ ∃ hp hq, f ⟨hp, hq⟩ :=
@@ -2333,10 +2459,12 @@ theorem forall_or_right {q : Prop} {p : α → Prop} : (∀ x, p x ∨ q) ↔ (
#align forall_or_distrib_right forall_or_right
-/
+#print exists_prop /-
@[simp]
theorem exists_prop {p q : Prop} : (∃ h : p, q) ↔ p ∧ q :=
⟨fun ⟨h₁, h₂⟩ => ⟨h₁, h₂⟩, fun ⟨h₁, h₂⟩ => ⟨h₁, h₂⟩⟩
#align exists_prop exists_prop
+-/
#print exists_unique_prop /-
theorem exists_unique_prop {p q : Prop} : (∃! h : p, q) ↔ p ∧ q := by simp
@@ -2460,6 +2588,7 @@ theorem forall_true_left (p : True → Prop) : (∀ x, p x) ↔ p True.intro :=
#align forall_true_left forall_true_left
-/
+#print ExistsUnique.elim₂ /-
theorem ExistsUnique.elim₂ {α : Sort _} {p : α → Sort _} [∀ x, Subsingleton (p x)]
{q : ∀ (x) (h : p x), Prop} {b : Prop} (h₂ : ∃! (x : _) (h : p x), q x h)
(h₁ : ∀ (x) (h : p x), q x h → (∀ (y) (hy : p y), q y hy → y = x) → b) : b :=
@@ -2468,7 +2597,9 @@ theorem ExistsUnique.elim₂ {α : Sort _} {p : α → Sort _} [∀ x, Subsingle
apply h₂.elim
exact fun x ⟨hxp, hxq⟩ H => h₁ x hxp hxq fun y hyp hyq => H y ⟨hyp, hyq⟩
#align exists_unique.elim2 ExistsUnique.elim₂
+-/
+#print ExistsUnique.intro₂ /-
theorem ExistsUnique.intro₂ {α : Sort _} {p : α → Sort _} [∀ x, Subsingleton (p x)]
{q : ∀ (x : α) (h : p x), Prop} (w : α) (hp : p w) (hq : q w hp)
(H : ∀ (y) (hy : p y), q y hy → y = w) : ∃! (x : _) (hx : p x), q x hx :=
@@ -2476,12 +2607,16 @@ theorem ExistsUnique.intro₂ {α : Sort _} {p : α → Sort _} [∀ x, Subsingl
simp only [exists_unique_iff_exists]
exact ExistsUnique.intro w ⟨hp, hq⟩ fun y ⟨hyp, hyq⟩ => H y hyp hyq
#align exists_unique.intro2 ExistsUnique.intro₂
+-/
+#print ExistsUnique.exists₂ /-
theorem ExistsUnique.exists₂ {α : Sort _} {p : α → Sort _} {q : ∀ (x : α) (h : p x), Prop}
(h : ∃! (x : _) (hx : p x), q x hx) : ∃ (x : _) (hx : p x), q x hx :=
h.exists.imp fun x hx => hx.exists
#align exists_unique.exists2 ExistsUnique.exists₂
+-/
+#print ExistsUnique.unique₂ /-
theorem ExistsUnique.unique₂ {α : Sort _} {p : α → Sort _} [∀ x, Subsingleton (p x)]
{q : ∀ (x : α) (hx : p x), Prop} (h : ∃! (x : _) (hx : p x), q x hx) {y₁ y₂ : α} (hpy₁ : p y₁)
(hqy₁ : q y₁ hpy₁) (hpy₂ : p y₂) (hqy₂ : q y₂ hpy₂) : y₁ = y₂ :=
@@ -2489,6 +2624,7 @@ theorem ExistsUnique.unique₂ {α : Sort _} {p : α → Sort _} [∀ x, Subsing
simp only [exists_unique_iff_exists] at h
exact h.unique ⟨hpy₁, hqy₁⟩ ⟨hpy₂, hqy₂⟩
#align exists_unique.unique2 ExistsUnique.unique₂
+-/
end Quantifiers
@@ -2529,12 +2665,14 @@ noncomputable def decEq (α : Sort _) : DecidableEq α := by infer_instance
#align classical.dec_eq Classical.decEq
-/
+#print Classical.existsCases /-
/-- Construct a function from a default value `H0`, and a function to use if there exists a value
satisfying the predicate. -/
@[elab_as_elim]
noncomputable def existsCases.{u} {C : Sort u} (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
+-/
#print Classical.some_spec₂ /-
theorem some_spec₂ {α : Sort _} {p : α → Prop} {h : ∃ a, p a} (q : α → Prop)
@@ -2566,6 +2704,7 @@ def choice_of_byContradiction' {α : Sort _} (contra : ¬(α → False) → α)
end Classical
+#print Exists.classicalRecOn /-
/-- This function has the same type as `exists.rec_on`, and can be used to case on an equality,
but `exists.rec_on` can only eliminate into Prop, while this version eliminates into any universe
using the axiom of choice. -/
@@ -2574,6 +2713,7 @@ noncomputable def Exists.classicalRecOn.{u} {α} {p : α → Prop} (h : ∃ a, p
(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 -/
@@ -2899,16 +3039,21 @@ theorem ite_eq_or_eq : ite P a b = a ∨ ite P a b = b :=
#align ite_eq_or_eq ite_eq_or_eq
-/
+#print apply_dite /-
/-- A function applied to a `dite` is a `dite` of that function applied to each of the branches. -/
theorem apply_dite (x : P → α) (y : ¬P → α) :
f (dite P x y) = dite P (fun h => f (x h)) fun h => f (y h) := by by_cases h : P <;> simp [h]
#align apply_dite apply_dite
+-/
+#print apply_ite /-
/-- A function applied to a `ite` is a `ite` of that function applied to each of the branches. -/
theorem apply_ite : f (ite P a b) = ite P (f a) (f b) :=
apply_dite f P (fun _ => a) fun _ => b
#align apply_ite apply_ite
+-/
+#print apply_dite₂ /-
/-- A two-argument function applied to two `dite`s is a `dite` of that two-argument function
applied to each of the branches. -/
theorem apply_dite₂ (f : α → β → γ) (P : Prop) [Decidable P] (a : P → α) (b : ¬P → α) (c : P → β)
@@ -2916,13 +3061,16 @@ theorem apply_dite₂ (f : α → β → γ) (P : Prop) [Decidable P] (a : P →
f (dite P a b) (dite P c d) = dite P (fun h => f (a h) (c h)) fun h => f (b h) (d h) := by
by_cases h : P <;> simp [h]
#align apply_dite2 apply_dite₂
+-/
+#print apply_ite₂ /-
/-- A two-argument function applied to two `ite`s is a `ite` of that two-argument function
applied to each of the branches. -/
theorem apply_ite₂ (f : α → β → γ) (P : Prop) [Decidable P] (a b : α) (c d : β) :
f (ite P a b) (ite P c d) = ite P (f a c) (f b d) :=
apply_dite₂ f P (fun _ => a) (fun _ => b) (fun _ => c) fun _ => d
#align apply_ite2 apply_ite₂
+-/
#print dite_apply /-
/-- A 'dite' producing a `Pi` type `Π a, σ a`, applied to a value `a : α` is a `dite` that applies
mathlib commit https://github.com/leanprover-community/mathlib/commit/31c24aa72e7b3e5ed97a8412470e904f82b81004
@@ -1625,7 +1625,7 @@ theorem ball_cond_comm {α} {s : α → Prop} {p : α → α → Prop} :
#align ball_cond_comm ball_cond_comm
-/
-/- ./././Mathport/Syntax/Translate/Basic.lean:635:2: warning: expanding binder collection (a b «expr ∈ » s) -/
+/- ./././Mathport/Syntax/Translate/Basic.lean:638:2: warning: expanding binder collection (a b «expr ∈ » s) -/
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
@@ -2106,13 +2106,13 @@ theorem forall_eq' {a' : α} : (∀ a, a' = a → p a) ↔ p a' := by simp [@eq_
#align forall_eq' forall_eq'
-/
-/- ./././Mathport/Syntax/Translate/Basic.lean:635:2: warning: expanding binder collection (b «expr ≠ » a) -/
+/- ./././Mathport/Syntax/Translate/Basic.lean:638:2: warning: expanding binder collection (b «expr ≠ » a) -/
theorem Decidable.and_forall_ne [DecidableEq α] (a : α) :
(p a ∧ ∀ (b) (_ : b ≠ a), p b) ↔ ∀ b, p b := by
simp only [← @forall_eq _ p a, ← forall_and, ← or_imp, Decidable.em, forall_const]
#align decidable.and_forall_ne Decidable.and_forall_ne
-/- ./././Mathport/Syntax/Translate/Basic.lean:635:2: warning: expanding binder collection (b «expr ≠ » a) -/
+/- ./././Mathport/Syntax/Translate/Basic.lean:638:2: warning: expanding binder collection (b «expr ≠ » a) -/
#print and_forall_ne /-
theorem and_forall_ne (a : α) : (p a ∧ ∀ (b) (_ : b ≠ a), p b) ↔ ∀ b, p b :=
Decidable.and_forall_ne a
mathlib commit https://github.com/leanprover-community/mathlib/commit/5f25c089cb34db4db112556f23c50d12da81b297
@@ -271,7 +271,7 @@ theorem eq_iff_eq_cancel_right {α : Sort _} {a b : α} : (∀ {c}, a = c ↔ b
-/
#print Fact /-
-/- ./././Mathport/Syntax/Translate/Command.lean:393:30: infer kinds are unsupported in Lean 4: #[`out] [] -/
+/- ./././Mathport/Syntax/Translate/Command.lean:394:30: infer kinds are unsupported in Lean 4: #[`out] [] -/
/-- Wrapper for adding elementary propositions to the type class systems.
Warning: this can easily be abused. See the rest of this docstring for details.
mathlib commit https://github.com/leanprover-community/mathlib/commit/cca40788df1b8755d5baf17ab2f27dacc2e17acb
@@ -37,8 +37,8 @@ section Miscellany
/- We add the `inline` attribute to optimize VM computation using these declarations. For example,
`if p ∧ q then ... else ...` will not evaluate the decidability of `q` if `p` is false. -/
-attribute [inline]
- And.decidable Or.decidable decidableFalse Xor'.decidable Iff.decidable decidableTrue Implies.decidable Not.decidable Ne.decidable Bool.decidableEq Decidable.decide
+attribute [inline] And.decidable Or.decidable decidableFalse Xor'.decidable Iff.decidable
+ decidableTrue Implies.decidable Not.decidable Ne.decidable Bool.decidableEq Decidable.decide
attribute [simp] cast_eq cast_hEq
@@ -66,7 +66,7 @@ instance : Subsingleton Empty :=
instance Subsingleton.prod {α β : Type _} [Subsingleton α] [Subsingleton β] :
Subsingleton (α × β) :=
- ⟨by intro a b; cases a; cases b; congr ⟩
+ ⟨by intro a b; cases a; cases b; congr⟩
#align subsingleton.prod Subsingleton.prod
instance : DecidableEq Empty := fun a => a.elim
@@ -704,7 +704,7 @@ theorem imp_not_self : a → ¬a ↔ ¬a :=
#print Decidable.not_imp_self /-
theorem Decidable.not_imp_self [Decidable a] : ¬a → a ↔ a := by have := @imp_not_self ¬a;
- rwa [Decidable.not_not] at this
+ rwa [Decidable.not_not] at this
#align decidable.not_imp_self Decidable.not_imp_self
-/
@@ -938,7 +938,7 @@ theorem iff_and_self {p q : Prop} : (p ↔ q ∧ p) ↔ p → q := by rw [and_co
#print and_congr_right_iff /-
@[simp]
theorem and_congr_right_iff : (a ∧ b ↔ a ∧ c) ↔ a → (b ↔ c) :=
- ⟨fun h ha => by simp [ha] at h <;> exact h, and_congr_right⟩
+ ⟨fun h ha => by simp [ha] at h <;> exact h, and_congr_right⟩
#align and.congr_right_iff and_congr_right_iff
-/
@@ -1398,8 +1398,13 @@ theorem iff_not_comm : (a ↔ ¬b) ↔ (b ↔ ¬a) :=
protected theorem Decidable.iff_iff_and_or_not_and_not [Decidable b] : (a ↔ b) ↔ a ∧ b ∨ ¬a ∧ ¬b :=
by
constructor <;> intro h
- · rw [h] <;> by_cases b <;> [left;right] <;> constructor <;> assumption
- · cases' h with h h <;> cases h <;> constructor <;> intro <;> · first |contradiction|assumption
+ · rw [h] <;> by_cases b <;> [left; right] <;> constructor <;> assumption
+ ·
+ cases' h with h h <;> cases h <;> constructor <;> intro <;>
+ ·
+ first
+ | contradiction
+ | assumption
#align decidable.iff_iff_and_or_not_and_not Decidable.iff_iff_and_or_not_and_not
#print iff_iff_and_or_not_and_not /-
@@ -2456,32 +2461,32 @@ theorem forall_true_left (p : True → Prop) : (∀ x, p x) ↔ p True.intro :=
-/
theorem ExistsUnique.elim₂ {α : Sort _} {p : α → Sort _} [∀ x, Subsingleton (p x)]
- {q : ∀ (x) (h : p x), Prop} {b : Prop} (h₂ : ∃! (x : _)(h : p x), q x h)
+ {q : ∀ (x) (h : p x), Prop} {b : Prop} (h₂ : ∃! (x : _) (h : p x), q x h)
(h₁ : ∀ (x) (h : p x), q x h → (∀ (y) (hy : p y), q y hy → y = x) → b) : b :=
by
- simp only [exists_unique_iff_exists] at h₂
+ simp only [exists_unique_iff_exists] at h₂
apply h₂.elim
exact fun x ⟨hxp, hxq⟩ H => h₁ x hxp hxq fun y hyp hyq => H y ⟨hyp, hyq⟩
#align exists_unique.elim2 ExistsUnique.elim₂
theorem ExistsUnique.intro₂ {α : Sort _} {p : α → Sort _} [∀ x, Subsingleton (p x)]
{q : ∀ (x : α) (h : p x), Prop} (w : α) (hp : p w) (hq : q w hp)
- (H : ∀ (y) (hy : p y), q y hy → y = w) : ∃! (x : _)(hx : p x), q x hx :=
+ (H : ∀ (y) (hy : p y), q y hy → y = w) : ∃! (x : _) (hx : p x), q x hx :=
by
simp only [exists_unique_iff_exists]
exact ExistsUnique.intro w ⟨hp, hq⟩ fun y ⟨hyp, hyq⟩ => H y hyp hyq
#align exists_unique.intro2 ExistsUnique.intro₂
theorem ExistsUnique.exists₂ {α : Sort _} {p : α → Sort _} {q : ∀ (x : α) (h : p x), Prop}
- (h : ∃! (x : _)(hx : p x), q x hx) : ∃ (x : _)(hx : p x), q x hx :=
+ (h : ∃! (x : _) (hx : p x), q x hx) : ∃ (x : _) (hx : p x), q x hx :=
h.exists.imp fun x hx => hx.exists
#align exists_unique.exists2 ExistsUnique.exists₂
theorem ExistsUnique.unique₂ {α : Sort _} {p : α → Sort _} [∀ x, Subsingleton (p x)]
- {q : ∀ (x : α) (hx : p x), Prop} (h : ∃! (x : _)(hx : p x), q x hx) {y₁ y₂ : α} (hpy₁ : p y₁)
+ {q : ∀ (x : α) (hx : p x), Prop} (h : ∃! (x : _) (hx : p x), q x hx) {y₁ y₂ : α} (hpy₁ : p y₁)
(hqy₁ : q y₁ hpy₁) (hpy₂ : p y₂) (hqy₂ : q y₂ hpy₂) : y₁ = y₂ :=
by
- simp only [exists_unique_iff_exists] at h
+ simp only [exists_unique_iff_exists] at h
exact h.unique ⟨hpy₁, hqy₁⟩ ⟨hpy₂, hqy₂⟩
#align exists_unique.unique2 ExistsUnique.unique₂
@@ -2578,7 +2583,7 @@ section BoundedQuantifiers
variable {α : Sort _} {r p q : α → Prop} {P Q : ∀ x, p x → Prop} {b : Prop}
#print bex_def /-
-theorem bex_def : (∃ (x : _)(h : p x), q x) ↔ ∃ x, p x ∧ q x :=
+theorem bex_def : (∃ (x : _) (h : 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
-/
@@ -2590,7 +2595,7 @@ theorem BEx.elim {b : Prop} : (∃ x h, P x h) → (∀ a h, P a h → b) → b
-/
#print BEx.intro /-
-theorem BEx.intro (a : α) (h₁ : p a) (h₂ : P a h₁) : ∃ (x : _)(h : p x), P x h :=
+theorem BEx.intro (a : α) (h₁ : p a) (h₂ : P a h₁) : ∃ (x : _) (h : p x), P x h :=
⟨a, h₁, h₂⟩
#align bex.intro BEx.intro
-/
@@ -2608,7 +2613,7 @@ theorem bex_congr (H : ∀ x h, P x h ↔ Q x h) : (∃ x h, P x h) ↔ ∃ x h,
-/
#print bex_eq_left /-
-theorem bex_eq_left {a : α} : (∃ (x : _)(_ : x = a), p x) ↔ p a := by
+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
-/
@@ -2632,7 +2637,7 @@ theorem BAll.imp_left (H : ∀ x, p x → q x) (h₁ : ∀ x, q x → r x) (x) (
-/
#print BEx.imp_left /-
-theorem BEx.imp_left (H : ∀ x, p x → q x) : (∃ (x : _)(_ : p x), r x) → ∃ (x : _)(_ : q x), r x
+theorem BEx.imp_left (H : ∀ x, p x → q x) : (∃ (x : _) (_ : p x), r x) → ∃ (x : _) (_ : q x), r x
| ⟨x, hp, hr⟩ => ⟨x, H _ hp, hr⟩
#align bex.imp_left BEx.imp_left
-/
@@ -2650,13 +2655,13 @@ theorem forall_of_ball (H : ∀ x, p x) (h : ∀ x, p x → q x) (x) : q x :=
-/
#print bex_of_exists /-
-theorem bex_of_exists (H : ∀ x, p x) : (∃ x, q x) → ∃ (x : _)(_ : p x), q x
+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
-/
#print exists_of_bex /-
-theorem exists_of_bex : (∃ (x : _)(_ : p x), q x) → ∃ x, q x
+theorem exists_of_bex : (∃ (x : _) (_ : p x), q x) → ∃ x, q x
| ⟨x, _, hq⟩ => ⟨x, hq⟩
#align exists_of_bex exists_of_bex
-/
@@ -2720,7 +2725,7 @@ theorem ball_or_left : (∀ x, p x ∨ q x → r x) ↔ (∀ x, p x → r x) ∧
#print bex_or_left /-
theorem bex_or_left :
- (∃ (x : _)(_ : p x ∨ q x), r x) ↔ (∃ (x : _)(_ : p x), r x) ∨ ∃ (x : _)(_ : q x), r x := by
+ (∃ (x : _) (_ : p x ∨ q x), r x) ↔ (∃ (x : _) (_ : p x), r x) ∨ ∃ (x : _) (_ : q x), r x := by
simp only [exists_prop] <;> exact Iff.trans (exists_congr fun x => or_and_right) exists_or
#align bex_or_left_distrib bex_or_left
-/
@@ -2962,7 +2967,7 @@ 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) =
if q : Q then B q else if p : P then A p else C 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]⟩
+ ⟨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
-/
mathlib commit https://github.com/leanprover-community/mathlib/commit/917c3c072e487b3cccdbfeff17e75b40e45f66cb
@@ -216,43 +216,19 @@ theorem not_nonempty_pempty : ¬Nonempty PEmpty := fun ⟨h⟩ => h.elim
#align not_nonempty_pempty not_nonempty_pempty
-/
-/- warning: congr_heq -> congr_heq is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {β : Sort.{u1}} {γ : Sort.{u2}} {f : α -> γ} {g : β -> γ} {x : α} {y : β}, (HEq.{imax u1 u2} (α -> γ) f (β -> γ) g) -> (HEq.{u1} α x β y) -> (Eq.{u2} γ (f x) (g y))
-but is expected to have type
- forall {α : Sort.{u2}} {β : Sort.{u2}} {γ : Sort.{u1}} {f : α -> γ} {g : β -> γ} {x : α} {y : β}, (HEq.{imax u2 u1} (α -> γ) f (β -> γ) g) -> (HEq.{u2} α x β y) -> (Eq.{u1} γ (f x) (g y))
-Case conversion may be inaccurate. Consider using '#align congr_heq congr_heqₓ'. -/
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
-/- warning: congr_arg_heq -> congr_arg_heq is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {β : α -> Sort.{u2}} (f : forall (a : α), β a) {a₁ : α} {a₂ : α}, (Eq.{u1} α a₁ a₂) -> (HEq.{u2} (β a₁) (f a₁) (β a₂) (f a₂))
-but is expected to have type
- forall {α : Sort.{u2}} {β : α -> Sort.{u1}} (f : forall (a : α), β a) {a₁ : α} {a₂ : α}, (Eq.{u2} α a₁ a₂) -> (HEq.{u1} (β a₁) (f a₁) (β a₂) (f a₂))
-Case conversion may be inaccurate. Consider using '#align congr_arg_heq congr_arg_heqₓ'. -/
theorem congr_arg_heq {α} {β : α → Sort _} (f : ∀ a, β a) :
∀ {a₁ a₂ : α}, a₁ = a₂ → HEq (f a₁) (f a₂)
| a, _, rfl => HEq.rfl
#align congr_arg_heq congr_arg_heq
-/- warning: ulift.down_injective -> ULift.down_injective is a dubious translation:
-lean 3 declaration is
- forall {α : Type.{u1}}, Function.Injective.{succ (max u1 u2), succ u1} (ULift.{u2, u1} α) α (ULift.down.{u2, u1} α)
-but is expected to have type
- forall {α : Type.{u2}}, Function.Injective.{max (succ u2) (succ u1), succ u2} (ULift.{u1, u2} α) α (ULift.down.{u1, u2} α)
-Case conversion may be inaccurate. Consider using '#align ulift.down_injective ULift.down_injectiveₓ'. -/
theorem ULift.down_injective {α : Sort _} : Function.Injective (@ULift.down α)
| ⟨a⟩, ⟨b⟩, rfl => rfl
#align ulift.down_injective ULift.down_injective
-/- warning: ulift.down_inj -> ULift.down_inj is a dubious translation:
-lean 3 declaration is
- forall {α : Type.{u1}} {a : ULift.{u2, u1} α} {b : ULift.{u2, u1} α}, Iff (Eq.{succ u1} α (ULift.down.{u2, u1} α a) (ULift.down.{u2, u1} α b)) (Eq.{succ (max u1 u2)} (ULift.{u2, u1} α) a b)
-but is expected to have type
- forall {α : Type.{u2}} {a : ULift.{u1, u2} α} {b : ULift.{u1, u2} α}, Iff (Eq.{succ u2} α (ULift.down.{u1, u2} α a) (ULift.down.{u1, u2} α b)) (Eq.{max (succ u2) (succ u1)} (ULift.{u1, u2} α) a b)
-Case conversion may be inaccurate. Consider using '#align ulift.down_inj ULift.down_injₓ'. -/
@[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]⟩
@@ -739,12 +715,6 @@ theorem not_imp_self : ¬a → a ↔ a :=
#align not_imp_self not_imp_self
-/
-/- warning: imp.swap -> Imp.swap is a dubious translation:
-lean 3 declaration is
- forall {a : Prop} {b : Prop} {c : Prop}, Iff (a -> b -> c) (b -> a -> c)
-but is expected to have type
- forall {a : Sort.{u_1}} {b : Sort.{u_2}} {c : Prop}, Iff (a -> b -> c) (b -> a -> c)
-Case conversion may be inaccurate. Consider using '#align imp.swap Imp.swapₓ'. -/
theorem Imp.swap : a → b → c ↔ b → a → c :=
⟨swap, swap⟩
#align imp.swap Imp.swap
@@ -790,12 +760,6 @@ theorem xor_false : Xor' False = id :=
#align xor_false xor_false
-/
-/- warning: xor_comm -> xor_comm is a dubious translation:
-lean 3 declaration is
- forall (a : Prop) (b : Prop), Iff (Xor' a b) (Xor' b a)
-but is expected to have type
- forall (a : Prop) (b : Prop), Eq.{1} Prop (Xor' a b) (Xor' b a)
-Case conversion may be inaccurate. Consider using '#align xor_comm xor_commₓ'. -/
theorem xor_comm (a b) : Xor' a b ↔ Xor' b a :=
or_comm' _ _
#align xor_comm xor_comm
@@ -935,12 +899,6 @@ theorem and_iff_left_of_imp {a b : Prop} (h : a → b) : a ∧ b ↔ a :=
#align and_iff_left_of_imp and_iff_left_of_imp
-/
-/- warning: and_iff_right_of_imp -> and_iff_right_of_imp is a dubious translation:
-lean 3 declaration is
- forall {a : Prop} {b : Prop}, (b -> a) -> (Iff (And a b) b)
-but is expected to have type
- forall {a : Prop} {b : Prop}, (a -> b) -> (Iff (And b a) a)
-Case conversion may be inaccurate. Consider using '#align and_iff_right_of_imp and_iff_right_of_impₓ'. -/
theorem and_iff_right_of_imp {a b : Prop} (h : b → a) : a ∧ b ↔ b :=
Iff.intro And.right fun hb => ⟨h hb, hb⟩
#align and_iff_right_of_imp and_iff_right_of_imp
@@ -1062,22 +1020,10 @@ theorem or_of_or_of_imp_of_imp (h₁ : a ∨ b) (h₂ : a → c) (h₃ : b → d
#align or_of_or_of_imp_of_imp or_of_or_of_imp_of_imp
-/
-/- warning: or_of_or_of_imp_left -> or_of_or_of_imp_left is a dubious translation:
-lean 3 declaration is
- forall {a : Prop} {b : Prop} {c : Prop}, (Or a c) -> (a -> b) -> (Or b c)
-but is expected to have type
- forall {a : Prop} {b : Prop} {c : Prop}, (Or a b) -> (a -> c) -> (Or c b)
-Case conversion may be inaccurate. Consider using '#align or_of_or_of_imp_left or_of_or_of_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
-/- warning: or_of_or_of_imp_right -> or_of_or_of_imp_right is a dubious translation:
-lean 3 declaration is
- forall {a : Prop} {b : Prop} {c : Prop}, (Or c a) -> (a -> b) -> (Or c b)
-but is expected to have type
- forall {a : Prop} {b : Prop} {c : Prop}, (Or a b) -> (b -> c) -> (Or a c)
-Case conversion may be inaccurate. Consider using '#align or_of_or_of_imp_right or_of_or_of_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
@@ -1088,12 +1034,6 @@ theorem Or.elim3 (h : a ∨ b ∨ c) (ha : a → d) (hb : b → d) (hc : c → d
#align or.elim3 Or.elim3
-/
-/- warning: or.imp3 -> Or.imp3 is a dubious translation:
-lean 3 declaration is
- forall {a : Prop} {b : Prop} {c : Prop} {d : Prop} {e : Prop} {f : Prop}, (a -> d) -> (b -> e) -> (c -> f) -> (Or a (Or b c)) -> (Or d (Or e f))
-but is expected to have type
- forall {a : Prop} {b : Prop} {c : Prop} {d : Prop} {e : Prop} {f : Prop}, (a -> b) -> (c -> d) -> (e -> f) -> (Or a (Or c e)) -> (Or b (Or d f))
-Case conversion may be inaccurate. Consider using '#align or.imp3 Or.imp3ₓ'. -/
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
@@ -1172,12 +1112,6 @@ protected theorem Decidable.imp_iff_or_not [Decidable b] : b → a ↔ a ∨ ¬b
Decidable.imp_iff_not_or.trans or_comm
#align decidable.imp_iff_or_not Decidable.imp_iff_or_notₓ
-/- warning: imp_iff_or_not -> imp_iff_or_not is a dubious translation:
-lean 3 declaration is
- forall {a : Prop} {b : Prop}, Iff (b -> a) (Or a (Not b))
-but is expected to have type
- forall {a : Prop} {b : Prop}, Iff (a -> b) (Or b (Not a))
-Case conversion may be inaccurate. Consider using '#align imp_iff_or_not imp_iff_or_notₓ'. -/
theorem imp_iff_or_not : b → a ↔ a ∨ ¬b :=
Decidable.imp_iff_or_not
#align imp_iff_or_not imp_iff_or_not
@@ -1202,23 +1136,11 @@ protected theorem Function.mtr : (¬a → ¬b) → b → a :=
#align function.mtr Function.mtr
-/
-/- warning: decidable.or_congr_left -> Decidable.or_congr_left' is a dubious translation:
-lean 3 declaration is
- forall {a : Prop} {b : Prop} {c : Prop} [_inst_1 : Decidable c], ((Not c) -> (Iff a b)) -> (Iff (Or a c) (Or b c))
-but is expected to have type
- forall {a : Prop} {b : Prop} {c : Prop} [_inst_1 : Decidable a], ((Not a) -> (Iff b c)) -> (Iff (Or b a) (Or c a))
-Case conversion may be inaccurate. Consider using '#align decidable.or_congr_left Decidable.or_congr_left'ₓ'. -/
-- See Note [decidable namespace]
protected theorem Decidable.or_congr_left' [Decidable c] (h : ¬c → (a ↔ b)) : a ∨ c ↔ b ∨ c := by
rw [Decidable.or_iff_not_imp_right, Decidable.or_iff_not_imp_right]; exact imp_congr_right h
#align decidable.or_congr_left Decidable.or_congr_left'
-/- warning: or_congr_left -> or_congr_left' is a dubious translation:
-lean 3 declaration is
- forall {a : Prop} {b : Prop} {c : Prop}, ((Not c) -> (Iff a b)) -> (Iff (Or a c) (Or b c))
-but is expected to have type
- forall {a : Prop} {b : Prop} {c : Prop}, ((Not a) -> (Iff b c)) -> (Iff (Or b a) (Or c a))
-Case conversion may be inaccurate. Consider using '#align or_congr_left or_congr_left'ₓ'. -/
theorem or_congr_left' (h : ¬c → (a ↔ b)) : a ∨ c ↔ b ∨ c :=
Decidable.or_congr_left' h
#align or_congr_left or_congr_left'
@@ -1369,12 +1291,6 @@ theorem imp_or : a → b ∨ c ↔ (a → b) ∨ (a → c) :=
#align imp_or_distrib imp_or
-/
-/- warning: decidable.imp_or_distrib' -> Decidable.imp_or' is a dubious translation:
-lean 3 declaration is
- forall {a : Prop} {b : Prop} {c : Prop} [_inst_1 : Decidable b], Iff (a -> (Or b c)) (Or (a -> b) (a -> c))
-but is expected to have type
- forall {a : Prop} {b : Sort.{u_1}} {c : Prop} [_inst_1 : Decidable a], Iff (b -> (Or a c)) (Or (b -> a) (b -> c))
-Case conversion may be inaccurate. Consider using '#align decidable.imp_or_distrib' Decidable.imp_or'ₓ'. -/
-- See Note [decidable namespace]
protected theorem Decidable.imp_or' [Decidable b] : a → b ∨ c ↔ (a → b) ∨ (a → c) := by
by_cases b <;> simp [h, or_iff_right_of_imp ((· ∘ ·) False.elim)]
@@ -1454,12 +1370,6 @@ theorem not_iff_comm : (¬a ↔ b) ↔ (¬b ↔ a) :=
#align not_iff_comm not_iff_comm
-/
-/- warning: decidable.not_iff -> Decidable.not_iff is a dubious translation:
-lean 3 declaration is
- forall {a : Prop} {b : Prop} [_inst_1 : Decidable b], Iff (Not (Iff a b)) (Iff (Not a) b)
-but is expected to have type
- forall {a : Prop} {b : Prop} [_inst_1 : Decidable a], Iff (Not (Iff b a)) (Iff (Not b) a)
-Case conversion may be inaccurate. Consider using '#align decidable.not_iff Decidable.not_iffₓ'. -/
-- See Note [decidable namespace]
protected theorem Decidable.not_iff : ∀ [Decidable b], ¬(a ↔ b) ↔ (¬a ↔ b) := by
intro h <;> cases h <;> simp only [h, iff_true_iff, iff_false_iff]
@@ -1484,12 +1394,6 @@ theorem iff_not_comm : (a ↔ ¬b) ↔ (b ↔ ¬a) :=
#align iff_not_comm iff_not_comm
-/
-/- warning: decidable.iff_iff_and_or_not_and_not -> Decidable.iff_iff_and_or_not_and_not is a dubious translation:
-lean 3 declaration is
- forall {a : Prop} {b : Prop} [_inst_1 : Decidable b], Iff (Iff a b) (Or (And a b) (And (Not a) (Not b)))
-but is expected to have type
- forall {a : Prop} {b : Prop} [_inst_1 : Decidable a], Iff (Iff b a) (Or (And b a) (And (Not b) (Not a)))
-Case conversion may be inaccurate. Consider using '#align decidable.iff_iff_and_or_not_and_not Decidable.iff_iff_and_or_not_and_notₓ'. -/
-- See Note [decidable namespace]
protected theorem Decidable.iff_iff_and_or_not_and_not [Decidable b] : (a ↔ b) ↔ a ∧ b ∨ ¬a ∧ ¬b :=
by
@@ -1519,12 +1423,6 @@ theorem iff_iff_not_or_and_or_not : (a ↔ b) ↔ (¬a ∨ b) ∧ (a ∨ ¬b) :=
#align iff_iff_not_or_and_or_not iff_iff_not_or_and_or_not
-/
-/- warning: decidable.not_and_not_right -> Decidable.not_and_not_right is a dubious translation:
-lean 3 declaration is
- forall {a : Prop} {b : Prop} [_inst_1 : Decidable b], Iff (Not (And a (Not b))) (a -> b)
-but is expected to have type
- forall {a : Prop} {b : Prop} [_inst_1 : Decidable a], Iff (Not (And b (Not a))) (b -> a)
-Case conversion may be inaccurate. Consider using '#align decidable.not_and_not_right Decidable.not_and_not_rightₓ'. -/
-- See Note [decidable namespace]
protected theorem Decidable.not_and_not_right [Decidable b] : ¬(a ∧ ¬b) ↔ a → b :=
⟨fun h ha => h.decidable_imp_symm <| And.intro ha, fun h ⟨ha, hb⟩ => hb <| h ha⟩
@@ -1581,12 +1479,6 @@ protected theorem Decidable.not_and [Decidable a] : ¬(a ∧ b) ↔ ¬a ∨ ¬b
#align decidable.not_and_distrib Decidable.not_and
-/
-/- warning: decidable.not_and_distrib' -> Decidable.not_and' is a dubious translation:
-lean 3 declaration is
- forall {a : Prop} {b : Prop} [_inst_1 : Decidable b], Iff (Not (And a b)) (Or (Not a) (Not b))
-but is expected to have type
- forall {a : Prop} {b : Prop} [_inst_1 : Decidable a], Iff (Not (And b a)) (Or (Not b) (Not a))
-Case conversion may be inaccurate. Consider using '#align decidable.not_and_distrib' Decidable.not_and'ₓ'. -/
-- See Note [decidable namespace]
protected theorem Decidable.not_and' [Decidable b] : ¬(a ∧ b) ↔ ¬a ∨ ¬b :=
⟨fun h => if hb : b then Or.inl fun ha => h ⟨ha, hb⟩ else Or.inr hb, not_and_of_not_or_not⟩
@@ -1681,43 +1573,19 @@ section Mem
variable {α β : Type _} [Membership α β] {s t : β} {a b : α}
-/- warning: ne_of_mem_of_not_mem -> ne_of_mem_of_not_mem is a dubious translation:
-lean 3 declaration is
- forall {α : Type.{u1}} {β : Type.{u2}} [_inst_1 : Membership.{u1, u2} α β] {s : β} {a : α} {b : α}, (Membership.Mem.{u1, u2} α β _inst_1 a s) -> (Not (Membership.Mem.{u1, u2} α β _inst_1 b s)) -> (Ne.{succ u1} α a b)
-but is expected to have type
- forall {α : Type.{u2}} {β : Type.{u1}} [_inst_1 : Membership.{u2, u1} α β] {s : β} {a : α} {b : α}, (Membership.mem.{u2, u1} α β _inst_1 a s) -> (Not (Membership.mem.{u2, u1} α β _inst_1 b s)) -> (Ne.{succ u2} α a b)
-Case conversion may be inaccurate. Consider using '#align ne_of_mem_of_not_mem ne_of_mem_of_not_memₓ'. -/
theorem ne_of_mem_of_not_mem (h : a ∈ s) : b ∉ s → a ≠ b :=
mt fun e => e ▸ h
#align ne_of_mem_of_not_mem ne_of_mem_of_not_mem
-/- warning: ne_of_mem_of_not_mem' -> ne_of_mem_of_not_mem' is a dubious translation:
-lean 3 declaration is
- forall {α : Type.{u1}} {β : Type.{u2}} [_inst_1 : Membership.{u1, u2} α β] {s : β} {t : β} {a : α}, (Membership.Mem.{u1, u2} α β _inst_1 a s) -> (Not (Membership.Mem.{u1, u2} α β _inst_1 a t)) -> (Ne.{succ u2} β s t)
-but is expected to have type
- forall {α : Type.{u2}} {β : Type.{u1}} [_inst_1 : Membership.{u2, u1} α β] {s : β} {t : β} {a : α}, (Membership.mem.{u2, u1} α β _inst_1 a s) -> (Not (Membership.mem.{u2, u1} α β _inst_1 a t)) -> (Ne.{succ u1} β s t)
-Case conversion may be inaccurate. Consider using '#align ne_of_mem_of_not_mem' ne_of_mem_of_not_mem'ₓ'. -/
theorem ne_of_mem_of_not_mem' (h : a ∈ s) : a ∉ t → s ≠ t :=
mt fun e => e ▸ h
#align ne_of_mem_of_not_mem' ne_of_mem_of_not_mem'
-/- warning: has_mem.mem.ne_of_not_mem -> Membership.mem.ne_of_not_mem is a dubious translation:
-lean 3 declaration is
- forall {α : Type.{u1}} {β : Type.{u2}} [_inst_1 : Membership.{u1, u2} α β] {s : β} {a : α} {b : α}, (Membership.Mem.{u1, u2} α β _inst_1 a s) -> (Not (Membership.Mem.{u1, u2} α β _inst_1 b s)) -> (Ne.{succ u1} α a b)
-but is expected to have type
- forall {α : Type.{u2}} {β : Type.{u1}} [_inst_1 : Membership.{u2, u1} α β] {s : β} {a : α} {b : α}, (Membership.mem.{u2, u1} α β _inst_1 a s) -> (Not (Membership.mem.{u2, u1} α β _inst_1 b s)) -> (Ne.{succ u2} α a b)
-Case conversion may be inaccurate. Consider using '#align has_mem.mem.ne_of_not_mem Membership.mem.ne_of_not_memₓ'. -/
/-- **Alias** of `ne_of_mem_of_not_mem`. -/
theorem Membership.mem.ne_of_not_mem : a ∈ s → b ∉ s → a ≠ b :=
ne_of_mem_of_not_mem
#align has_mem.mem.ne_of_not_mem Membership.mem.ne_of_not_mem
-/- warning: has_mem.mem.ne_of_not_mem' -> Membership.mem.ne_of_not_mem' is a dubious translation:
-lean 3 declaration is
- forall {α : Type.{u1}} {β : Type.{u2}} [_inst_1 : Membership.{u1, u2} α β] {s : β} {t : β} {a : α}, (Membership.Mem.{u1, u2} α β _inst_1 a s) -> (Not (Membership.Mem.{u1, u2} α β _inst_1 a t)) -> (Ne.{succ u2} β s t)
-but is expected to have type
- forall {α : Type.{u2}} {β : Type.{u1}} [_inst_1 : Membership.{u2, u1} α β] {s : β} {t : β} {a : α}, (Membership.mem.{u2, u1} α β _inst_1 a s) -> (Not (Membership.mem.{u2, u1} α β _inst_1 a t)) -> (Ne.{succ u1} β s t)
-Case conversion may be inaccurate. Consider using '#align has_mem.mem.ne_of_not_mem' Membership.mem.ne_of_not_mem'ₓ'. -/
/-- **Alias** of `ne_of_mem_of_not_mem'`. -/
theorem Membership.mem.ne_of_not_mem' : a ∈ s → a ∉ t → s ≠ t :=
ne_of_mem_of_not_mem'
@@ -1752,24 +1620,12 @@ theorem ball_cond_comm {α} {s : α → Prop} {p : α → α → Prop} :
#align ball_cond_comm ball_cond_comm
-/
-/- warning: ball_mem_comm -> ball_mem_comm is a dubious translation:
-lean 3 declaration is
- forall {α : outParam.{succ (succ u1)} Type.{u1}} {β : Type.{u2}} [_inst_1 : Membership.{u1, u2} α β] {s : β} {p : α -> α -> Prop}, Iff (forall (a : α), (Membership.Mem.{u1, u2} α β _inst_1 a s) -> (forall (b : α), (Membership.Mem.{u1, u2} α β _inst_1 b s) -> (p a b))) (forall (a : α) (b : α), (Membership.Mem.{u1, u2} α β _inst_1 a s) -> (Membership.Mem.{u1, u2} α β _inst_1 b s) -> (p a b))
-but is expected to have type
- forall {α : Type.{u2}} {β : Type.{u1}} [_inst_1 : Membership.{u2, u1} α β] {s : β} {p : α -> α -> Prop}, Iff (forall (a : α), (Membership.mem.{u2, u1} α β _inst_1 a s) -> (forall (b : α), (Membership.mem.{u2, u1} α β _inst_1 b s) -> (p a b))) (forall (a : α) (b : α), (Membership.mem.{u2, u1} α β _inst_1 a s) -> (Membership.mem.{u2, u1} α β _inst_1 b s) -> (p a b))
-Case conversion may be inaccurate. Consider using '#align ball_mem_comm ball_mem_commₓ'. -/
/- ./././Mathport/Syntax/Translate/Basic.lean:635:2: warning: expanding binder collection (a b «expr ∈ » s) -/
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
-/- warning: ne_of_apply_ne -> ne_of_apply_ne is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {β : Sort.{u2}} (f : α -> β) {x : α} {y : α}, (Ne.{u2} β (f x) (f y)) -> (Ne.{u1} α x y)
-but is expected to have type
- forall {α : Sort.{u2}} {β : Sort.{u1}} (f : α -> β) {x : α} {y : α}, (Ne.{u1} β (f x) (f y)) -> (Ne.{u2} α x y)
-Case conversion may be inaccurate. Consider using '#align ne_of_apply_ne ne_of_apply_neₓ'. -/
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
@@ -1780,12 +1636,6 @@ theorem eq_equivalence : Equivalence (@Eq α) :=
#align eq_equivalence eq_equivalence
-/
-/- warning: eq_rec_constant -> eq_rec_constant is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {a : α} {a' : α} {β : Sort.{u2}} (y : β) (h : Eq.{u1} α a a'), Eq.{u2} ((fun (a : α) => β) a') (Eq.ndrec.{u2, u1} α a (fun (a : α) => β) y a' h) y
-but is expected to have type
- forall {α : Sort.{u2}} {a : α} {a' : α} {β : Sort.{u1}} (y : β) (h : Eq.{u2} α a a'), Eq.{u1} β (Eq.rec.{u1, u2} α a (fun (a_1 : α) (x._@.Std.Logic._hyg.12340 : Eq.{u2} α a a_1) => β) y a' h) y
-Case conversion may be inaccurate. Consider using '#align eq_rec_constant eq_rec_constantₓ'. -/
/-- Transport through trivial families is the identity. -/
@[simp]
theorem eq_rec_constant {α : Sort _} {a a' : α} {β : Sort _} (y : β) (h : a = a') :
@@ -1814,60 +1664,30 @@ theorem cast_cast :
#align cast_cast cast_cast
-/
-/- warning: congr_refl_left -> congr_refl_left is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {β : Sort.{u2}} (f : α -> β) {a : α} {b : α} (h : Eq.{u1} α a b), Eq.{0} (Eq.{u2} β (f a) (f b)) (congr.{u1, u2} α β f f a b (rfl.{imax u1 u2} (α -> β) f) h) (congr_arg.{u1, u2} α β a b f h)
-but is expected to have type
- forall {α : Sort.{u2}} {β : Sort.{u1}} (f : α -> β) {a : α} {b : α} (h : Eq.{u2} α a b), Eq.{0} (Eq.{u1} β (f a) (f b)) (congr.{u2, u1} α β f f a b (Eq.refl.{imax u2 u1} (α -> β) f) h) (congr_arg.{u2, u1} α β a b f h)
-Case conversion may be inaccurate. Consider using '#align congr_refl_left congr_refl_leftₓ'. -/
@[simp]
theorem congr_refl_left {α β : Sort _} (f : α → β) {a b : α} (h : a = b) :
congr (Eq.refl f) h = congr_arg f h :=
rfl
#align congr_refl_left congr_refl_left
-/- warning: congr_refl_right -> congr_refl_right is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {β : Sort.{u2}} {f : α -> β} {g : α -> β} (h : Eq.{imax u1 u2} (α -> β) f g) (a : α), Eq.{0} (Eq.{u2} β (f a) (g a)) (congr.{u1, u2} α β f g a a h (rfl.{u1} α a)) (congr_fun.{u1, u2} α (fun (x : α) => β) f (fun (a : α) => g a) h a)
-but is expected to have type
- forall {α : Sort.{u2}} {β : Sort.{u1}} {f : α -> β} {g : α -> β} (h : Eq.{imax u2 u1} (α -> β) f g) (a : α), Eq.{0} (Eq.{u1} β (f a) (g a)) (congr.{u2, u1} α β f g a a h (Eq.refl.{u2} α a)) (congr_fun.{u2, u1} α (fun (x : α) => β) f g h a)
-Case conversion may be inaccurate. Consider using '#align congr_refl_right congr_refl_rightₓ'. -/
@[simp]
theorem congr_refl_right {α β : Sort _} {f g : α → β} (h : f = g) (a : α) :
congr h (Eq.refl a) = congr_fun h a :=
rfl
#align congr_refl_right congr_refl_right
-/- warning: congr_arg_refl -> congr_arg_refl is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {β : Sort.{u2}} (f : α -> β) (a : α), Eq.{0} (Eq.{u2} β (f a) (f a)) (congr_arg.{u1, u2} α β a a f (rfl.{u1} α a)) (rfl.{u2} β (f a))
-but is expected to have type
- forall {α : Sort.{u2}} {β : Sort.{u1}} (f : α -> β) (a : α), Eq.{0} (Eq.{u1} β (f a) (f a)) (congr_arg.{u2, u1} α β a a f (Eq.refl.{u2} α a)) (Eq.refl.{u1} β (f a))
-Case conversion may be inaccurate. Consider using '#align congr_arg_refl congr_arg_reflₓ'. -/
@[simp]
theorem congr_arg_refl {α β : Sort _} (f : α → β) (a : α) :
congr_arg f (Eq.refl a) = Eq.refl (f a) :=
rfl
#align congr_arg_refl congr_arg_refl
-/- warning: congr_fun_rfl -> congr_fun_rfl is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {β : Sort.{u2}} (f : α -> β) (a : α), Eq.{0} (Eq.{u2} β (f a) (f a)) (congr_fun.{u1, u2} α (fun (ᾰ : α) => β) f f (rfl.{imax u1 u2} (α -> β) f) a) (rfl.{u2} β (f a))
-but is expected to have type
- forall {α : Sort.{u2}} {β : Sort.{u1}} (f : α -> β) (a : α), Eq.{0} (Eq.{u1} β (f a) (f a)) (congr_fun.{u2, u1} α (fun (ᾰ : α) => β) f f (Eq.refl.{imax u2 u1} (α -> β) f) a) (Eq.refl.{u1} β (f a))
-Case conversion may be inaccurate. Consider using '#align congr_fun_rfl congr_fun_rflₓ'. -/
@[simp]
theorem congr_fun_rfl {α β : Sort _} (f : α → β) (a : α) :
congr_fun (Eq.refl f) a = Eq.refl (f a) :=
rfl
#align congr_fun_rfl congr_fun_rfl
-/- warning: congr_fun_congr_arg -> congr_fun_congr_arg is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {β : Sort.{u2}} {γ : Sort.{u3}} (f : α -> β -> γ) {a : α} {a' : α} (p : Eq.{u1} α a a') (b : β), Eq.{0} (Eq.{u3} γ (f a b) (f a' b)) (congr_fun.{u2, u3} β (fun (ᾰ : β) => γ) (f a) (f a') (congr_arg.{u1, imax u2 u3} α (β -> γ) a a' f p) b) (congr_arg.{u1, u3} α γ a a' (fun (a : α) => f a b) p)
-but is expected to have type
- forall {α : Sort.{u3}} {β : Sort.{u2}} {γ : Sort.{u1}} (f : α -> β -> γ) {a : α} {a' : α} (p : Eq.{u3} α a a') (b : β), Eq.{0} (Eq.{u1} γ (f a b) (f a' b)) (congr_fun.{u2, u1} β (fun (ᾰ : β) => γ) (f a) (f a') (congr_arg.{u3, imax u2 u1} α (β -> γ) a a' f p) b) (congr_arg.{u3, u1} α γ a a' (fun (a : α) => f a b) p)
-Case conversion may be inaccurate. Consider using '#align congr_fun_congr_arg congr_fun_congr_argₓ'. -/
@[simp]
theorem congr_fun_congr_arg {α β γ : Sort _} (f : α → β → γ) {a a' : α} (p : a = a') (b : β) :
congr_fun (congr_arg f p) b = congr_arg (fun a => f a b) p :=
@@ -1881,52 +1701,22 @@ theorem heq_of_cast_eq :
#align heq_of_cast_eq heq_of_cast_eq
-/
-/- warning: cast_eq_iff_heq -> cast_eq_iff_heq is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {β : Sort.{u1}} {a : α} {a' : β} {e : Eq.{succ u1} Sort.{u1} α β}, Iff (Eq.{u1} β (cast.{u1} α β e a) a') (HEq.{u1} α a β a')
-but is expected to have type
- forall {α : Sort.{u1}} {β : Sort.{u1}} {a : Eq.{succ u1} Sort.{u1} α β} {a' : α} {e : β}, Iff (Eq.{u1} β (cast.{u1} α β a a') e) (HEq.{u1} α a' β e)
-Case conversion may be inaccurate. Consider using '#align cast_eq_iff_heq cast_eq_iff_heqₓ'. -/
theorem cast_eq_iff_heq {α β : Sort _} {a : α} {a' : β} {e : α = β} : 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
-/- warning: rec_heq_of_heq -> rec_heq_of_heq is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {a : α} {b : α} {β : Sort.{u2}} {C : α -> Sort.{u2}} {x : C a} {y : β} (e : Eq.{u1} α a b), (HEq.{u2} (C a) x β y) -> (HEq.{u2} (C b) (Eq.ndrec.{u2, u1} α a C x b e) β y)
-but is expected to have type
- forall {α : Sort.{u2}} {a : α} {b : Sort.{u1}} {β : α} {C : α -> Sort.{u1}} {x : C a} {y : b} (e : Eq.{u2} α a β), (HEq.{u1} (C a) x b y) -> (HEq.{u1} (C β) (Eq.rec.{u1, u2} α a (fun (x._@.Mathlib.Logic.Basic._hyg.4255 : α) (h._@.Mathlib.Logic.Basic._hyg.4256 : Eq.{u2} α a x._@.Mathlib.Logic.Basic._hyg.4255) => C x._@.Mathlib.Logic.Basic._hyg.4255) x β e) b y)
-Case conversion may be inaccurate. Consider using '#align rec_heq_of_heq rec_heq_of_heqₓ'. -/
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
-/- warning: rec_heq_iff_heq -> rec_heq_iff_heq is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {a : α} {b : α} {β : Sort.{u2}} {C : α -> Sort.{u2}} {x : C a} {y : β} {e : Eq.{u1} α a b}, Iff (HEq.{u2} (C b) (Eq.ndrec.{u2, u1} α a C x b e) β y) (HEq.{u2} (C a) x β y)
-but is expected to have type
- forall {α : Sort.{u2}} {a : α} {b : Sort.{u1}} {β : α} {C : α -> Sort.{u1}} {x : C a} {y : b} {e : Eq.{u2} α a β}, Iff (HEq.{u1} (C β) (Eq.rec.{u1, u2} α a (fun (x._@.Mathlib.Logic.Basic._hyg.4315 : α) (h._@.Mathlib.Logic.Basic._hyg.4316 : Eq.{u2} α a x._@.Mathlib.Logic.Basic._hyg.4315) => C x._@.Mathlib.Logic.Basic._hyg.4315) x β e) b y) (HEq.{u1} (C a) x b y)
-Case conversion may be inaccurate. Consider using '#align rec_heq_iff_heq rec_heq_iff_heqₓ'. -/
theorem rec_heq_iff_heq {β} {C : α → Sort _} {x : C a} {y : β} {e : a = b} :
HEq (@Eq.ndrec α a C x b e) y ↔ HEq x y := by subst e
#align rec_heq_iff_heq rec_heq_iff_heq
-/- warning: heq_rec_iff_heq -> heq_rec_iff_heq is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {a : α} {b : α} {β : Sort.{u2}} {C : α -> Sort.{u2}} {x : β} {y : C a} {e : Eq.{u1} α a b}, Iff (HEq.{u2} β x (C b) (Eq.ndrec.{u2, u1} α a C y b e)) (HEq.{u2} β x (C a) y)
-but is expected to have type
- forall {α : Sort.{u2}} {a : Sort.{u1}} {b : α} {β : α} {C : α -> Sort.{u1}} {x : a} {y : C b} {e : Eq.{u2} α b β}, Iff (HEq.{u1} a x (C β) (Eq.rec.{u1, u2} α b (fun (x._@.Mathlib.Logic.Basic._hyg.4377 : α) (h._@.Mathlib.Logic.Basic._hyg.4378 : Eq.{u2} α b x._@.Mathlib.Logic.Basic._hyg.4377) => C x._@.Mathlib.Logic.Basic._hyg.4377) y β e)) (HEq.{u1} a x (C b) y)
-Case conversion may be inaccurate. Consider using '#align heq_rec_iff_heq heq_rec_iff_heqₓ'. -/
theorem heq_rec_iff_heq {β} {C : α → Sort _} {x : β} {y : C a} {e : a = b} :
HEq x (@Eq.ndrec α a C y b e) ↔ HEq x y := by subst e
#align heq_rec_iff_heq heq_rec_iff_heq
-/- warning: eq.congr -> Eq.congr is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {x₁ : α} {x₂ : α} {y₁ : α} {y₂ : α}, (Eq.{u1} α x₁ y₁) -> (Eq.{u1} α x₂ y₂) -> (Iff (Eq.{u1} α x₁ x₂) (Eq.{u1} α y₁ y₂))
-but is expected to have type
- forall {α : Sort.{u1}} {x₁ : α} {x₂ : α} {y₁ : α} {y₂ : α}, (Eq.{u1} α x₁ x₂) -> (Eq.{u1} α y₁ y₂) -> (Iff (Eq.{u1} α x₁ y₁) (Eq.{u1} α x₂ y₂))
-Case conversion may be inaccurate. Consider using '#align eq.congr Eq.congrₓ'. -/
protected theorem Eq.congr {x₁ x₂ y₁ y₂ : α} (h₁ : x₁ = y₁) (h₂ : x₂ = y₂) : x₁ = x₂ ↔ y₁ = y₂ := by
subst h₁; subst h₂
#align eq.congr Eq.congr
@@ -1941,55 +1731,25 @@ theorem Eq.congr_right {x y z : α} (h : x = y) : z = x ↔ z = y := by rw [h]
#align eq.congr_right Eq.congr_right
-/
-/- warning: congr_arg2 -> congr_arg₂ is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {β : Sort.{u2}} {γ : Sort.{u3}} (f : α -> β -> γ) {x : α} {x' : α} {y : β} {y' : β}, (Eq.{u1} α x x') -> (Eq.{u2} β y y') -> (Eq.{u3} γ (f x y) (f x' y'))
-but is expected to have type
- forall {α : Sort.{u3}} {β : Sort.{u2}} {γ : Sort.{u1}} (f : α -> β -> γ) {x : α} {x' : α} {y : β} {y' : β}, (Eq.{u3} α x x') -> (Eq.{u2} β y y') -> (Eq.{u1} γ (f x y) (f x' y'))
-Case conversion may be inaccurate. Consider using '#align congr_arg2 congr_arg₂ₓ'. -/
theorem congr_arg₂ {α β γ : Sort _} (f : α → β → γ) {x x' : α} {y y' : β} (hx : x = x')
(hy : y = y') : f x y = f x' y' := by subst hx; subst hy
#align congr_arg2 congr_arg₂
variable {β : α → Sort _} {γ : ∀ a, β a → Sort _} {δ : ∀ a b, γ a b → Sort _}
-/- warning: congr_fun₂ -> congr_fun₂ is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {β : α -> Sort.{u2}} {γ : forall (a : α), (β a) -> Sort.{u3}} {f : forall (a : α) (b : β a), γ a b} {g : forall (a : α) (b : β a), γ a b}, (Eq.{imax u1 u2 u3} (forall (a : α) (b : β a), γ a b) f g) -> (forall (a : α) (b : β a), Eq.{u3} (γ a b) (f a b) (g a b))
-but is expected to have type
- forall {α : Sort.{u3}} {β : α -> Sort.{u2}} {γ : forall (a : α), (β a) -> Sort.{u1}} {f : forall (a : α) (b : β a), γ a b} {g : forall (a : α) (b : β a), γ a b}, (Eq.{imax u3 u2 u1} (forall (a : α) (b : β a), γ a b) f g) -> (forall (a : α) (b : β a), Eq.{u1} (γ a b) (f a b) (g a b))
-Case conversion may be inaccurate. Consider using '#align congr_fun₂ congr_fun₂ₓ'. -/
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₂
-/- warning: congr_fun₃ -> congr_fun₃ is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {β : α -> Sort.{u2}} {γ : forall (a : α), (β a) -> Sort.{u3}} {δ : forall (a : α) (b : β a), (γ a b) -> Sort.{u4}} {f : forall (a : α) (b : β a) (c : γ a b), δ a b c} {g : forall (a : α) (b : β a) (c : γ a b), δ a b c}, (Eq.{imax u1 u2 u3 u4} (forall (a : α) (b : β a) (c : γ a b), δ a b c) f g) -> (forall (a : α) (b : β a) (c : γ a b), Eq.{u4} (δ a b c) (f a b c) (g a b c))
-but is expected to have type
- forall {α : Sort.{u4}} {β : α -> Sort.{u3}} {γ : forall (a : α), (β a) -> Sort.{u2}} {δ : forall (a : α) (b : β a), (γ a b) -> Sort.{u1}} {f : forall (a : α) (b : β a) (c : γ a b), δ a b c} {g : forall (a : α) (b : β a) (c : γ a b), δ a b c}, (Eq.{imax u4 u3 u2 u1} (forall (a : α) (b : β a) (c : γ a b), δ a b c) f g) -> (forall (a : α) (b : β a) (c : γ a b), Eq.{u1} (δ a b c) (f a b c) (g a b c))
-Case conversion may be inaccurate. Consider using '#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₃
-/- warning: funext₂ -> funext₂ is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {β : α -> Sort.{u2}} {γ : forall (a : α), (β a) -> Sort.{u3}} {f : forall (a : α) (b : β a), γ a b} {g : forall (a : α) (b : β a), γ a b}, (forall (a : α) (b : β a), Eq.{u3} (γ a b) (f a b) (g a b)) -> (Eq.{imax u1 u2 u3} (forall (a : α) (b : β a), γ a b) f g)
-but is expected to have type
- forall {α : Sort.{u2}} {β : α -> Sort.{u1}} {γ : forall (a : α), (β a) -> Sort.{u3}} {f : forall (a : α) (b : β a), γ a b} {g : forall (a : α) (b : β a), γ a b}, (forall (a : α) (b : β a), Eq.{u3} (γ a b) (f a b) (g a b)) -> (Eq.{imax u2 u1 u3} (forall (a : α) (b : β a), γ a b) f g)
-Case conversion may be inaccurate. Consider using '#align funext₂ funext₂ₓ'. -/
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₂
-/- warning: funext₃ -> funext₃ is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {β : α -> Sort.{u2}} {γ : forall (a : α), (β a) -> Sort.{u3}} {δ : forall (a : α) (b : β a), (γ a b) -> Sort.{u4}} {f : forall (a : α) (b : β a) (c : γ a b), δ a b c} {g : forall (a : α) (b : β a) (c : γ a b), δ a b c}, (forall (a : α) (b : β a) (c : γ a b), Eq.{u4} (δ a b c) (f a b c) (g a b c)) -> (Eq.{imax u1 u2 u3 u4} (forall (a : α) (b : β a) (c : γ a b), δ a b c) f g)
-but is expected to have type
- forall {α : Sort.{u3}} {β : α -> Sort.{u2}} {γ : forall (a : α), (β a) -> Sort.{u1}} {δ : forall (a : α) (b : β a), (γ a b) -> Sort.{u4}} {f : forall (a : α) (b : β a) (c : γ a b), δ a b c} {g : forall (a : α) (b : β a) (c : γ a b), δ a b c}, (forall (a : α) (b : β a) (c : γ a b), Eq.{u4} (δ a b c) (f a b c) (g a b c)) -> (Eq.{imax u3 u2 u1 u4} (forall (a : α) (b : β a) (c : γ a b), δ a b c) f g)
-Case conversion may be inaccurate. Consider using '#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₃
@@ -2043,45 +1803,21 @@ theorem forall₅_congr {p q : ∀ a b c d, ε a b c d → Prop}
#align forall₅_congr forall₅_congr
-/
-/- warning: exists₂_congr -> exists₂_congr is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {β : α -> Sort.{u2}} {p : forall (a : α), (β a) -> Prop} {q : forall (a : α), (β a) -> Prop}, (forall (a : α) (b : β a), Iff (p a b) (q a b)) -> (Iff (Exists.{u1} α (fun (a : α) => Exists.{u2} (β a) (fun (b : β a) => p a b))) (Exists.{u1} α (fun (a : α) => Exists.{u2} (β a) (fun (b : β a) => q a b))))
-but is expected to have type
- forall {α : Sort.{u2}} {β : α -> Sort.{u1}} {p : forall (a : α), (β a) -> Prop} {q : forall (a : α), (β a) -> Prop}, (forall (a : α) (b : β a), Iff (p a b) (q a b)) -> (Iff (Exists.{u2} α (fun (a : α) => Exists.{u1} (β a) (fun (b : β a) => p a b))) (Exists.{u2} α (fun (a : α) => Exists.{u1} (β a) (fun (b : β a) => q a b))))
-Case conversion may be inaccurate. Consider using '#align exists₂_congr exists₂_congrₓ'. -/
theorem exists₂_congr {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_congr fun a => exists_congr <| h a
#align exists₂_congr exists₂_congr
-/- warning: exists₃_congr -> exists₃_congr is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {β : α -> Sort.{u2}} {γ : forall (a : α), (β a) -> Sort.{u3}} {p : forall (a : α) (b : β a), (γ a b) -> Prop} {q : forall (a : α) (b : β a), (γ a b) -> Prop}, (forall (a : α) (b : β a) (c : γ a b), Iff (p a b c) (q a b c)) -> (Iff (Exists.{u1} α (fun (a : α) => Exists.{u2} (β a) (fun (b : β a) => Exists.{u3} (γ a b) (fun (c : γ a b) => p a b c)))) (Exists.{u1} α (fun (a : α) => Exists.{u2} (β a) (fun (b : β a) => Exists.{u3} (γ a b) (fun (c : γ a b) => q a b c)))))
-but is expected to have type
- forall {α : Sort.{u3}} {β : α -> Sort.{u2}} {γ : forall (a : α), (β a) -> Sort.{u1}} {p : forall (a : α) (b : β a), (γ a b) -> Prop} {q : forall (a : α) (b : β a), (γ a b) -> Prop}, (forall (a : α) (b : β a) (c : γ a b), Iff (p a b c) (q a b c)) -> (Iff (Exists.{u3} α (fun (a : α) => Exists.{u2} (β a) (fun (b : β a) => Exists.{u1} (γ a b) (fun (c : γ a b) => p a b c)))) (Exists.{u3} α (fun (a : α) => Exists.{u2} (β a) (fun (b : β a) => Exists.{u1} (γ a b) (fun (c : γ a b) => q a b c)))))
-Case conversion may be inaccurate. Consider using '#align exists₃_congr exists₃_congrₓ'. -/
theorem exists₃_congr {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_congr fun a => exists₂_congr <| h a
#align exists₃_congr exists₃_congr
-/- warning: exists₄_congr -> exists₄_congr is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {β : α -> Sort.{u2}} {γ : forall (a : α), (β a) -> Sort.{u3}} {δ : forall (a : α) (b : β a), (γ a b) -> Sort.{u4}} {p : forall (a : α) (b : β a) (c : γ a b), (δ a b c) -> Prop} {q : forall (a : α) (b : β a) (c : γ a b), (δ a b c) -> Prop}, (forall (a : α) (b : β a) (c : γ a b) (d : δ a b c), Iff (p a b c d) (q a b c d)) -> (Iff (Exists.{u1} α (fun (a : α) => Exists.{u2} (β a) (fun (b : β a) => Exists.{u3} (γ a b) (fun (c : γ a b) => Exists.{u4} (δ a b c) (fun (d : δ a b c) => p a b c d))))) (Exists.{u1} α (fun (a : α) => Exists.{u2} (β a) (fun (b : β a) => Exists.{u3} (γ a b) (fun (c : γ a b) => Exists.{u4} (δ a b c) (fun (d : δ a b c) => q a b c d))))))
-but is expected to have type
- forall {α : Sort.{u4}} {β : α -> Sort.{u3}} {γ : forall (a : α), (β a) -> Sort.{u2}} {δ : forall (a : α) (b : β a), (γ a b) -> Sort.{u1}} {p : forall (a : α) (b : β a) (c : γ a b), (δ a b c) -> Prop} {q : forall (a : α) (b : β a) (c : γ a b), (δ a b c) -> Prop}, (forall (a : α) (b : β a) (c : γ a b) (d : δ a b c), Iff (p a b c d) (q a b c d)) -> (Iff (Exists.{u4} α (fun (a : α) => Exists.{u3} (β a) (fun (b : β a) => Exists.{u2} (γ a b) (fun (c : γ a b) => Exists.{u1} (δ a b c) (fun (d : δ a b c) => p a b c d))))) (Exists.{u4} α (fun (a : α) => Exists.{u3} (β a) (fun (b : β a) => Exists.{u2} (γ a b) (fun (c : γ a b) => Exists.{u1} (δ a b c) (fun (d : δ a b c) => q a b c d))))))
-Case conversion may be inaccurate. Consider using '#align exists₄_congr exists₄_congrₓ'. -/
theorem exists₄_congr {p q : ∀ a b c, δ a b c → Prop} (h : ∀ a b c d, p a b c d ↔ q a b c d) :
(∃ a b c d, p a b c d) ↔ ∃ a b c d, q a b c d :=
exists_congr fun a => exists₃_congr <| h a
#align exists₄_congr exists₄_congr
-/- warning: exists₅_congr -> exists₅_congr is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {β : α -> Sort.{u2}} {γ : forall (a : α), (β a) -> Sort.{u3}} {δ : forall (a : α) (b : β a), (γ a b) -> Sort.{u4}} {ε : forall (a : α) (b : β a) (c : γ a b), (δ a b c) -> Sort.{u5}} {p : forall (a : α) (b : β a) (c : γ a b) (d : δ a b c), (ε a b c d) -> Prop} {q : forall (a : α) (b : β a) (c : γ a b) (d : δ a b c), (ε a b c d) -> Prop}, (forall (a : α) (b : β a) (c : γ a b) (d : δ a b c) (e : ε a b c d), Iff (p a b c d e) (q a b c d e)) -> (Iff (Exists.{u1} α (fun (a : α) => Exists.{u2} (β a) (fun (b : β a) => Exists.{u3} (γ a b) (fun (c : γ a b) => Exists.{u4} (δ a b c) (fun (d : δ a b c) => Exists.{u5} (ε a b c d) (fun (e : ε a b c d) => p a b c d e)))))) (Exists.{u1} α (fun (a : α) => Exists.{u2} (β a) (fun (b : β a) => Exists.{u3} (γ a b) (fun (c : γ a b) => Exists.{u4} (δ a b c) (fun (d : δ a b c) => Exists.{u5} (ε a b c d) (fun (e : ε a b c d) => q a b c d e)))))))
-but is expected to have type
- forall {α : Sort.{u5}} {β : α -> Sort.{u4}} {γ : forall (a : α), (β a) -> Sort.{u3}} {δ : forall (a : α) (b : β a), (γ a b) -> Sort.{u2}} {ε : forall (a : α) (b : β a) (c : γ a b), (δ a b c) -> Sort.{u1}} {p : forall (a : α) (b : β a) (c : γ a b) (d : δ a b c), (ε a b c d) -> Prop} {q : forall (a : α) (b : β a) (c : γ a b) (d : δ a b c), (ε a b c d) -> Prop}, (forall (a : α) (b : β a) (c : γ a b) (d : δ a b c) (e : ε a b c d), Iff (p a b c d e) (q a b c d e)) -> (Iff (Exists.{u5} α (fun (a : α) => Exists.{u4} (β a) (fun (b : β a) => Exists.{u3} (γ a b) (fun (c : γ a b) => Exists.{u2} (δ a b c) (fun (d : δ a b c) => Exists.{u1} (ε a b c d) (fun (e : ε a b c d) => p a b c d e)))))) (Exists.{u5} α (fun (a : α) => Exists.{u4} (β a) (fun (b : β a) => Exists.{u3} (γ a b) (fun (c : γ a b) => Exists.{u2} (δ a b c) (fun (d : δ a b c) => Exists.{u1} (ε a b c d) (fun (e : ε a b c d) => q a b c d e)))))))
-Case conversion may be inaccurate. Consider using '#align exists₅_congr exists₅_congrₓ'. -/
theorem exists₅_congr {p q : ∀ a b c d, ε a b c d → Prop}
(h : ∀ a b c d e, p a b c d e ↔ q a b c d e) :
(∃ a b c d e, p a b c d e) ↔ ∃ a b c d e, q a b c d e :=
@@ -2114,23 +1850,11 @@ theorem Exists.imp {p q : α → Prop} (h : ∀ a, p a → q a) : (∃ a, p a)
#align Exists.imp Exists.imp
-/
-/- warning: Exists₂.imp -> Exists₂.imp is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {β : α -> Sort.{u2}} {p : forall (a : α), (β a) -> Prop} {q : forall (a : α), (β a) -> Prop}, (forall (a : α) (b : β a), (p a b) -> (q a b)) -> (Exists.{u1} α (fun (a : α) => Exists.{u2} (β a) (fun (b : β a) => p a b))) -> (Exists.{u1} α (fun (a : α) => Exists.{u2} (β a) (fun (b : β a) => q a b)))
-but is expected to have type
- forall {α : Sort.{u2}} {β : α -> Sort.{u1}} {p : forall (a : α), (β a) -> Prop} {q : forall (a : α), (β a) -> Prop}, (forall (a : α) (b : β a), (p a b) -> (q a b)) -> (Exists.{u2} α (fun (a : α) => Exists.{u1} (β a) (fun (b : β a) => p a b))) -> (Exists.{u2} α (fun (a : α) => Exists.{u1} (β a) (fun (b : β a) => q a b)))
-Case conversion may be inaccurate. Consider using '#align Exists₂.imp Exists₂.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
-/- warning: Exists₃.imp -> Exists₃.imp is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {β : α -> Sort.{u2}} {γ : forall (a : α), (β a) -> Sort.{u3}} {p : forall (a : α) (b : β a), (γ a b) -> Prop} {q : forall (a : α) (b : β a), (γ a b) -> Prop}, (forall (a : α) (b : β a) (c : γ a b), (p a b c) -> (q a b c)) -> (Exists.{u1} α (fun (a : α) => Exists.{u2} (β a) (fun (b : β a) => Exists.{u3} (γ a b) (fun (c : γ a b) => p a b c)))) -> (Exists.{u1} α (fun (a : α) => Exists.{u2} (β a) (fun (b : β a) => Exists.{u3} (γ a b) (fun (c : γ a b) => q a b c))))
-but is expected to have type
- forall {α : Sort.{u3}} {β : α -> Sort.{u2}} {γ : forall (a : α), (β a) -> Sort.{u1}} {p : forall (a : α) (b : β a), (γ a b) -> Prop} {q : forall (a : α) (b : β a), (γ a b) -> Prop}, (forall (a : α) (b : β a) (c : γ a b), (p a b c) -> (q a b c)) -> (Exists.{u3} α (fun (a : α) => Exists.{u2} (β a) (fun (b : β a) => Exists.{u1} (γ a b) (fun (c : γ a b) => p a b c)))) -> (Exists.{u3} α (fun (a : α) => Exists.{u2} (β a) (fun (b : β a) => Exists.{u1} (γ a b) (fun (c : γ a b) => q a b c))))
-Case conversion may be inaccurate. Consider using '#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
@@ -2140,33 +1864,15 @@ end Dependent
variable {ι β : Sort _} {κ : ι → Sort _} {p q : α → Prop} {b : Prop}
-/- warning: exists_imp_exists' -> Exists.imp' is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {β : Sort.{u2}} {p : α -> Prop} {q : β -> Prop} (f : α -> β), (forall (a : α), (p a) -> (q (f a))) -> (Exists.{u1} α (fun (a : α) => p a)) -> (Exists.{u2} β (fun (b : β) => q b))
-but is expected to have type
- forall {α : Sort.{u1}} {β : α -> Prop} {p : Sort.{u2}} {q : p -> Prop} (f : α -> p), (forall (a : α), (β a) -> (q (f a))) -> (Exists.{u1} α (fun (a : α) => β a)) -> (Exists.{u2} p (fun (b : p) => q b))
-Case conversion may be inaccurate. Consider using '#align exists_imp_exists' Exists.imp'ₓ'. -/
theorem Exists.imp' {p : α → Prop} {q : β → Prop} (f : α → β) (hpq : ∀ a, p a → q (f a))
(hp : ∃ a, p a) : ∃ b, q b :=
Exists.elim hp fun a hp' => ⟨_, hpq _ hp'⟩
#align exists_imp_exists' Exists.imp'
-/- warning: forall_swap -> forall_swap is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {β : Sort.{u2}} {p : α -> β -> Prop}, Iff (forall (x : α) (y : β), p x y) (forall (y : β) (x : α), p x y)
-but is expected to have type
- forall {α : Sort.{u2}} {β : Sort.{u1}} {p : α -> β -> Prop}, Iff (forall (x : α) (y : β), p x y) (forall (y : β) (x : α), p x y)
-Case conversion may be inaccurate. Consider using '#align forall_swap forall_swapₓ'. -/
theorem forall_swap {p : α → β → Prop} : (∀ x y, p x y) ↔ ∀ y x, p x y :=
⟨swap, swap⟩
#align forall_swap forall_swap
-/- warning: forall₂_swap -> forall₂_swap is a dubious translation:
-lean 3 declaration is
- forall {ι₁ : Sort.{u1}} {ι₂ : Sort.{u2}} {κ₁ : ι₁ -> Sort.{u3}} {κ₂ : ι₂ -> Sort.{u4}} {p : forall (i₁ : ι₁), (κ₁ i₁) -> (forall (i₂ : ι₂), (κ₂ i₂) -> Prop)}, Iff (forall (i₁ : ι₁) (j₁ : κ₁ i₁) (i₂ : ι₂) (j₂ : κ₂ i₂), p i₁ j₁ i₂ j₂) (forall (i₂ : ι₂) (j₂ : κ₂ i₂) (i₁ : ι₁) (j₁ : κ₁ i₁), p i₁ j₁ i₂ j₂)
-but is expected to have type
- forall {ι₁ : Sort.{u4}} {ι₂ : Sort.{u3}} {κ₁ : ι₁ -> Sort.{u2}} {κ₂ : ι₂ -> Sort.{u1}} {p : forall (i₁ : ι₁), (κ₁ i₁) -> (forall (i₂ : ι₂), (κ₂ i₂) -> Prop)}, Iff (forall (i₁ : ι₁) (j₁ : κ₁ i₁) (i₂ : ι₂) (j₂ : κ₂ i₂), p i₁ j₁ i₂ j₂) (forall (i₂ : ι₂) (j₂ : κ₂ i₂) (i₁ : ι₁) (j₁ : κ₁ i₁), p i₁ j₁ i₂ j₂)
-Case conversion may be inaccurate. Consider using '#align forall₂_swap forall₂_swapₓ'. -/
theorem forall₂_swap {ι₁ ι₂ : Sort _} {κ₁ : ι₁ → 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₂ :=
@@ -2218,12 +1924,6 @@ theorem Exists.choose_spec {p : α → Prop} (P : ∃ a, p a) : p P.some :=
#align Exists.some_spec Exists.choose_spec
-/
-/- warning: not_exists_of_forall_not -> not_exists_of_forall_not is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {p : α -> Prop}, (forall (x : α), Not (p x)) -> (Not (Exists.{u1} α (fun (x : α) => p x)))
-but is expected to have type
- forall {α : Sort.{u1}} {p : α -> Prop} {h : Prop}, (forall (x : α), (p x) -> h) -> (Exists.{u1} α (fun (x : α) => p x)) -> h
-Case conversion may be inaccurate. Consider using '#align not_exists_of_forall_not not_exists_of_forall_notₓ'. -/
--theorem forall_not_of_not_exists (h : ¬ ∃ x, p x) : ∀ x, ¬ p x :=
--forall_imp_of_exists_imp h
theorem not_exists_of_forall_not (h : ∀ x, ¬p x) : ¬∃ x, p x :=
@@ -2311,23 +2011,11 @@ theorem forall_true_iff' (h : ∀ a, p a ↔ True) : (∀ a, p a) ↔ True :=
#align forall_true_iff' forall_true_iff'
-/
-/- warning: forall_2_true_iff -> forall₂_true_iff is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {β : α -> Sort.{u2}}, Iff (forall (a : α), (β a) -> True) True
-but is expected to have type
- forall {α : Sort.{u2}} {β : α -> Sort.{u1}}, Iff (forall (a : α), (β a) -> True) True
-Case conversion may be inaccurate. Consider using '#align forall_2_true_iff forall₂_true_iffₓ'. -/
@[simp]
theorem forall₂_true_iff {β : α → Sort _} : (∀ a, β a → True) ↔ True :=
forall_true_iff' fun _ => forall_true_iff
#align forall_2_true_iff forall₂_true_iff
-/- warning: forall_3_true_iff -> forall₃_true_iff is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {β : α -> Sort.{u2}} {γ : forall (a : α), (β a) -> Sort.{u3}}, Iff (forall (a : α) (b : β a), (γ a b) -> True) True
-but is expected to have type
- forall {α : Sort.{u3}} {β : α -> Sort.{u2}} {γ : forall (a : α), (β a) -> Sort.{u1}}, Iff (forall (a : α) (b : β a), (γ a b) -> True) True
-Case conversion may be inaccurate. Consider using '#align forall_3_true_iff forall₃_true_iffₓ'. -/
@[simp]
theorem forall₃_true_iff {β : α → Sort _} {γ : ∀ a, β a → Sort _} :
(∀ (a) (b : β a), γ a b → True) ↔ True :=
@@ -2390,23 +2078,11 @@ theorem exists_or : (∃ x, p x ∨ q x) ↔ (∃ x, p x) ∨ ∃ x, q x :=
#align exists_or_distrib exists_or
-/
-/- warning: exists_and_distrib_left -> exists_and_left is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {q : Prop} {p : α -> Prop}, Iff (Exists.{u1} α (fun (x : α) => And q (p x))) (And q (Exists.{u1} α (fun (x : α) => p x)))
-but is expected to have type
- forall {α : Sort.{u1}} {q : α -> Prop} {p : Prop}, Iff (Exists.{u1} α (fun (x : α) => And p (q x))) (And p (Exists.{u1} α (fun (x : α) => q x)))
-Case conversion may be inaccurate. Consider using '#align exists_and_distrib_left exists_and_leftₓ'. -/
@[simp]
theorem exists_and_left {q : Prop} {p : α → Prop} : (∃ x, q ∧ p x) ↔ q ∧ ∃ x, p x :=
⟨fun ⟨x, hq, hp⟩ => ⟨hq, x, hp⟩, fun ⟨hq, x, hp⟩ => ⟨x, hq, hp⟩⟩
#align exists_and_distrib_left exists_and_left
-/- warning: exists_and_distrib_right -> exists_and_right is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {q : Prop} {p : α -> Prop}, Iff (Exists.{u1} α (fun (x : α) => And (p x) q)) (And (Exists.{u1} α (fun (x : α) => p x)) q)
-but is expected to have type
- forall {α : Sort.{u1}} {q : α -> Prop} {p : Prop}, Iff (Exists.{u1} α (fun (x : α) => And (q x) p)) (And (Exists.{u1} α (fun (x : α) => q x)) p)
-Case conversion may be inaccurate. Consider using '#align exists_and_distrib_right exists_and_rightₓ'. -/
@[simp]
theorem exists_and_right {q : Prop} {p : α → Prop} : (∃ x, p x ∧ q) ↔ (∃ x, p x) ∧ q := by
simp [and_comm']
@@ -2425,12 +2101,6 @@ theorem forall_eq' {a' : α} : (∀ a, a' = a → p a) ↔ p a' := by simp [@eq_
#align forall_eq' forall_eq'
-/
-/- warning: decidable.and_forall_ne -> Decidable.and_forall_ne is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {p : α -> Prop} [_inst_1 : DecidableEq.{u1} α] (a : α), Iff (And (p a) (forall (b : α), (Ne.{u1} α b a) -> (p b))) (forall (b : α), p b)
-but is expected to have type
- forall {α : Sort.{u1}} [p : DecidableEq.{u1} α] (_inst_1 : α) {a : α -> Prop}, Iff (And (a _inst_1) (forall (b : α), (Ne.{u1} α b _inst_1) -> (a b))) (forall (b : α), a b)
-Case conversion may be inaccurate. Consider using '#align decidable.and_forall_ne Decidable.and_forall_neₓ'. -/
/- ./././Mathport/Syntax/Translate/Basic.lean:635:2: warning: expanding binder collection (b «expr ≠ » a) -/
theorem Decidable.and_forall_ne [DecidableEq α] (a : α) :
(p a ∧ ∀ (b) (_ : b ≠ a), p b) ↔ ∀ b, p b := by
@@ -2499,23 +2169,11 @@ theorem exists_eq_right {a' : α} : (∃ a, p a ∧ a = a') ↔ p a' :=
#align exists_eq_right exists_eq_right
-/
-/- warning: exists_eq_right_right -> exists_eq_right_right is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {p : α -> Prop} {q : α -> Prop} {a' : α}, Iff (Exists.{u1} α (fun (a : α) => And (p a) (And (q a) (Eq.{u1} α a a')))) (And (p a') (q a'))
-but is expected to have type
- forall {α : Sort.{u1}} {p : α -> Prop} {q : Prop} {a' : α}, Iff (Exists.{u1} α (fun (a : α) => And (p a) (And q (Eq.{u1} α a a')))) (And (p a') q)
-Case conversion may be inaccurate. Consider using '#align exists_eq_right_right exists_eq_right_rightₓ'. -/
@[simp]
theorem exists_eq_right_right {a' : α} : (∃ a : α, p a ∧ q a ∧ a = a') ↔ p a' ∧ q a' :=
⟨fun ⟨_, hp, hq, rfl⟩ => ⟨hp, hq⟩, fun ⟨hp, hq⟩ => ⟨a', hp, hq, rfl⟩⟩
#align exists_eq_right_right exists_eq_right_right
-/- warning: exists_eq_right_right' -> exists_eq_right_right' is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {p : α -> Prop} {q : α -> Prop} {a' : α}, Iff (Exists.{u1} α (fun (a : α) => And (p a) (And (q a) (Eq.{u1} α a' a)))) (And (p a') (q a'))
-but is expected to have type
- forall {α : Sort.{u1}} {p : α -> Prop} {q : Prop} {a' : α}, Iff (Exists.{u1} α (fun (a : α) => And (p a) (And q (Eq.{u1} α a' a)))) (And (p a') q)
-Case conversion may be inaccurate. Consider using '#align exists_eq_right_right' exists_eq_right_right'ₓ'. -/
@[simp]
theorem exists_eq_right_right' {a' : α} : (∃ a : α, p a ∧ q a ∧ a' = a) ↔ p a' ∧ q a' :=
⟨fun ⟨_, hp, hq, rfl⟩ => ⟨hp, hq⟩, fun ⟨hp, hq⟩ => ⟨a', hp, hq, rfl⟩⟩
@@ -2579,57 +2237,27 @@ theorem exists_or_eq_right' (y : α) (p : α → Prop) : ∃ x : α, p x ∨ y =
#align exists_or_eq_right' exists_or_eq_right'
-/
-/- warning: forall_apply_eq_imp_iff -> forall_apply_eq_imp_iff is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {β : Sort.{u2}} {f : α -> β} {p : β -> Prop}, Iff (forall (a : α) (b : β), (Eq.{u2} β (f a) b) -> (p b)) (forall (a : α), p (f a))
-but is expected to have type
- forall {α : Sort.{u2}} {β : Sort.{u1}} {f : α -> β} {p : β -> Prop}, Iff (forall (a : α) (b : β), (Eq.{u1} β (f a) b) -> (p b)) (forall (a : α), p (f a))
-Case conversion may be inaccurate. Consider using '#align forall_apply_eq_imp_iff forall_apply_eq_imp_iffₓ'. -/
@[simp]
theorem forall_apply_eq_imp_iff {f : α → β} {p : β → Prop} :
(∀ a, ∀ b, f a = b → p b) ↔ ∀ a, p (f a) :=
⟨fun h a => h a (f a) rfl, fun h a b hab => hab ▸ h a⟩
#align forall_apply_eq_imp_iff forall_apply_eq_imp_iff
-/- warning: forall_apply_eq_imp_iff' -> forall_apply_eq_imp_iff' is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {β : Sort.{u2}} {f : α -> β} {p : β -> Prop}, Iff (forall (b : β) (a : α), (Eq.{u2} β (f a) b) -> (p b)) (forall (a : α), p (f a))
-but is expected to have type
- forall {α : Sort.{u2}} {β : Sort.{u1}} {f : α -> β} {p : β -> Prop}, Iff (forall (b : β) (a : α), (Eq.{u1} β (f a) b) -> (p b)) (forall (a : α), p (f a))
-Case conversion may be inaccurate. Consider using '#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 rw [forall_swap]; simp
#align forall_apply_eq_imp_iff' forall_apply_eq_imp_iff'
-/- warning: forall_eq_apply_imp_iff -> forall_eq_apply_imp_iff is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {β : Sort.{u2}} {f : α -> β} {p : β -> Prop}, Iff (forall (a : α) (b : β), (Eq.{u2} β b (f a)) -> (p b)) (forall (a : α), p (f a))
-but is expected to have type
- forall {α : Sort.{u2}} {β : Sort.{u1}} {f : α -> β} {p : β -> Prop}, Iff (forall (a : α) (b : β), (Eq.{u1} β b (f a)) -> (p b)) (forall (a : α), p (f a))
-Case conversion may be inaccurate. Consider using '#align forall_eq_apply_imp_iff forall_eq_apply_imp_iffₓ'. -/
@[simp]
theorem forall_eq_apply_imp_iff {f : α → β} {p : β → Prop} :
(∀ a, ∀ b, b = f a → p b) ↔ ∀ a, p (f a) := by simp [@eq_comm _ _ (f _)]
#align forall_eq_apply_imp_iff forall_eq_apply_imp_iff
-/- warning: forall_eq_apply_imp_iff' -> forall_eq_apply_imp_iff' is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {β : Sort.{u2}} {f : α -> β} {p : β -> Prop}, Iff (forall (b : β) (a : α), (Eq.{u2} β b (f a)) -> (p b)) (forall (a : α), p (f a))
-but is expected to have type
- forall {α : Sort.{u2}} {β : Sort.{u1}} {f : α -> β} {p : β -> Prop}, Iff (forall (b : β) (a : α), (Eq.{u1} β b (f a)) -> (p b)) (forall (a : α), p (f a))
-Case conversion may be inaccurate. Consider using '#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 rw [forall_swap]; simp
#align forall_eq_apply_imp_iff' forall_eq_apply_imp_iff'
-/- warning: forall_apply_eq_imp_iff₂ -> forall_apply_eq_imp_iff₂ is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {β : Sort.{u2}} {f : α -> β} {p : α -> Prop} {q : β -> Prop}, Iff (forall (b : β) (a : α), (p a) -> (Eq.{u2} β (f a) b) -> (q b)) (forall (a : α), (p a) -> (q (f a)))
-but is expected to have type
- forall {α : Sort.{u2}} {β : Sort.{u1}} {f : α -> β} {p : α -> Prop} {q : β -> Prop}, Iff (forall (b : β) (a : α), (p a) -> (Eq.{u1} β (f a) b) -> (q b)) (forall (a : α), (p a) -> (q (f a)))
-Case conversion may be inaccurate. Consider using '#align forall_apply_eq_imp_iff₂ forall_apply_eq_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) :=
@@ -2654,12 +2282,6 @@ theorem exists_comm {p : α → β → Prop} : (∃ a b, p a b) ↔ ∃ b a, p a
#align exists_comm exists_comm
-/
-/- warning: exists₂_comm -> exists₂_comm is a dubious translation:
-lean 3 declaration is
- forall {ι₁ : Sort.{u1}} {ι₂ : Sort.{u2}} {κ₁ : ι₁ -> Sort.{u3}} {κ₂ : ι₂ -> Sort.{u4}} {p : forall (i₁ : ι₁), (κ₁ i₁) -> (forall (i₂ : ι₂), (κ₂ i₂) -> Prop)}, Iff (Exists.{u1} ι₁ (fun (i₁ : ι₁) => Exists.{u3} (κ₁ i₁) (fun (j₁ : κ₁ i₁) => Exists.{u2} ι₂ (fun (i₂ : ι₂) => Exists.{u4} (κ₂ i₂) (fun (j₂ : κ₂ i₂) => p i₁ j₁ i₂ j₂))))) (Exists.{u2} ι₂ (fun (i₂ : ι₂) => Exists.{u4} (κ₂ i₂) (fun (j₂ : κ₂ i₂) => Exists.{u1} ι₁ (fun (i₁ : ι₁) => Exists.{u3} (κ₁ i₁) (fun (j₁ : κ₁ i₁) => p i₁ j₁ i₂ j₂)))))
-but is expected to have type
- forall {ι₁ : Sort.{u4}} {ι₂ : Sort.{u3}} {κ₁ : ι₁ -> Sort.{u2}} {κ₂ : ι₂ -> Sort.{u1}} {p : forall (i₁ : ι₁), (κ₁ i₁) -> (forall (i₂ : ι₂), (κ₂ i₂) -> Prop)}, Iff (Exists.{u4} ι₁ (fun (i₁ : ι₁) => Exists.{u2} (κ₁ i₁) (fun (j₁ : κ₁ i₁) => Exists.{u3} ι₂ (fun (i₂ : ι₂) => Exists.{u1} (κ₂ i₂) (fun (j₂ : κ₂ i₂) => p i₁ j₁ i₂ j₂))))) (Exists.{u3} ι₂ (fun (i₂ : ι₂) => Exists.{u1} (κ₂ i₂) (fun (j₂ : κ₂ i₂) => Exists.{u4} ι₁ (fun (i₁ : ι₁) => Exists.{u2} (κ₁ i₁) (fun (j₁ : κ₁ i₁) => p i₁ j₁ i₂ j₂)))))
-Case conversion may be inaccurate. Consider using '#align exists₂_comm exists₂_commₓ'. -/
theorem exists₂_comm {ι₁ ι₂ : Sort _} {κ₁ : ι₁ → 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
@@ -2706,12 +2328,6 @@ theorem forall_or_right {q : Prop} {p : α → Prop} : (∀ x, p x ∨ q) ↔ (
#align forall_or_distrib_right forall_or_right
-/
-/- warning: exists_prop -> exists_prop is a dubious translation:
-lean 3 declaration is
- forall {p : Prop} {q : Prop}, Iff (Exists.{0} p (fun (h : p) => q)) (And p q)
-but is expected to have type
- forall {p : Prop} {q : Prop}, Iff (Exists.{0} q (fun (h : q) => p)) (And q p)
-Case conversion may be inaccurate. Consider using '#align exists_prop exists_propₓ'. -/
@[simp]
theorem exists_prop {p q : Prop} : (∃ h : p, q) ↔ p ∧ q :=
⟨fun ⟨h₁, h₂⟩ => ⟨h₁, h₂⟩, fun ⟨h₁, h₂⟩ => ⟨h₁, h₂⟩⟩
@@ -2839,12 +2455,6 @@ theorem forall_true_left (p : True → Prop) : (∀ x, p x) ↔ p True.intro :=
#align forall_true_left forall_true_left
-/
-/- warning: exists_unique.elim2 -> ExistsUnique.elim₂ is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {p : α -> Sort.{u2}} [_inst_1 : forall (x : α), Subsingleton.{u2} (p x)] {q : forall (x : α), (p x) -> Prop} {b : Prop}, (ExistsUnique.{u1} α (fun (x : α) => ExistsUnique.{u2} (p x) (fun (h : p x) => q x h))) -> (forall (x : α) (h : p x), (q x h) -> (forall (y : α) (hy : p y), (q y hy) -> (Eq.{u1} α y x)) -> b) -> b
-but is expected to have type
- forall {α : Sort.{u2}} {p : α -> Sort.{u1}} [_inst_1 : forall (x : α), Subsingleton.{u1} (p x)] {q : forall (x : α), (p x) -> Prop} {b : Prop}, (ExistsUnique.{u2} α (fun (x : α) => ExistsUnique.{u1} (p x) (fun (h : p x) => q x h))) -> (forall (x : α) (h : p x), (q x h) -> (forall (y : α) (hy : p y), (q y hy) -> (Eq.{u2} α y x)) -> b) -> b
-Case conversion may be inaccurate. Consider using '#align exists_unique.elim2 ExistsUnique.elim₂ₓ'. -/
theorem ExistsUnique.elim₂ {α : Sort _} {p : α → Sort _} [∀ x, Subsingleton (p x)]
{q : ∀ (x) (h : p x), Prop} {b : Prop} (h₂ : ∃! (x : _)(h : p x), q x h)
(h₁ : ∀ (x) (h : p x), q x h → (∀ (y) (hy : p y), q y hy → y = x) → b) : b :=
@@ -2854,12 +2464,6 @@ theorem ExistsUnique.elim₂ {α : Sort _} {p : α → Sort _} [∀ x, Subsingle
exact fun x ⟨hxp, hxq⟩ H => h₁ x hxp hxq fun y hyp hyq => H y ⟨hyp, hyq⟩
#align exists_unique.elim2 ExistsUnique.elim₂
-/- warning: exists_unique.intro2 -> ExistsUnique.intro₂ is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {p : α -> Sort.{u2}} [_inst_1 : forall (x : α), Subsingleton.{u2} (p x)] {q : forall (x : α), (p x) -> Prop} (w : α) (hp : p w), (q w hp) -> (forall (y : α) (hy : p y), (q y hy) -> (Eq.{u1} α y w)) -> (ExistsUnique.{u1} α (fun (x : α) => ExistsUnique.{u2} (p x) (fun (hx : p x) => q x hx)))
-but is expected to have type
- forall {α : Sort.{u2}} {p : α -> Sort.{u1}} [_inst_1 : forall (x : α), Subsingleton.{u1} (p x)] {q : forall (x : α), (p x) -> Prop} (w : α) (hp : p w), (q w hp) -> (forall (y : α) (hy : p y), (q y hy) -> (Eq.{u2} α y w)) -> (ExistsUnique.{u2} α (fun (x : α) => ExistsUnique.{u1} (p x) (fun (hx : p x) => q x hx)))
-Case conversion may be inaccurate. Consider using '#align exists_unique.intro2 ExistsUnique.intro₂ₓ'. -/
theorem ExistsUnique.intro₂ {α : Sort _} {p : α → Sort _} [∀ x, Subsingleton (p x)]
{q : ∀ (x : α) (h : p x), Prop} (w : α) (hp : p w) (hq : q w hp)
(H : ∀ (y) (hy : p y), q y hy → y = w) : ∃! (x : _)(hx : p x), q x hx :=
@@ -2868,23 +2472,11 @@ theorem ExistsUnique.intro₂ {α : Sort _} {p : α → Sort _} [∀ x, Subsingl
exact ExistsUnique.intro w ⟨hp, hq⟩ fun y ⟨hyp, hyq⟩ => H y hyp hyq
#align exists_unique.intro2 ExistsUnique.intro₂
-/- warning: exists_unique.exists2 -> ExistsUnique.exists₂ is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {p : α -> Sort.{u2}} {q : forall (x : α), (p x) -> Prop}, (ExistsUnique.{u1} α (fun (x : α) => ExistsUnique.{u2} (p x) (fun (hx : p x) => q x hx))) -> (Exists.{u1} α (fun (x : α) => Exists.{u2} (p x) (fun (hx : p x) => q x hx)))
-but is expected to have type
- forall {α : Sort.{u2}} {p : α -> Sort.{u1}} {q : forall (x : α), (p x) -> Prop}, (ExistsUnique.{u2} α (fun (x : α) => ExistsUnique.{u1} (p x) (fun (hx : p x) => q x hx))) -> (Exists.{u2} α (fun (x : α) => Exists.{u1} (p x) (fun (hx : p x) => q x hx)))
-Case conversion may be inaccurate. Consider using '#align exists_unique.exists2 ExistsUnique.exists₂ₓ'. -/
theorem ExistsUnique.exists₂ {α : Sort _} {p : α → Sort _} {q : ∀ (x : α) (h : p x), Prop}
(h : ∃! (x : _)(hx : p x), q x hx) : ∃ (x : _)(hx : p x), q x hx :=
h.exists.imp fun x hx => hx.exists
#align exists_unique.exists2 ExistsUnique.exists₂
-/- warning: exists_unique.unique2 -> ExistsUnique.unique₂ is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {p : α -> Sort.{u2}} [_inst_1 : forall (x : α), Subsingleton.{u2} (p x)] {q : forall (x : α), (p x) -> Prop}, (ExistsUnique.{u1} α (fun (x : α) => ExistsUnique.{u2} (p x) (fun (hx : p x) => q x hx))) -> (forall {y₁ : α} {y₂ : α} (hpy₁ : p y₁), (q y₁ hpy₁) -> (forall (hpy₂ : p y₂), (q y₂ hpy₂) -> (Eq.{u1} α y₁ y₂)))
-but is expected to have type
- forall {α : Sort.{u2}} {p : α -> Sort.{u1}} [_inst_1 : forall (x : α), Subsingleton.{u1} (p x)] {q : forall (x : α), (p x) -> Prop}, (ExistsUnique.{u2} α (fun (x : α) => ExistsUnique.{u1} (p x) (fun (hx : p x) => q x hx))) -> (forall {y₁ : α} {y₂ : α} (hpy₁ : p y₁), (q y₁ hpy₁) -> (forall (hpy₂ : p y₂), (q y₂ hpy₂) -> (Eq.{u2} α y₁ y₂)))
-Case conversion may be inaccurate. Consider using '#align exists_unique.unique2 ExistsUnique.unique₂ₓ'. -/
theorem ExistsUnique.unique₂ {α : Sort _} {p : α → Sort _} [∀ x, Subsingleton (p x)]
{q : ∀ (x : α) (hx : p x), Prop} (h : ∃! (x : _)(hx : p x), q x hx) {y₁ y₂ : α} (hpy₁ : p y₁)
(hqy₁ : q y₁ hpy₁) (hpy₂ : p y₂) (hqy₂ : q y₂ hpy₂) : y₁ = y₂ :=
@@ -2932,12 +2524,6 @@ noncomputable def decEq (α : Sort _) : DecidableEq α := by infer_instance
#align classical.dec_eq Classical.decEq
-/
-/- warning: classical.exists_cases -> Classical.existsCases is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u_1}} {p : α -> Prop} {C : Sort.{u}}, C -> (forall (a : α), (p a) -> C) -> C
-but is expected to have type
- forall {α : Prop} {p : α -> Prop} {C : Sort.{u_1}}, C -> (forall (a : α), (p a) -> C) -> C
-Case conversion may be inaccurate. Consider using '#align classical.exists_cases Classical.existsCasesₓ'. -/
/-- Construct a function from a default value `H0`, and a function to use if there exists a value
satisfying the predicate. -/
@[elab_as_elim]
@@ -2975,12 +2561,6 @@ def choice_of_byContradiction' {α : Sort _} (contra : ¬(α → False) → α)
end Classical
-/- warning: exists.classical_rec_on -> Exists.classicalRecOn is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u2}} {p : α -> Prop}, (Exists.{u2} α (fun (a : α) => p a)) -> (forall {C : Sort.{u1}}, (forall (a : α), (p a) -> C) -> C)
-but is expected to have type
- forall {α : Sort.{u1}} {p : α -> Prop}, (Exists.{u1} α (fun (a : α) => p a)) -> (forall {C : Sort.{u2}}, (forall (a : α), (p a) -> C) -> C)
-Case conversion may be inaccurate. Consider using '#align exists.classical_rec_on Exists.classicalRecOnₓ'. -/
/-- This function has the same type as `exists.rec_on`, and can be used to case on an equality,
but `exists.rec_on` can only eliminate into Prop, while this version eliminates into any universe
using the axiom of choice. -/
@@ -3314,34 +2894,16 @@ theorem ite_eq_or_eq : ite P a b = a ∨ ite P a b = b :=
#align ite_eq_or_eq ite_eq_or_eq
-/
-/- warning: apply_dite -> apply_dite is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {β : Sort.{u2}} (f : α -> β) (P : Prop) [_inst_1 : Decidable P] (x : P -> α) (y : (Not P) -> α), Eq.{u2} β (f (dite.{u1} α P _inst_1 x y)) (dite.{u2} β P _inst_1 (fun (h : P) => f (x h)) (fun (h : Not P) => f (y h)))
-but is expected to have type
- forall {α : Sort.{u2}} {β : Sort.{u1}} (f : α -> β) (P : Prop) [_inst_1 : Decidable P] (x : P -> α) (y : (Not P) -> α), Eq.{u1} β (f (dite.{u2} α P _inst_1 x y)) (dite.{u1} β P _inst_1 (fun (h : P) => f (x h)) (fun (h : Not P) => f (y h)))
-Case conversion may be inaccurate. Consider using '#align apply_dite apply_diteₓ'. -/
/-- A function applied to a `dite` is a `dite` of that function applied to each of the branches. -/
theorem apply_dite (x : P → α) (y : ¬P → α) :
f (dite P x y) = dite P (fun h => f (x h)) fun h => f (y h) := by by_cases h : P <;> simp [h]
#align apply_dite apply_dite
-/- warning: apply_ite -> apply_ite is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {β : Sort.{u2}} (f : α -> β) (P : Prop) [_inst_1 : Decidable P] (a : α) (b : α), Eq.{u2} β (f (ite.{u1} α P _inst_1 a b)) (ite.{u2} β P _inst_1 (f a) (f b))
-but is expected to have type
- forall {α : Sort.{u2}} {β : Sort.{u1}} (f : α -> β) (P : Prop) [_inst_1 : Decidable P] (a : α) (b : α), Eq.{u1} β (f (ite.{u2} α P _inst_1 a b)) (ite.{u1} β P _inst_1 (f a) (f b))
-Case conversion may be inaccurate. Consider using '#align apply_ite apply_iteₓ'. -/
/-- A function applied to a `ite` is a `ite` of that function applied to each of the branches. -/
theorem apply_ite : f (ite P a b) = ite P (f a) (f b) :=
apply_dite f P (fun _ => a) fun _ => b
#align apply_ite apply_ite
-/- warning: apply_dite2 -> apply_dite₂ is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {β : Sort.{u2}} {γ : Sort.{u3}} (f : α -> β -> γ) (P : Prop) [_inst_3 : Decidable P] (a : P -> α) (b : (Not P) -> α) (c : P -> β) (d : (Not P) -> β), Eq.{u3} γ (f (dite.{u1} α P _inst_3 a b) (dite.{u2} β P _inst_3 c d)) (dite.{u3} γ P _inst_3 (fun (h : P) => f (a h) (c h)) (fun (h : Not P) => f (b h) (d h)))
-but is expected to have type
- forall {α : Sort.{u2}} {β : Sort.{u1}} {γ : Sort.{u3}} (f : α -> β -> γ) (P : Prop) [_inst_3 : Decidable P] (a : P -> α) (b : (Not P) -> α) (c : P -> β) (d : (Not P) -> β), Eq.{u3} γ (f (dite.{u2} α P _inst_3 a b) (dite.{u1} β P _inst_3 c d)) (dite.{u3} γ P _inst_3 (fun (h : P) => f (a h) (c h)) (fun (h : Not P) => f (b h) (d h)))
-Case conversion may be inaccurate. Consider using '#align apply_dite2 apply_dite₂ₓ'. -/
/-- A two-argument function applied to two `dite`s is a `dite` of that two-argument function
applied to each of the branches. -/
theorem apply_dite₂ (f : α → β → γ) (P : Prop) [Decidable P] (a : P → α) (b : ¬P → α) (c : P → β)
@@ -3350,12 +2912,6 @@ theorem apply_dite₂ (f : α → β → γ) (P : Prop) [Decidable P] (a : P →
by_cases h : P <;> simp [h]
#align apply_dite2 apply_dite₂
-/- warning: apply_ite2 -> apply_ite₂ is a dubious translation:
-lean 3 declaration is
- forall {α : Sort.{u1}} {β : Sort.{u2}} {γ : Sort.{u3}} (f : α -> β -> γ) (P : Prop) [_inst_3 : Decidable P] (a : α) (b : α) (c : β) (d : β), Eq.{u3} γ (f (ite.{u1} α P _inst_3 a b) (ite.{u2} β P _inst_3 c d)) (ite.{u3} γ P _inst_3 (f a c) (f b d))
-but is expected to have type
- forall {α : Sort.{u2}} {β : Sort.{u1}} {γ : Sort.{u3}} (f : α -> β -> γ) (P : Prop) [_inst_3 : Decidable P] (a : α) (b : α) (c : β) (d : β), Eq.{u3} γ (f (ite.{u2} α P _inst_3 a b) (ite.{u1} β P _inst_3 c d)) (ite.{u3} γ P _inst_3 (f a c) (f b d))
-Case conversion may be inaccurate. Consider using '#align apply_ite2 apply_ite₂ₓ'. -/
/-- A two-argument function applied to two `ite`s is a `ite` of that two-argument function
applied to each of the branches. -/
theorem apply_ite₂ (f : α → β → γ) (P : Prop) [Decidable P] (a b : α) (c d : β) :
mathlib commit https://github.com/leanprover-community/mathlib/commit/917c3c072e487b3cccdbfeff17e75b40e45f66cb
@@ -66,11 +66,7 @@ instance : Subsingleton Empty :=
instance Subsingleton.prod {α β : Type _} [Subsingleton α] [Subsingleton β] :
Subsingleton (α × β) :=
- ⟨by
- intro a b
- cases a
- cases b
- congr ⟩
+ ⟨by intro a b; cases a; cases b; congr ⟩
#align subsingleton.prod Subsingleton.prod
instance : DecidableEq Empty := fun a => a.elim
@@ -227,10 +223,7 @@ but is expected to have type
forall {α : Sort.{u2}} {β : Sort.{u2}} {γ : Sort.{u1}} {f : α -> γ} {g : β -> γ} {x : α} {y : β}, (HEq.{imax u2 u1} (α -> γ) f (β -> γ) g) -> (HEq.{u2} α x β y) -> (Eq.{u1} γ (f x) (g y))
Case conversion may be inaccurate. Consider using '#align congr_heq congr_heqₓ'. -/
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
+ (h₂ : HEq x y) : f x = g y := by cases h₂; cases h₁; rfl
#align congr_heq congr_heq
/- warning: congr_arg_heq -> congr_arg_heq is a dubious translation:
@@ -434,8 +427,7 @@ theorem Iff.imp (h₁ : a ↔ b) (h₂ : c ↔ d) : a → c ↔ b → d :=
#print eq_true_eq_id /-
@[simp]
-theorem eq_true_eq_id : Eq True = id := by
- funext
+theorem eq_true_eq_id : Eq True = id := by funext;
simp only [true_iff_iff, id.def, iff_self_iff, eq_iff_iff]
#align eq_true_eq_id eq_true_eq_id
-/
@@ -735,9 +727,7 @@ theorem imp_not_self : a → ¬a ↔ ¬a :=
-/
#print Decidable.not_imp_self /-
-theorem Decidable.not_imp_self [Decidable a] : ¬a → a ↔ a :=
- by
- have := @imp_not_self ¬a
+theorem Decidable.not_imp_self [Decidable a] : ¬a → a ↔ a := by have := @imp_not_self ¬a;
rwa [Decidable.not_not] at this
#align decidable.not_imp_self Decidable.not_imp_self
-/
@@ -1219,10 +1209,8 @@ but is expected to have type
forall {a : Prop} {b : Prop} {c : Prop} [_inst_1 : Decidable a], ((Not a) -> (Iff b c)) -> (Iff (Or b a) (Or c a))
Case conversion may be inaccurate. Consider using '#align decidable.or_congr_left Decidable.or_congr_left'ₓ'. -/
-- See Note [decidable namespace]
-protected theorem Decidable.or_congr_left' [Decidable c] (h : ¬c → (a ↔ b)) : a ∨ c ↔ b ∨ c :=
- by
- rw [Decidable.or_iff_not_imp_right, Decidable.or_iff_not_imp_right]
- exact imp_congr_right h
+protected theorem Decidable.or_congr_left' [Decidable c] (h : ¬c → (a ↔ b)) : a ∨ c ↔ b ∨ c := by
+ rw [Decidable.or_iff_not_imp_right, Decidable.or_iff_not_imp_right]; exact imp_congr_right h
#align decidable.or_congr_left Decidable.or_congr_left'
/- warning: or_congr_left -> or_congr_left' is a dubious translation:
@@ -1237,10 +1225,8 @@ theorem or_congr_left' (h : ¬c → (a ↔ b)) : a ∨ c ↔ b ∨ c :=
#print Decidable.or_congr_right' /-
-- See Note [decidable namespace]
-protected theorem Decidable.or_congr_right' [Decidable a] (h : ¬a → (b ↔ c)) : a ∨ b ↔ a ∨ c :=
- by
- rw [Decidable.or_iff_not_imp_left, Decidable.or_iff_not_imp_left]
- exact imp_congr_right h
+protected theorem Decidable.or_congr_right' [Decidable a] (h : ¬a → (b ↔ c)) : a ∨ b ↔ a ∨ c := by
+ rw [Decidable.or_iff_not_imp_left, Decidable.or_iff_not_imp_left]; exact imp_congr_right h
#align decidable.or_congr_right Decidable.or_congr_right'
-/
@@ -1803,9 +1789,7 @@ Case conversion may be inaccurate. Consider using '#align eq_rec_constant eq_rec
/-- Transport through trivial families is the identity. -/
@[simp]
theorem eq_rec_constant {α : Sort _} {a a' : α} {β : Sort _} (y : β) (h : a = a') :
- @Eq.ndrec α a (fun a => β) y a' h = y := by
- cases h
- rfl
+ @Eq.ndrec α a (fun a => β) y a' h = y := by cases h; rfl
#align eq_rec_constant eq_rec_constant
#print eq_mp_eq_cast /-
@@ -1943,10 +1927,8 @@ lean 3 declaration is
but is expected to have type
forall {α : Sort.{u1}} {x₁ : α} {x₂ : α} {y₁ : α} {y₂ : α}, (Eq.{u1} α x₁ x₂) -> (Eq.{u1} α y₁ y₂) -> (Iff (Eq.{u1} α x₁ y₁) (Eq.{u1} α x₂ y₂))
Case conversion may be inaccurate. Consider using '#align eq.congr Eq.congrₓ'. -/
-protected theorem Eq.congr {x₁ x₂ y₁ y₂ : α} (h₁ : x₁ = y₁) (h₂ : x₂ = y₂) : x₁ = x₂ ↔ y₁ = y₂ :=
- by
- subst h₁
- subst h₂
+protected theorem Eq.congr {x₁ x₂ y₁ y₂ : α} (h₁ : x₁ = y₁) (h₂ : x₂ = y₂) : x₁ = x₂ ↔ y₁ = y₂ := by
+ subst h₁; subst h₂
#align eq.congr Eq.congr
#print Eq.congr_left /-
@@ -1966,9 +1948,7 @@ but is expected to have type
forall {α : Sort.{u3}} {β : Sort.{u2}} {γ : Sort.{u1}} (f : α -> β -> γ) {x : α} {x' : α} {y : β} {y' : β}, (Eq.{u3} α x x') -> (Eq.{u2} β y y') -> (Eq.{u1} γ (f x y) (f x' y'))
Case conversion may be inaccurate. Consider using '#align congr_arg2 congr_arg₂ₓ'. -/
theorem congr_arg₂ {α β γ : Sort _} (f : α → β → γ) {x x' : α} {y y' : β} (hx : x = x')
- (hy : y = y') : f x y = f x' y' := by
- subst hx
- subst hy
+ (hy : y = y') : f x y = f x' y' := by subst hx; subst hy
#align congr_arg2 congr_arg₂
variable {β : α → Sort _} {γ : ∀ a, β a → Sort _} {δ : ∀ a b, γ a b → Sort _}
@@ -2619,10 +2599,7 @@ but is expected to have type
Case conversion may be inaccurate. Consider using '#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
- rw [forall_swap]
- simp
+ (∀ b, ∀ a, f a = b → p b) ↔ ∀ a, p (f a) := by rw [forall_swap]; simp
#align forall_apply_eq_imp_iff' forall_apply_eq_imp_iff'
/- warning: forall_eq_apply_imp_iff -> forall_eq_apply_imp_iff is a dubious translation:
@@ -2644,10 +2621,7 @@ but is expected to have type
Case conversion may be inaccurate. Consider using '#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
- rw [forall_swap]
- simp
+ (∀ b, ∀ a, b = f a → p b) ↔ ∀ a, p (f a) := by rw [forall_swap]; simp
#align forall_eq_apply_imp_iff' forall_eq_apply_imp_iff'
/- warning: forall_apply_eq_imp_iff₂ -> forall_apply_eq_imp_iff₂ is a dubious translation:
@@ -3247,10 +3221,8 @@ theorem ite_eq_right_iff : ite P a b = b ↔ P → a = b :=
-/
#print dite_ne_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]
+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
-/
@@ -3434,11 +3406,7 @@ 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) =
if q : Q then B q else if p : P then A p else C 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]⟩
+ ⟨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
-/
mathlib commit https://github.com/leanprover-community/mathlib/commit/8d33f09cd7089ecf074b4791907588245aec5d1b
@@ -1508,7 +1508,7 @@ Case conversion may be inaccurate. Consider using '#align decidable.iff_iff_and_
protected theorem Decidable.iff_iff_and_or_not_and_not [Decidable b] : (a ↔ b) ↔ a ∧ b ∨ ¬a ∧ ¬b :=
by
constructor <;> intro h
- · rw [h] <;> by_cases b <;> [left, right] <;> constructor <;> assumption
+ · rw [h] <;> by_cases b <;> [left;right] <;> constructor <;> assumption
· cases' h with h h <;> cases h <;> constructor <;> intro <;> · first |contradiction|assumption
#align decidable.iff_iff_and_or_not_and_not Decidable.iff_iff_and_or_not_and_not
@@ -1798,7 +1798,7 @@ theorem eq_equivalence : Equivalence (@Eq α) :=
lean 3 declaration is
forall {α : Sort.{u1}} {a : α} {a' : α} {β : Sort.{u2}} (y : β) (h : Eq.{u1} α a a'), Eq.{u2} ((fun (a : α) => β) a') (Eq.ndrec.{u2, u1} α a (fun (a : α) => β) y a' h) y
but is expected to have type
- forall {α : Sort.{u2}} {a : α} {a' : α} {β : Sort.{u1}} (y : β) (h : Eq.{u2} α a a'), Eq.{u1} β (Eq.rec.{u1, u2} α a (fun (a_1 : α) (x._@.Std.Logic._hyg.12313 : Eq.{u2} α a a_1) => β) y a' h) y
+ forall {α : Sort.{u2}} {a : α} {a' : α} {β : Sort.{u1}} (y : β) (h : Eq.{u2} α a a'), Eq.{u1} β (Eq.rec.{u1, u2} α a (fun (a_1 : α) (x._@.Std.Logic._hyg.12340 : Eq.{u2} α a a_1) => β) y a' h) y
Case conversion may be inaccurate. Consider using '#align eq_rec_constant eq_rec_constantₓ'. -/
/-- Transport through trivial families is the identity. -/
@[simp]
mathlib commit https://github.com/leanprover-community/mathlib/commit/8b8ba04e2f326f3f7cf24ad129beda58531ada61
@@ -38,7 +38,7 @@ section Miscellany
/- We add the `inline` attribute to optimize VM computation using these declarations. For example,
`if p ∧ q then ... else ...` will not evaluate the decidability of `q` if `p` is false. -/
attribute [inline]
- And.decidable Or.decidable Decidable.false Xor'.decidable Iff.decidable Decidable.true Implies.decidable Not.decidable Ne.decidable Bool.decidableEq Decidable.decide
+ And.decidable Or.decidable decidableFalse Xor'.decidable Iff.decidable decidableTrue Implies.decidable Not.decidable Ne.decidable Bool.decidableEq Decidable.decide
attribute [simp] cast_eq cast_hEq
mathlib commit https://github.com/leanprover-community/mathlib/commit/09079525fd01b3dda35e96adaa08d2f943e1648c
@@ -31,7 +31,7 @@ maybe it is useful for writing automation.
open Function
-attribute [local instance] Classical.propDecidable
+attribute [local instance 10] Classical.propDecidable
section Miscellany
@@ -156,7 +156,7 @@ theorem coeFn_coe_base' {α β} {γ : _} [Coe α β] [CoeFun β fun _ => γ] (x
-- This instance should have low priority, to ensure we follow the chain
-- `set_like → has_coe_to_sort`
-attribute [instance] coeSortTrans
+attribute [instance 10] coeSortTrans
theorem coeSort_coe_trans {α β γ δ} [Coe α β] [HasCoeTAux β γ] [CoeSort γ δ] (x : α) :
@coeSort α _ _ x = @coeSort β _ _ x :=
@@ -302,7 +302,7 @@ theorem eq_iff_eq_cancel_right {α : Sort _} {a b : α} : (∀ {c}, a = c ↔ b
-/
#print Fact /-
-/- ./././Mathport/Syntax/Translate/Command.lean:388:30: infer kinds are unsupported in Lean 4: #[`out] [] -/
+/- ./././Mathport/Syntax/Translate/Command.lean:393:30: infer kinds are unsupported in Lean 4: #[`out] [] -/
/-- Wrapper for adding elementary propositions to the type class systems.
Warning: this can easily be abused. See the rest of this docstring for details.
mathlib commit https://github.com/leanprover-community/mathlib/commit/02ba8949f486ebecf93fe7460f1ed0564b5e442c
@@ -1715,27 +1715,27 @@ theorem ne_of_mem_of_not_mem' (h : a ∈ s) : a ∉ t → s ≠ t :=
mt fun e => e ▸ h
#align ne_of_mem_of_not_mem' ne_of_mem_of_not_mem'
-/- warning: has_mem.mem.ne_of_not_mem -> Membership.Mem.ne_of_not_mem is a dubious translation:
+/- warning: has_mem.mem.ne_of_not_mem -> Membership.mem.ne_of_not_mem is a dubious translation:
lean 3 declaration is
forall {α : Type.{u1}} {β : Type.{u2}} [_inst_1 : Membership.{u1, u2} α β] {s : β} {a : α} {b : α}, (Membership.Mem.{u1, u2} α β _inst_1 a s) -> (Not (Membership.Mem.{u1, u2} α β _inst_1 b s)) -> (Ne.{succ u1} α a b)
but is expected to have type
forall {α : Type.{u2}} {β : Type.{u1}} [_inst_1 : Membership.{u2, u1} α β] {s : β} {a : α} {b : α}, (Membership.mem.{u2, u1} α β _inst_1 a s) -> (Not (Membership.mem.{u2, u1} α β _inst_1 b s)) -> (Ne.{succ u2} α a b)
-Case conversion may be inaccurate. Consider using '#align has_mem.mem.ne_of_not_mem Membership.Mem.ne_of_not_memₓ'. -/
+Case conversion may be inaccurate. Consider using '#align has_mem.mem.ne_of_not_mem Membership.mem.ne_of_not_memₓ'. -/
/-- **Alias** of `ne_of_mem_of_not_mem`. -/
-theorem Membership.Mem.ne_of_not_mem : a ∈ s → b ∉ s → a ≠ b :=
+theorem Membership.mem.ne_of_not_mem : a ∈ s → b ∉ s → a ≠ b :=
ne_of_mem_of_not_mem
-#align has_mem.mem.ne_of_not_mem Membership.Mem.ne_of_not_mem
+#align has_mem.mem.ne_of_not_mem Membership.mem.ne_of_not_mem
-/- warning: has_mem.mem.ne_of_not_mem' -> Membership.Mem.ne_of_not_mem' is a dubious translation:
+/- warning: has_mem.mem.ne_of_not_mem' -> Membership.mem.ne_of_not_mem' is a dubious translation:
lean 3 declaration is
forall {α : Type.{u1}} {β : Type.{u2}} [_inst_1 : Membership.{u1, u2} α β] {s : β} {t : β} {a : α}, (Membership.Mem.{u1, u2} α β _inst_1 a s) -> (Not (Membership.Mem.{u1, u2} α β _inst_1 a t)) -> (Ne.{succ u2} β s t)
but is expected to have type
forall {α : Type.{u2}} {β : Type.{u1}} [_inst_1 : Membership.{u2, u1} α β] {s : β} {t : β} {a : α}, (Membership.mem.{u2, u1} α β _inst_1 a s) -> (Not (Membership.mem.{u2, u1} α β _inst_1 a t)) -> (Ne.{succ u1} β s t)
-Case conversion may be inaccurate. Consider using '#align has_mem.mem.ne_of_not_mem' Membership.Mem.ne_of_not_mem'ₓ'. -/
+Case conversion may be inaccurate. Consider using '#align has_mem.mem.ne_of_not_mem' Membership.mem.ne_of_not_mem'ₓ'. -/
/-- **Alias** of `ne_of_mem_of_not_mem'`. -/
-theorem Membership.Mem.ne_of_not_mem' : a ∈ s → a ∉ t → s ≠ t :=
+theorem Membership.mem.ne_of_not_mem' : a ∈ s → a ∉ t → s ≠ t :=
ne_of_mem_of_not_mem'
-#align has_mem.mem.ne_of_not_mem' Membership.Mem.ne_of_not_mem'
+#align has_mem.mem.ne_of_not_mem' Membership.mem.ne_of_not_mem'
end Mem
@@ -1911,7 +1911,7 @@ theorem cast_eq_iff_heq {α β : Sort _} {a : α} {a' : β} {e : α = β} : cast
lean 3 declaration is
forall {α : Sort.{u1}} {a : α} {b : α} {β : Sort.{u2}} {C : α -> Sort.{u2}} {x : C a} {y : β} (e : Eq.{u1} α a b), (HEq.{u2} (C a) x β y) -> (HEq.{u2} (C b) (Eq.ndrec.{u2, u1} α a C x b e) β y)
but is expected to have type
- forall {α : Sort.{u2}} {a : α} {b : Sort.{u1}} {β : α} {C : α -> Sort.{u1}} {x : C a} {y : b} (e : Eq.{u2} α a β), (HEq.{u1} (C a) x b y) -> (HEq.{u1} (C β) (Eq.rec.{u1, u2} α a (fun (x._@.Mathlib.Logic.Basic._hyg.4253 : α) (h._@.Mathlib.Logic.Basic._hyg.4254 : Eq.{u2} α a x._@.Mathlib.Logic.Basic._hyg.4253) => C x._@.Mathlib.Logic.Basic._hyg.4253) x β e) b y)
+ forall {α : Sort.{u2}} {a : α} {b : Sort.{u1}} {β : α} {C : α -> Sort.{u1}} {x : C a} {y : b} (e : Eq.{u2} α a β), (HEq.{u1} (C a) x b y) -> (HEq.{u1} (C β) (Eq.rec.{u1, u2} α a (fun (x._@.Mathlib.Logic.Basic._hyg.4255 : α) (h._@.Mathlib.Logic.Basic._hyg.4256 : Eq.{u2} α a x._@.Mathlib.Logic.Basic._hyg.4255) => C x._@.Mathlib.Logic.Basic._hyg.4255) x β e) b y)
Case conversion may be inaccurate. Consider using '#align rec_heq_of_heq rec_heq_of_heqₓ'. -/
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
@@ -1921,7 +1921,7 @@ theorem rec_heq_of_heq {β} {C : α → Sort _} {x : C a} {y : β} (e : a = b) (
lean 3 declaration is
forall {α : Sort.{u1}} {a : α} {b : α} {β : Sort.{u2}} {C : α -> Sort.{u2}} {x : C a} {y : β} {e : Eq.{u1} α a b}, Iff (HEq.{u2} (C b) (Eq.ndrec.{u2, u1} α a C x b e) β y) (HEq.{u2} (C a) x β y)
but is expected to have type
- forall {α : Sort.{u2}} {a : α} {b : Sort.{u1}} {β : α} {C : α -> Sort.{u1}} {x : C a} {y : b} {e : Eq.{u2} α a β}, Iff (HEq.{u1} (C β) (Eq.rec.{u1, u2} α a (fun (x._@.Mathlib.Logic.Basic._hyg.4313 : α) (h._@.Mathlib.Logic.Basic._hyg.4314 : Eq.{u2} α a x._@.Mathlib.Logic.Basic._hyg.4313) => C x._@.Mathlib.Logic.Basic._hyg.4313) x β e) b y) (HEq.{u1} (C a) x b y)
+ forall {α : Sort.{u2}} {a : α} {b : Sort.{u1}} {β : α} {C : α -> Sort.{u1}} {x : C a} {y : b} {e : Eq.{u2} α a β}, Iff (HEq.{u1} (C β) (Eq.rec.{u1, u2} α a (fun (x._@.Mathlib.Logic.Basic._hyg.4315 : α) (h._@.Mathlib.Logic.Basic._hyg.4316 : Eq.{u2} α a x._@.Mathlib.Logic.Basic._hyg.4315) => C x._@.Mathlib.Logic.Basic._hyg.4315) x β e) b y) (HEq.{u1} (C a) x b y)
Case conversion may be inaccurate. Consider using '#align rec_heq_iff_heq rec_heq_iff_heqₓ'. -/
theorem rec_heq_iff_heq {β} {C : α → Sort _} {x : C a} {y : β} {e : a = b} :
HEq (@Eq.ndrec α a C x b e) y ↔ HEq x y := by subst e
@@ -1931,7 +1931,7 @@ theorem rec_heq_iff_heq {β} {C : α → Sort _} {x : C a} {y : β} {e : a = b}
lean 3 declaration is
forall {α : Sort.{u1}} {a : α} {b : α} {β : Sort.{u2}} {C : α -> Sort.{u2}} {x : β} {y : C a} {e : Eq.{u1} α a b}, Iff (HEq.{u2} β x (C b) (Eq.ndrec.{u2, u1} α a C y b e)) (HEq.{u2} β x (C a) y)
but is expected to have type
- forall {α : Sort.{u2}} {a : Sort.{u1}} {b : α} {β : α} {C : α -> Sort.{u1}} {x : a} {y : C b} {e : Eq.{u2} α b β}, Iff (HEq.{u1} a x (C β) (Eq.rec.{u1, u2} α b (fun (x._@.Mathlib.Logic.Basic._hyg.4375 : α) (h._@.Mathlib.Logic.Basic._hyg.4376 : Eq.{u2} α b x._@.Mathlib.Logic.Basic._hyg.4375) => C x._@.Mathlib.Logic.Basic._hyg.4375) y β e)) (HEq.{u1} a x (C b) y)
+ forall {α : Sort.{u2}} {a : Sort.{u1}} {b : α} {β : α} {C : α -> Sort.{u1}} {x : a} {y : C b} {e : Eq.{u2} α b β}, Iff (HEq.{u1} a x (C β) (Eq.rec.{u1, u2} α b (fun (x._@.Mathlib.Logic.Basic._hyg.4377 : α) (h._@.Mathlib.Logic.Basic._hyg.4378 : Eq.{u2} α b x._@.Mathlib.Logic.Basic._hyg.4377) => C x._@.Mathlib.Logic.Basic._hyg.4377) y β e)) (HEq.{u1} a x (C b) y)
Case conversion may be inaccurate. Consider using '#align heq_rec_iff_heq heq_rec_iff_heqₓ'. -/
theorem heq_rec_iff_heq {β} {C : α → Sort _} {x : β} {y : C a} {e : a = b} :
HEq x (@Eq.ndrec α a C y b e) ↔ HEq x y := by subst e
mathlib commit https://github.com/leanprover-community/mathlib/commit/3180fab693e2cee3bff62675571264cb8778b212
@@ -1798,7 +1798,7 @@ theorem eq_equivalence : Equivalence (@Eq α) :=
lean 3 declaration is
forall {α : Sort.{u1}} {a : α} {a' : α} {β : Sort.{u2}} (y : β) (h : Eq.{u1} α a a'), Eq.{u2} ((fun (a : α) => β) a') (Eq.ndrec.{u2, u1} α a (fun (a : α) => β) y a' h) y
but is expected to have type
- forall {α : Sort.{u2}} {a : α} {a' : α} {β : Sort.{u1}} (y : β) (h : Eq.{u2} α a a'), Eq.{u1} β (Eq.rec.{u1, u2} α a (fun (a_1 : α) (x._@.Std.Logic._hyg.12312 : Eq.{u2} α a a_1) => β) y a' h) y
+ forall {α : Sort.{u2}} {a : α} {a' : α} {β : Sort.{u1}} (y : β) (h : Eq.{u2} α a a'), Eq.{u1} β (Eq.rec.{u1, u2} α a (fun (a_1 : α) (x._@.Std.Logic._hyg.12313 : Eq.{u2} α a a_1) => β) y a' h) y
Case conversion may be inaccurate. Consider using '#align eq_rec_constant eq_rec_constantₓ'. -/
/-- Transport through trivial families is the identity. -/
@[simp]
mathlib commit https://github.com/leanprover-community/mathlib/commit/641b6a82006416ec431b2987b354af9311fed4f2
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Jeremy Avigad, Leonardo de Moura
! This file was ported from Lean 3 source module logic.basic
-! leanprover-community/mathlib commit 13cd3e89b30352d5b1b7349f5537ea18ba878e40
+! leanprover-community/mathlib commit d2d8742b0c21426362a9dacebc6005db895ca963
! Please do not edit these lines, except to modify the commit id
! if you have ported upstream changes.
-/
@@ -287,27 +287,19 @@ theorem ne_comm {α} {a b : α} : a ≠ b ↔ b ≠ a :=
#align ne_comm ne_comm
-/
-/- warning: eq_iff_eq_cancel_left -> eq_iff_eq_cancel_left is a dubious translation:
-lean 3 declaration is
- forall {α : Type.{u1}} {b : α} {c : α}, Iff (forall {a : α}, Iff (Eq.{succ u1} α a b) (Eq.{succ u1} α a c)) (Eq.{succ u1} α b c)
-but is expected to have type
- forall {α : Sort.{u1}} {b : α} {c : α}, Iff (forall {a : α}, Iff (Eq.{u1} α a b) (Eq.{u1} α a c)) (Eq.{u1} α b c)
-Case conversion may be inaccurate. Consider using '#align eq_iff_eq_cancel_left eq_iff_eq_cancel_leftₓ'. -/
+#print eq_iff_eq_cancel_left /-
@[simp]
-theorem eq_iff_eq_cancel_left {b c : α} : (∀ {a}, a = b ↔ a = c) ↔ b = c :=
+theorem eq_iff_eq_cancel_left {α : Sort _} {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
+-/
-/- warning: eq_iff_eq_cancel_right -> eq_iff_eq_cancel_right is a dubious translation:
-lean 3 declaration is
- forall {α : Type.{u1}} {a : α} {b : α}, Iff (forall {c : α}, Iff (Eq.{succ u1} α a c) (Eq.{succ u1} α b c)) (Eq.{succ u1} α a b)
-but is expected to have type
- forall {α : Sort.{u1}} {a : α} {b : α}, Iff (forall {c : α}, Iff (Eq.{u1} α a c) (Eq.{u1} α b c)) (Eq.{u1} α a b)
-Case conversion may be inaccurate. Consider using '#align eq_iff_eq_cancel_right eq_iff_eq_cancel_rightₓ'. -/
+#print eq_iff_eq_cancel_right /-
@[simp]
-theorem eq_iff_eq_cancel_right {a b : α} : (∀ {c}, a = c ↔ b = c) ↔ a = b :=
+theorem eq_iff_eq_cancel_right {α : Sort _} {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
+-/
#print Fact /-
/- ./././Mathport/Syntax/Translate/Command.lean:388:30: infer kinds are unsupported in Lean 4: #[`out] [] -/
mathlib commit https://github.com/leanprover-community/mathlib/commit/4c586d291f189eecb9d00581aeb3dd998ac34442
@@ -1780,7 +1780,7 @@ lean 3 declaration is
but is expected to have type
forall {α : Type.{u2}} {β : Type.{u1}} [_inst_1 : Membership.{u2, u1} α β] {s : β} {p : α -> α -> Prop}, Iff (forall (a : α), (Membership.mem.{u2, u1} α β _inst_1 a s) -> (forall (b : α), (Membership.mem.{u2, u1} α β _inst_1 b s) -> (p a b))) (forall (a : α) (b : α), (Membership.mem.{u2, u1} α β _inst_1 a s) -> (Membership.mem.{u2, u1} α β _inst_1 b s) -> (p a b))
Case conversion may be inaccurate. Consider using '#align ball_mem_comm ball_mem_commₓ'. -/
-/- ./././Mathport/Syntax/Translate/Basic.lean:628:2: warning: expanding binder collection (a b «expr ∈ » s) -/
+/- ./././Mathport/Syntax/Translate/Basic.lean:635:2: warning: expanding binder collection (a b «expr ∈ » s) -/
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
@@ -2459,13 +2459,13 @@ lean 3 declaration is
but is expected to have type
forall {α : Sort.{u1}} [p : DecidableEq.{u1} α] (_inst_1 : α) {a : α -> Prop}, Iff (And (a _inst_1) (forall (b : α), (Ne.{u1} α b _inst_1) -> (a b))) (forall (b : α), a b)
Case conversion may be inaccurate. Consider using '#align decidable.and_forall_ne Decidable.and_forall_neₓ'. -/
-/- ./././Mathport/Syntax/Translate/Basic.lean:628:2: warning: expanding binder collection (b «expr ≠ » a) -/
+/- ./././Mathport/Syntax/Translate/Basic.lean:635:2: warning: expanding binder collection (b «expr ≠ » a) -/
theorem Decidable.and_forall_ne [DecidableEq α] (a : α) :
(p a ∧ ∀ (b) (_ : b ≠ a), p b) ↔ ∀ b, p b := by
simp only [← @forall_eq _ p a, ← forall_and, ← or_imp, Decidable.em, forall_const]
#align decidable.and_forall_ne Decidable.and_forall_ne
-/- ./././Mathport/Syntax/Translate/Basic.lean:628:2: warning: expanding binder collection (b «expr ≠ » a) -/
+/- ./././Mathport/Syntax/Translate/Basic.lean:635:2: warning: expanding binder collection (b «expr ≠ » a) -/
#print and_forall_ne /-
theorem and_forall_ne (a : α) : (p a ∧ ∀ (b) (_ : b ≠ a), p b) ↔ ∀ b, p b :=
Decidable.and_forall_ne a
mathlib commit https://github.com/leanprover-community/mathlib/commit/9da1b3534b65d9661eb8f42443598a92bbb49211
@@ -1919,19 +1919,31 @@ theorem cast_eq_iff_heq {α β : Sort _} {a : α} {a' : β} {e : α = β} : cast
lean 3 declaration is
forall {α : Sort.{u1}} {a : α} {b : α} {β : Sort.{u2}} {C : α -> Sort.{u2}} {x : C a} {y : β} (e : Eq.{u1} α a b), (HEq.{u2} (C a) x β y) -> (HEq.{u2} (C b) (Eq.ndrec.{u2, u1} α a C x b e) β y)
but is expected to have type
- forall {α : Sort.{u2}} {a : α} {b : Sort.{u1}} {β : α} {C : α -> Sort.{u1}} {x : C a} {y : b} (e : Eq.{u2} α a β), (HEq.{u1} (C a) x b y) -> (HEq.{u1} (C β) (Eq.ndrec.{u1, u2} α a C x β e) b y)
+ forall {α : Sort.{u2}} {a : α} {b : Sort.{u1}} {β : α} {C : α -> Sort.{u1}} {x : C a} {y : b} (e : Eq.{u2} α a β), (HEq.{u1} (C a) x b y) -> (HEq.{u1} (C β) (Eq.rec.{u1, u2} α a (fun (x._@.Mathlib.Logic.Basic._hyg.4253 : α) (h._@.Mathlib.Logic.Basic._hyg.4254 : Eq.{u2} α a x._@.Mathlib.Logic.Basic._hyg.4253) => C x._@.Mathlib.Logic.Basic._hyg.4253) x β e) b y)
Case conversion may be inaccurate. Consider using '#align rec_heq_of_heq rec_heq_of_heqₓ'. -/
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
-theorem ndrec_hEq_iff_hEq {β} {C : α → Sort _} {x : C a} {y : β} {e : a = b} :
+/- warning: rec_heq_iff_heq -> rec_heq_iff_heq is a dubious translation:
+lean 3 declaration is
+ forall {α : Sort.{u1}} {a : α} {b : α} {β : Sort.{u2}} {C : α -> Sort.{u2}} {x : C a} {y : β} {e : Eq.{u1} α a b}, Iff (HEq.{u2} (C b) (Eq.ndrec.{u2, u1} α a C x b e) β y) (HEq.{u2} (C a) x β y)
+but is expected to have type
+ forall {α : Sort.{u2}} {a : α} {b : Sort.{u1}} {β : α} {C : α -> Sort.{u1}} {x : C a} {y : b} {e : Eq.{u2} α a β}, Iff (HEq.{u1} (C β) (Eq.rec.{u1, u2} α a (fun (x._@.Mathlib.Logic.Basic._hyg.4313 : α) (h._@.Mathlib.Logic.Basic._hyg.4314 : Eq.{u2} α a x._@.Mathlib.Logic.Basic._hyg.4313) => C x._@.Mathlib.Logic.Basic._hyg.4313) x β e) b y) (HEq.{u1} (C a) x b y)
+Case conversion may be inaccurate. Consider using '#align rec_heq_iff_heq rec_heq_iff_heqₓ'. -/
+theorem rec_heq_iff_heq {β} {C : α → Sort _} {x : C a} {y : β} {e : a = b} :
HEq (@Eq.ndrec α a C x b e) y ↔ HEq x y := by subst e
-#align rec_heq_iff_heq ndrec_hEq_iff_hEq
+#align rec_heq_iff_heq rec_heq_iff_heq
-theorem hEq_ndrec_iff_hEq {β} {C : α → Sort _} {x : β} {y : C a} {e : a = b} :
+/- warning: heq_rec_iff_heq -> heq_rec_iff_heq is a dubious translation:
+lean 3 declaration is
+ forall {α : Sort.{u1}} {a : α} {b : α} {β : Sort.{u2}} {C : α -> Sort.{u2}} {x : β} {y : C a} {e : Eq.{u1} α a b}, Iff (HEq.{u2} β x (C b) (Eq.ndrec.{u2, u1} α a C y b e)) (HEq.{u2} β x (C a) y)
+but is expected to have type
+ forall {α : Sort.{u2}} {a : Sort.{u1}} {b : α} {β : α} {C : α -> Sort.{u1}} {x : a} {y : C b} {e : Eq.{u2} α b β}, Iff (HEq.{u1} a x (C β) (Eq.rec.{u1, u2} α b (fun (x._@.Mathlib.Logic.Basic._hyg.4375 : α) (h._@.Mathlib.Logic.Basic._hyg.4376 : Eq.{u2} α b x._@.Mathlib.Logic.Basic._hyg.4375) => C x._@.Mathlib.Logic.Basic._hyg.4375) y β e)) (HEq.{u1} a x (C b) y)
+Case conversion may be inaccurate. Consider using '#align heq_rec_iff_heq heq_rec_iff_heqₓ'. -/
+theorem heq_rec_iff_heq {β} {C : α → Sort _} {x : β} {y : C a} {e : a = b} :
HEq x (@Eq.ndrec α a C y b e) ↔ HEq x y := by subst e
-#align heq_rec_iff_heq hEq_ndrec_iff_hEq
+#align heq_rec_iff_heq heq_rec_iff_heq
/- warning: eq.congr -> Eq.congr is a dubious translation:
lean 3 declaration is
mathlib commit https://github.com/leanprover-community/mathlib/commit/bd9851ca476957ea4549eb19b40e7b5ade9428cc
In one case, replacing this naively errored, so I just merged the single declaration using it.
Delete two unused variables in Logic/Basic.
@@ -23,8 +23,6 @@ Theorems that require decidability hypotheses are in the namespace `Decidable`.
Classical versions are in the namespace `Classical`.
-/
-set_option autoImplicit true
-
open Function
attribute [local instance 10] Classical.propDecidable
@@ -47,11 +45,13 @@ as done in the `elide` and `unelide` tactics. -/
@[reducible] def hidden {α : Sort*} {a : α} := a
#align hidden hidden
+variable {α : Sort*}
+
instance (priority := 10) decidableEq_of_subsingleton [Subsingleton α] : DecidableEq α :=
fun a b ↦ isTrue (Subsingleton.elim a b)
#align decidable_eq_of_subsingleton decidableEq_of_subsingleton
-instance (α : Sort*) [Subsingleton α] (p : α → Prop) : Subsingleton (Subtype p) :=
+instance [Subsingleton α] (p : α → Prop) : Subsingleton (Subtype p) :=
⟨fun ⟨x, _⟩ ⟨y, _⟩ ↦ by cases Subsingleton.elim x y; rfl⟩
#align pempty PEmpty
@@ -61,7 +61,7 @@ theorem congr_heq {α β γ : Sort _} {f : α → γ} {g : β → γ} {x : α} {
cases h₂; cases h₁; rfl
#align congr_heq congr_heq
-theorem congr_arg_heq {α} {β : α → Sort*} (f : ∀ a, β a) :
+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
@@ -74,11 +74,11 @@ theorem ULift.down_injective {α : Sort _} : Function.Injective (@ULift.down α)
⟨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 α)
+theorem PLift.down_injective : 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 :=
+@[simp] theorem PLift.down_inj {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
@@ -90,7 +90,7 @@ theorem PLift.down_injective {α : Sort*} : Function.Injective (@PLift.down α)
⟨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 :=
+lemma ne_and_eq_iff_right {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
@@ -129,7 +129,7 @@ theorem fact_iff {p : Prop} : Fact p ↔ p := ⟨fun h ↦ h.1, fun h ↦ ⟨h
#align fact.elim Fact.elim
/-- Swaps two pairs of arguments to a function. -/
-@[reducible] def Function.swap₂ {κ₁ : ι₁ → Sort*} {κ₂ : ι₂ → Sort*}
+@[reducible] def Function.swap₂ {ι₁ ι₂ : Sort*} {κ₁ : ι₁ → 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₂
@@ -172,14 +172,14 @@ alias Iff.imp := imp_congr
#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
+@[simp] theorem imp_iff_right_iff {a b : Prop} : (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
+@[simp] theorem and_or_imp {a b c : Prop} : 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
+protected theorem Function.mt {a b : Prop} : (a → b) → ¬b → ¬a := mt
#align function.mt Function.mt
/-! ### Declarations about `not` -/
@@ -199,23 +199,25 @@ theorem em' (p : Prop) : ¬p ∨ p := (em p).symm
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
+theorem Decidable.eq_or_ne {α : Sort*} (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
+theorem Decidable.ne_or_eq {α : Sort*} (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
+theorem eq_or_ne {α : Sort*} (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
+theorem ne_or_eq {α : Sort*} (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
+theorem by_contradiction {p : Prop} : (¬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 :=
+theorem by_cases {p q : Prop} (hpq : p → q) (hnpq : ¬p → q) : q :=
if hp : p then hpq hp else hnpq hp
#align classical.by_cases by_cases
@@ -246,13 +248,15 @@ export Classical (not_not)
attribute [simp] not_not
#align not_not Classical.not_not
-theorem of_not_not : ¬¬a → a := by_contra
+variable {a b : Prop}
+
+theorem of_not_not {a : Prop} : ¬¬a → a := by_contra
#align of_not_not of_not_not
-theorem not_ne_iff : ¬a ≠ b ↔ a = b := not_not
+theorem not_ne_iff {α : Sort*} {a b : α} : ¬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
+theorem of_not_imp : ¬(a → b) → a := Decidable.of_not_imp
#align of_not_imp of_not_imp
alias Not.decidable_imp_symm := Decidable.not_imp_symm
@@ -267,7 +271,7 @@ theorem not_imp_comm : ¬a → b ↔ ¬b → a := Decidable.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⟩
+theorem Imp.swap {a b : Sort*} {c : Prop} : a → b → c ↔ b → a → c := ⟨Function.swap, Function.swap⟩
#align imp.swap Imp.swap
alias Iff.not := not_congr
@@ -300,7 +304,7 @@ lemma Iff.ne_right {α β : Sort*} {a b : α} {c d : β} : (a ≠ b ↔ c = d)
@[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]
+theorem xor_comm (a b : Prop) : Xor' a b = Xor' b a := by simp [Xor', and_comm, or_comm]
#align xor_comm xor_comm
instance : Std.Commutative Xor' := ⟨xor_comm⟩
@@ -334,8 +338,8 @@ 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
-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]
+theorem and_symm_right {α : Sort*} (a b : α) (p : Prop) : p ∧ a = b ↔ p ∧ b = a := by simp [eq_comm]
+theorem and_symm_left {α : Sort*} (a b : α) (p : Prop) : a = b ∧ p ↔ b = a ∧ p := by simp [eq_comm]
/-! ### Declarations about `or` -/
@@ -348,22 +352,25 @@ alias ⟨Or.rotate, _⟩ := or_rotate
#align or.rotate Or.rotate
@[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₁
+theorem or_of_or_of_imp_of_imp {a b c d : Prop} (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₁
+theorem or_of_or_of_imp_left {a c b : Prop} (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₁
+theorem or_of_or_of_imp_right {c a b : Prop} (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 :=
+theorem Or.elim3 {c 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 :=
+theorem Or.imp3 {d e c f : Prop} (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
@@ -387,7 +394,7 @@ theorem or_not_of_imp : (a → b) → b ∨ ¬a := Decidable.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
+theorem imp_iff_or_not {b a : Prop} : 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
@@ -406,10 +413,11 @@ protected theorem Function.mtr : (¬a → ¬b) → b → a := not_imp_not.mp
#align decidable.or_iff_not_imp_right Decidable.or_iff_not_imp_rightₓ -- reorder implicits
#align decidable.imp_iff_or_not Decidable.imp_iff_or_notₓ -- reorder implicits
-theorem or_congr_left' (h : ¬c → (a ↔ b)) : a ∨ c ↔ b ∨ c := Decidable.or_congr_left' h
+theorem or_congr_left' {c a b : Prop} (h : ¬c → (a ↔ b)) : a ∨ c ↔ b ∨ c :=
+ Decidable.or_congr_left' h
#align or_congr_left or_congr_left'
-theorem or_congr_right' (h : ¬a → (b ↔ c)) : a ∨ b ↔ a ∨ c := Decidable.or_congr_right' h
+theorem or_congr_right' {c : Prop} (h : ¬a → (b ↔ c)) : a ∨ b ↔ a ∨ c := Decidable.or_congr_right' h
#align or_congr_right or_congr_right'ₓ -- reorder implicits
#align or_iff_left or_iff_leftₓ -- reorder implicits
@@ -427,7 +435,7 @@ alias Iff.iff := iff_congr
#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
+theorem iff_mpr_iff_true_intro {P : Prop} (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
@@ -437,7 +445,7 @@ theorem imp_or {a b c : Prop} : a → b ∨ c ↔ (a → b) ∨ (a → c) := Dec
#align decidable.imp_or_distrib' Decidable.imp_or'
-theorem imp_or' : a → b ∨ c ↔ (a → b) ∨ (a → c) := Decidable.imp_or'
+theorem imp_or' {a : Sort*} {b c : Prop} : 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_iff_and_not
@@ -533,15 +541,15 @@ theorem forall_mem_comm {α β} [Membership α β] {s : β} {p : α → α → P
#align ne_of_apply_ne ne_of_apply_ne
-lemma ne_of_eq_of_ne (h₁ : a = b) (h₂ : b ≠ c) : a ≠ c := h₁.symm ▸ h₂
-lemma ne_of_ne_of_eq (h₁ : a ≠ b) (h₂ : b = c) : a ≠ c := h₂ ▸ h₁
+lemma ne_of_eq_of_ne {α : Sort*} {a b c : α} (h₁ : a = b) (h₂ : b ≠ c) : a ≠ c := h₁.symm ▸ h₂
+lemma ne_of_ne_of_eq {α : Sort*} {a b c : α} (h₁ : a ≠ b) (h₂ : b = c) : a ≠ c := h₂ ▸ h₁
alias Eq.trans_ne := ne_of_eq_of_ne
alias Ne.trans_eq := ne_of_ne_of_eq
#align eq.trans_ne Eq.trans_ne
#align ne.trans_eq Ne.trans_eq
-theorem eq_equivalence : Equivalence (@Eq α) :=
+theorem eq_equivalence {α : Sort*} : Equivalence (@Eq α) :=
⟨Eq.refl, @Eq.symm _, @Eq.trans _⟩
#align eq_equivalence eq_equivalence
@@ -553,25 +561,28 @@ attribute [simp] eq_mp_eq_cast eq_mpr_eq_cast
#align cast_cast cast_cast
-- @[simp] -- FIXME simp ignores proof rewrites
-theorem congr_refl_left (f : α → β) {a b : α} (h : a = b) :
+theorem congr_refl_left {α β : Sort*} (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 : α) :
+theorem congr_refl_right {α β : Sort*} {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
+theorem congr_arg_refl {α β : Sort*} (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
+theorem congr_fun_rfl {α β : Sort*} (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 : β) :
+theorem congr_fun_congr_arg {α β γ : Sort*} (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
@@ -582,19 +593,22 @@ theorem Eq.rec_eq_cast {α : Sort _} {P : α → Sort _} {x y : α} (h : x = y)
h ▸ z = cast (congr_arg P h) z := by induction h; rfl
-- Porting note (#10756): new theorem. More general version of `eqRec_heq`
-theorem eqRec_heq' {α : Sort u_1} {a' : α} {motive : (a : α) → a' = a → Sort u}
+theorem eqRec_heq' {α : Sort*} {a' : α} {motive : (a : α) → a' = a → Sort*}
(p : motive a' (rfl : a' = a')) {a : α} (t : a' = a) :
HEq (@Eq.rec α a' motive p a t) p :=
by subst t; rfl
+set_option autoImplicit true in
theorem rec_heq_of_heq {C : α → Sort*} {x : C a} {y : β} (e : a = b) (h : HEq x y) :
HEq (e ▸ x) y := by subst e; exact h
#align rec_heq_of_heq rec_heq_of_heq
+set_option autoImplicit true in
theorem rec_heq_iff_heq {C : α → Sort*} {x : C a} {y : β} {e : a = b} :
HEq (e ▸ x) y ↔ HEq x y := by subst e; rfl
#align rec_heq_iff_heq rec_heq_iff_heq
+set_option autoImplicit true in
theorem heq_rec_iff_heq {C : α → Sort*} {x : β} {y : C a} {e : a = b} :
HEq x (e ▸ y) ↔ HEq x y := by subst e; rfl
#align heq_rec_iff_heq heq_rec_iff_heq
@@ -604,8 +618,6 @@ theorem heq_rec_iff_heq {C : α → Sort*} {x : β} {y : C a} {e : a = b} :
#align eq.congr_right Eq.congr_right
#align congr_arg2 congr_arg₂
-variable {β : α → Sort*} {γ : ∀ a, β a → Sort*} {δ : ∀ a b, γ a b → Sort*}
-
#align congr_fun₂ congr_fun₂
#align congr_fun₃ congr_fun₃
#align funext₂ funext₂
@@ -614,12 +626,10 @@ variable {β : α → Sort*} {γ : ∀ a, β a → Sort*} {δ : ∀ a b, γ a b
end Equality
/-! ### Declarations about quantifiers -/
-
-
section Quantifiers
section Dependent
-variable {β : α → Sort*} {γ : ∀ a, β a → Sort*} {δ : ∀ a b, γ a b → Sort*}
+variable {α : Sort*} {β : α → Sort*} {γ : ∀ a, β a → Sort*} {δ : ∀ a b, γ a b → Sort*}
{ε : ∀ a b c, δ a b c → Sort*}
theorem pi_congr {β' : α → Sort _} (h : ∀ a, β a = β' a) : (∀ a, β a) = ∀ a, β' a :=
@@ -651,14 +661,15 @@ theorem Exists₃.imp {p q : ∀ a b, γ a b → Prop} (h : ∀ a b c, p a b c
end Dependent
-variable {κ : ι → Sort*} {p q : α → Prop}
+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} :
+theorem forall₂_swap
+ {ι₁ ι₂ : Sort*} {κ₁ : ι₁ → 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
@@ -701,7 +712,8 @@ lemma forall_or_exists_not (P : α → Prop) : (∀ a, P a) ∨ ∃ a, ¬ P a :=
lemma exists_or_forall_not (P : α → Prop) : (∃ a, P a) ∨ ∀ a, ¬ P a := by
rw [← not_exists]; exact em _
-theorem forall_imp_iff_exists_imp [ha : Nonempty α] : (∀ x, p x) → b ↔ ∃ x, p x → b := by
+theorem forall_imp_iff_exists_imp {α : Sort*} {p : α → Prop} {b : Prop} [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
@@ -735,7 +747,7 @@ theorem forall₃_true_iff {β : α → Sort*} {γ : ∀ a, β a → Sort*} :
#align exists_const exists_const
-theorem exists_unique_const (α) [i : Nonempty α] [Subsingleton α] :
+theorem exists_unique_const {b : Prop} (α : Sort*) [i : Nonempty α] [Subsingleton α] :
(∃! _ : α, b) ↔ b := by simp
#align exists_unique_const exists_unique_const
@@ -825,7 +837,8 @@ theorem forall_eq_apply_imp_iff' {f : α → β} {p : β → Prop} :
#align exists_comm exists_comm
-theorem exists₂_comm {κ₁ : ι₁ → Sort*} {κ₂ : ι₂ → Sort*} {p : ∀ i₁, κ₁ i₁ → ∀ i₂, κ₂ i₂ → Prop} :
+theorem exists₂_comm
+ {ι₁ ι₂ : Sort*} {κ₁ : ι₁ → 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
@@ -834,7 +847,9 @@ theorem And.exists {p q : Prop} {f : p ∧ q → Prop} : (∃ h, f h) ↔ ∃ hp
⟨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
+theorem forall_or_of_or_forall {α : Sort*} {p : α → Prop} {b : Prop} (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]
@@ -935,13 +950,14 @@ lemma imp_congr_eq {a b c d : Prop} (h₁ : a = c) (h₂ : b = d) : (a → b) =
lemma imp_congr_ctx_eq {a b c d : Prop} (h₁ : a = c) (h₂ : c → b = d) : (a → b) = (c → d) :=
propext (imp_congr_ctx h₁.to_iff fun hc ↦ (h₂ hc).to_iff)
-lemma eq_true_intro (h : a) : a = True := propext (iff_true_intro h)
-lemma eq_false_intro (h : ¬a) : a = False := propext (iff_false_intro h)
+lemma eq_true_intro {a : Prop} (h : a) : a = True := propext (iff_true_intro h)
+
+lemma eq_false_intro {a : Prop} (h : ¬a) : a = False := propext (iff_false_intro h)
-- FIXME: `alias` creates `def Iff.eq := propext` instead of `lemma Iff.eq := propext`
@[nolint defLemma] alias Iff.eq := propext
-lemma iff_eq_eq : (a ↔ b) = (a = b) := propext ⟨propext, Eq.to_iff⟩
+lemma iff_eq_eq {a b : Prop} : (a ↔ b) = (a = b) := propext ⟨propext, Eq.to_iff⟩
-- They were not used in Lean 3 and there are already lemmas with those names in Lean 4
#noalign eq_false
@@ -984,13 +1000,14 @@ end Quantifiers
/-! ### Classical lemmas -/
namespace Classical
-variable {p : α → Prop}
-- use shortened names to avoid conflict when classical namespace is open.
/-- Any prop `p` is decidable classically. A shorthand for `Classical.propDecidable`. -/
noncomputable def dec (p : Prop) : Decidable p := by infer_instance
#align classical.dec Classical.dec
+variable {α : Sort*} {p : α → Prop}
+
/-- Any predicate `p` is decidable classically. -/
noncomputable def decPred (p : α → Prop) : DecidablePred p := by infer_instance
#align classical.dec_pred Classical.decPred
@@ -1000,13 +1017,13 @@ noncomputable def decRel (p : α → α → Prop) : DecidableRel p := by infer_i
#align classical.dec_rel Classical.decRel
/-- Any type `α` has decidable equality classically. -/
-noncomputable def decEq (α : Sort u) : DecidableEq α := by infer_instance
+noncomputable def decEq (α : Sort*) : 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 :=
+noncomputable def existsCases {α C : Sort*} {p : α → Prop} (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
@@ -1031,6 +1048,7 @@ def choice_of_byContradiction' {α : Sort*} (contra : ¬(α → False) → α) :
end Classical
+set_option autoImplicit true in
/-- This function has the same type as `Exists.recOn`, and can be used to case on an equality,
but `Exists.recOn` can only eliminate into Prop, while this version eliminates into any universe
using the axiom of choice. -/
@@ -1040,9 +1058,9 @@ noncomputable def Exists.classicalRecOn {p : α → Prop} (h : ∃ a, p a) {C} (
#align exists.classical_rec_on Exists.classicalRecOn
/-! ### Declarations about bounded quantifiers -/
-
section BoundedQuantifiers
-variable {r p q : α → Prop} {P Q : ∀ x, p x → Prop} {b : Prop}
+
+variable {α : Sort*} {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⟩⟩
@@ -1162,7 +1180,7 @@ end BoundedQuantifiers
section ite
-variable {σ : α → Sort*} (f : α → β) {P Q R : Prop} [Decidable P] [Decidable Q]
+variable {α : Sort*} {σ : α → Sort*} {P Q R : Prop} [Decidable P] [Decidable Q]
{a b c : α} {A : P → α} {B : ¬P → α}
theorem dite_eq_iff : dite P A B = c ↔ (∃ h, A h = c) ∨ ∃ h, B h = c := by
@@ -1252,15 +1270,15 @@ theorem ite_eq_or_eq : ite P a b = a ∨ ite P a b = b :=
/-- A two-argument function applied to two `dite`s is a `dite` of that two-argument function
applied to each of the branches. -/
-theorem apply_dite₂ (f : α → β → γ) (P : Prop) [Decidable P] (a : P → α) (b : ¬P → α)
- (c : P → β) (d : ¬P → β) :
+theorem apply_dite₂ {α β γ : Sort*} (f : α → β → γ) (P : Prop) [Decidable P]
+ (a : P → α) (b : ¬P → α) (c : P → β) (d : ¬P → β) :
f (dite P a b) (dite P c d) = dite P (fun h ↦ f (a h) (c h)) fun h ↦ f (b h) (d h) := by
by_cases h : P <;> simp [h]
#align apply_dite2 apply_dite₂
/-- A two-argument function applied to two `ite`s is a `ite` of that two-argument function
applied to each of the branches. -/
-theorem apply_ite₂ (f : α → β → γ) (P : Prop) [Decidable P] (a b : α) (c d : β) :
+theorem apply_ite₂ {α β γ : Sort*} (f : α → β → γ) (P : Prop) [Decidable P] (a b : α) (c d : β) :
f (ite P a b) (ite P c d) = ite P (f a c) (f b d) :=
apply_dite₂ f P (fun _ ↦ a) (fun _ ↦ b) (fun _ ↦ c) fun _ ↦ d
#align apply_ite2 apply_ite₂
@@ -1326,15 +1344,15 @@ theorem dite_prop_iff_and {Q : P → Prop} {R : ¬P → Prop} [Decidable P] :
end ite
-theorem not_beq_of_ne [BEq α] [LawfulBEq α] {a b : α} (ne : a ≠ b) : ¬(a == b) :=
+theorem not_beq_of_ne {α : Type*} [BEq α] [LawfulBEq α] {a b : α} (ne : a ≠ b) : ¬(a == b) :=
fun h => ne (eq_of_beq h)
-theorem beq_eq_decide [BEq α] [LawfulBEq α] {a b : α} : (a == b) = decide (a = b) := by
+theorem beq_eq_decide {α : Type*} [BEq α] [LawfulBEq α] {a b : α} : (a == b) = decide (a = b) := by
rw [← beq_iff_eq a b]
cases a == b <;> simp
@[ext]
-theorem beq_ext (inst1 : BEq α) (inst2 : BEq α)
+theorem beq_ext {α : Type*} (inst1 : BEq α) (inst2 : BEq α)
(h : ∀ x y, @BEq.beq _ inst1 x y = @BEq.beq _ inst2 x y) :
inst1 = inst2 := by
have ⟨beq1⟩ := inst1
@@ -1343,7 +1361,7 @@ theorem beq_ext (inst1 : BEq α) (inst2 : BEq α)
funext x y
exact h x y
-theorem lawful_beq_subsingleton (inst1 : BEq α) (inst2 : BEq α)
+theorem lawful_beq_subsingleton {α : Type*} (inst1 : BEq α) (inst2 : BEq α)
[@LawfulBEq α inst1] [@LawfulBEq α inst2] :
inst1 = inst2 := by
apply beq_ext
ExistsUnique
notation throw an error when used with more than one binder (#12218)
@@ -953,7 +953,7 @@ lemma iff_eq_eq : (a ↔ b) = (a = b) := propext ⟨propext, Eq.to_iff⟩
#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)
+ {q : ∀ (x) (_ : p x), Prop} {b : Prop} (h₂ : ∃! x, ∃! h : p x, q x h)
(h₁ : ∀ (x) (h : p x), q x h → (∀ (y) (hy : p y), q y hy → y = x) → b) : b := by
simp only [exists_unique_iff_exists] at h₂
apply h₂.elim
@@ -962,18 +962,18 @@ theorem ExistsUnique.elim₂ {α : Sort*} {p : α → Sort*} [∀ x, Subsingleto
theorem ExistsUnique.intro₂ {α : Sort*} {p : α → Sort*} [∀ x, Subsingleton (p x)]
{q : ∀ (x : α) (_ : p x), Prop} (w : α) (hp : p w) (hq : q w hp)
- (H : ∀ (y) (hy : p y), q y hy → y = w) : ∃! (x : _) (hx : p x), q x hx := by
+ (H : ∀ (y) (hy : p y), q y hy → y = w) : ∃! x, ∃! hx : p x, q x hx := by
simp only [exists_unique_iff_exists]
exact ExistsUnique.intro w ⟨hp, hq⟩ fun y ⟨hyp, hyq⟩ ↦ H y hyp hyq
#align exists_unique.intro2 ExistsUnique.intro₂
theorem ExistsUnique.exists₂ {α : Sort*} {p : α → Sort*} {q : ∀ (x : α) (_ : p x), Prop}
- (h : ∃! (x : _) (hx : p x), q x hx) : ∃ (x : _) (hx : p x), q x hx :=
+ (h : ∃! x, ∃! hx : p x, q x hx) : ∃ (x : _) (hx : p x), q x hx :=
h.exists.imp fun _ hx ↦ hx.exists
#align exists_unique.exists2 ExistsUnique.exists₂
theorem ExistsUnique.unique₂ {α : Sort*} {p : α → Sort*} [∀ x, Subsingleton (p x)]
- {q : ∀ (x : α) (_ : p x), Prop} (h : ∃! (x : _) (hx : p x), q x hx) {y₁ y₂ : α}
+ {q : ∀ (x : α) (_ : p x), Prop} (h : ∃! x, ∃! hx : p x, q x hx) {y₁ y₂ : α}
(hpy₁ : p y₁) (hqy₁ : q y₁ hpy₁) (hpy₂ : p y₂) (hqy₂ : q y₂ hpy₂) : y₁ = y₂ := by
simp only [exists_unique_iff_exists] at h
exact h.unique ⟨hpy₁, hqy₁⟩ ⟨hpy₂, hqy₂⟩
These are changes from #11997, the latest adaptation PR for nightly-2024-04-07, which can be made directly on master.
Co-authored-by: Scott Morrison <scott.morrison@gmail.com> Co-authored-by: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com>
@@ -165,7 +165,7 @@ alias Iff.imp := imp_congr
#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]
+ funext _; simp only [true_iff, id, eq_iff_iff]
#align eq_true_eq_id eq_true_eq_id
#align imp_and_distrib imp_and
bex
and ball
from lemma names (#11615)
Follow-up to #10816.
Remaining places containing such lemmas are
Option.bex_ne_none
and Option.ball_ne_none
: defined in Lean coreNat.decidableBallLT
and Nat.decidableBallLE
: defined in Lean corebef_def
is still used in a number of places and could be renamedBAll.imp_{left,right}
, BEx.imp_{left,right}
, BEx.intro
and BEx.elim
I only audited the first ~150 lemmas mentioning "ball"; too many lemmas named after Metric.ball/openBall/closedBall.
Co-authored-by: Yaël Dillies <yael.dillies@gmail.com>
@@ -517,15 +517,19 @@ alias Membership.mem.ne_of_not_mem' := ne_of_mem_of_not_mem'
section Equality
-- todo: change name
-theorem ball_cond_comm {α} {s : α → Prop} {p : α → α → Prop} :
+theorem forall_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
+#align ball_cond_comm forall_cond_comm
-theorem ball_mem_comm {α β} [Membership α β] {s : β} {p : α → α → Prop} :
+theorem forall_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
+ forall_cond_comm
+#align ball_mem_comm forall_mem_comm
+
+-- 2024-03-23
+@[deprecated] alias ball_cond_comm := forall_cond_comm
+@[deprecated] alias ball_mem_comm := forall_mem_comm
#align ne_of_apply_ne ne_of_apply_ne
@@ -1052,18 +1056,18 @@ theorem BEx.intro (a : α) (h₁ : p a) (h₂ : P a h₁) : ∃ (x : _) (h : p x
⟨a, h₁, h₂⟩
#align bex.intro BEx.intro
-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
+#align ball_congr forall₂_congr
+#align bex_congr exists₂_congr
+@[deprecated exists_eq_left] -- 2024-04-06
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
+-- 2024-04-06
+@[deprecated] alias ball_congr := forall₂_congr
+@[deprecated] alias bex_congr := exists₂_congr
+
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₁ _ _
#align ball.imp_right BAll.imp_right
@@ -1080,61 +1084,77 @@ theorem BEx.imp_left (H : ∀ x, p x → q x) : (∃ (x : _) (_ : p x), r x) →
| ⟨x, hp, hr⟩ => ⟨x, H _ hp, hr⟩
#align bex.imp_left BEx.imp_left
+@[deprecated id] -- 2024-03-23
theorem ball_of_forall (h : ∀ x, p x) (x) : p x := h x
#align ball_of_forall ball_of_forall
+@[deprecated forall_imp] -- 2024-03-23
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
+theorem exists_mem_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
+#align bex_of_exists exists_mem_of_exists
-theorem exists_of_bex : (∃ (x : _) (_ : p x), q x) → ∃ x, q x
+theorem exists_of_exists_mem : (∃ (x : _) (_ : p x), q x) → ∃ x, q x
| ⟨x, _, hq⟩ => ⟨x, hq⟩
-#align exists_of_bex exists_of_bex
+#align exists_of_bex exists_of_exists_mem
+
+theorem exists₂_imp : (∃ x h, P x h) → b ↔ ∀ x h, P x h → b := by simp
+#align bex_imp_distrib exists₂_imp
-theorem bex_imp : (∃ x h, P x h) → b ↔ ∀ x h, P x h → b := by simp
-#align bex_imp_distrib bex_imp
+-- 2024-03-23
+@[deprecated] alias bex_of_exists := exists_mem_of_exists
+@[deprecated] alias exists_of_bex := exists_of_exists_mem
+@[deprecated] alias bex_imp := exists₂_imp
-theorem not_bex : (¬∃ x h, P x h) ↔ ∀ x h, ¬P x h := bex_imp
-#align not_bex not_bex
+theorem not_exists_mem : (¬∃ x h, P x h) ↔ ∀ x h, ¬P x h := exists₂_imp
+#align not_bex not_exists_mem
-theorem not_ball_of_bex_not : (∃ x h, ¬P x h) → ¬∀ x h, P x h
+theorem not_forall₂_of_exists₂_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
+#align not_ball_of_bex_not not_forall₂_of_exists₂_not
-- See Note [decidable namespace]
-protected theorem Decidable.not_ball [Decidable (∃ x h, ¬P x h)] [∀ x h, Decidable (P x h)] :
+protected theorem Decidable.not_forall₂ [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
+ fun h' ↦ ⟨x, h, h'⟩, not_forall₂_of_exists₂_not⟩
+#align decidable.not_ball Decidable.not_forall₂
-theorem not_ball : (¬∀ x h, P x h) ↔ ∃ x h, ¬P x h := Decidable.not_ball
-#align not_ball not_ball
+theorem not_forall₂ : (¬∀ x h, P x h) ↔ ∃ x h, ¬P x h := Decidable.not_forall₂
+#align not_ball not_forall₂
-theorem ball_true_iff (p : α → Prop) : (∀ x, p x → True) ↔ True :=
- iff_true_intro fun _ _ ↦ trivial
-#align ball_true_iff ball_true_iff
+#align ball_true_iff forall₂_true_iff
-theorem ball_and : (∀ x h, P x h ∧ Q x h) ↔ (∀ x h, P x h) ∧ ∀ x h, Q x h :=
+theorem forall₂_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
-#align ball_and_distrib ball_and
+#align ball_and_distrib forall₂_and
-theorem bex_or : (∃ x h, P x h ∨ Q x h) ↔ (∃ x h, P x h) ∨ ∃ x h, Q x h :=
+theorem exists_mem_or : (∃ x h, P x h ∨ Q x h) ↔ (∃ x h, P x h) ∨ ∃ x h, Q x h :=
Iff.trans (exists_congr fun _ ↦ exists_or) exists_or
-#align bex_or_distrib bex_or
+#align bex_or_distrib exists_mem_or
-theorem ball_or_left : (∀ x, p x ∨ q x → r x) ↔ (∀ x, p x → r x) ∧ ∀ x, q x → r x :=
+theorem forall₂_or_left : (∀ x, p x ∨ q x → r x) ↔ (∀ x, p x → r x) ∧ ∀ x, q x → r x :=
Iff.trans (forall_congr' fun _ ↦ or_imp) forall_and
-#align ball_or_left_distrib ball_or_left
+#align ball_or_left_distrib forall₂_or_left
-theorem bex_or_left :
+theorem exists_mem_or_left :
(∃ (x : _) (_ : p x ∨ q x), r x) ↔ (∃ (x : _) (_ : p x), r x) ∨ ∃ (x : _) (_ : q x), r x := by
simp only [exists_prop]
exact Iff.trans (exists_congr fun x ↦ or_and_right) exists_or
-#align bex_or_left_distrib bex_or_left
+#align bex_or_left_distrib exists_mem_or_left
+
+-- 2023-03-23
+@[deprecated] alias not_ball_of_bex_not := not_forall₂_of_exists₂_not
+@[deprecated] alias Decidable.not_ball := Decidable.not_forall₂
+@[deprecated] alias not_ball := not_forall₂
+@[deprecated] alias ball_true_iff := forall₂_true_iff
+@[deprecated] alias ball_and := forall₂_and
+@[deprecated] alias not_bex := not_exists_mem
+@[deprecated] alias bex_or := exists_mem_or
+@[deprecated] alias ball_or_left := forall₂_or_left
+@[deprecated] alias bex_or_left := exists_mem_or_left
end BoundedQuantifiers
@@ -923,6 +923,26 @@ theorem forall_prop_congr' {p p' : Prop} {q q' : p → Prop} (hq : ∀ h, q h
propext (forall_prop_congr hq hp)
#align forall_prop_congr' forall_prop_congr'
+#align forall_congr_eq forall_congr
+
+lemma imp_congr_eq {a b c d : Prop} (h₁ : a = c) (h₂ : b = d) : (a → b) = (c → d) :=
+ propext (imp_congr h₁.to_iff h₂.to_iff)
+
+lemma imp_congr_ctx_eq {a b c d : Prop} (h₁ : a = c) (h₂ : c → b = d) : (a → b) = (c → d) :=
+ propext (imp_congr_ctx h₁.to_iff fun hc ↦ (h₂ hc).to_iff)
+
+lemma eq_true_intro (h : a) : a = True := propext (iff_true_intro h)
+lemma eq_false_intro (h : ¬a) : a = False := propext (iff_false_intro h)
+
+-- FIXME: `alias` creates `def Iff.eq := propext` instead of `lemma Iff.eq := propext`
+@[nolint defLemma] alias Iff.eq := propext
+
+lemma iff_eq_eq : (a ↔ b) = (a = b) := propext ⟨propext, Eq.to_iff⟩
+
+-- They were not used in Lean 3 and there are already lemmas with those names in Lean 4
+#noalign eq_false
+#noalign eq_true
+
/-- 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 _
@@ -1150,12 +1150,12 @@ theorem ite_eq_iff' : ite P a b = c ↔ (P → a = c) ∧ (¬P → b = c) := dit
#align ite_eq_right_iff ite_eq_right_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]
+ rw [Ne, 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]
+ simp only [Ne, 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 :=
congr!
discharge equalities of BEq
instances (#11179)
Adds a congruence lemma for BEq
instances that makes use of LawfulBEq
instances, and gives congr!
the ability to use this congruence lemma. This is meant to help with diamonds that arise from interactions between Decidable and BEq instances.
This feature can be turned off using the beqEq
configuration setting, like congr! (config := { beqEq := false })
.
@@ -1292,3 +1292,20 @@ theorem not_beq_of_ne [BEq α] [LawfulBEq α] {a b : α} (ne : a ≠ b) : ¬(a =
theorem beq_eq_decide [BEq α] [LawfulBEq α] {a b : α} : (a == b) = decide (a = b) := by
rw [← beq_iff_eq a b]
cases a == b <;> simp
+
+@[ext]
+theorem beq_ext (inst1 : BEq α) (inst2 : BEq α)
+ (h : ∀ x y, @BEq.beq _ inst1 x y = @BEq.beq _ inst2 x y) :
+ inst1 = inst2 := by
+ have ⟨beq1⟩ := inst1
+ have ⟨beq2⟩ := inst2
+ congr
+ funext x y
+ exact h x y
+
+theorem lawful_beq_subsingleton (inst1 : BEq α) (inst2 : BEq α)
+ [@LawfulBEq α inst1] [@LawfulBEq α inst2] :
+ inst1 = inst2 := by
+ apply beq_ext
+ intro x y
+ simp only [beq_eq_decide]
@@ -577,7 +577,7 @@ theorem congr_fun_congr_arg (f : α → β → γ) {a a' : α} (p : a = a') (b :
theorem Eq.rec_eq_cast {α : Sort _} {P : α → Sort _} {x y : α} (h : x = y) (z : P x) :
h ▸ z = cast (congr_arg P h) z := by induction h; rfl
--- Porting note: new theorem. More general version of `eqRec_heq`
+-- Porting note (#10756): new theorem. More general version of `eqRec_heq`
theorem eqRec_heq' {α : Sort u_1} {a' : α} {motive : (a : α) → a' = a → Sort u}
(p : motive a' (rfl : a' = a')) {a : α} (t : a' = a) :
HEq (@Eq.rec α a' motive p a t) p :=
Homogenises porting notes via capitalisation and addition of whitespace.
It makes the following changes:
@@ -577,7 +577,7 @@ theorem congr_fun_congr_arg (f : α → β → γ) {a a' : α} (p : a = a') (b :
theorem Eq.rec_eq_cast {α : Sort _} {P : α → Sort _} {x y : α} (h : x = y) (z : P x) :
h ▸ z = cast (congr_arg P h) z := by induction h; rfl
---Porting note: new theorem. More general version of `eqRec_heq`
+-- Porting note: new theorem. More general version of `eqRec_heq`
theorem eqRec_heq' {α : Sort u_1} {a' : α} {motive : (a : α) → a' = a → Sort u}
(p : motive a' (rfl : a' = a')) {a : α} (t : a' = a) :
HEq (@Eq.rec α a' motive p a t) p :=
@@ -766,7 +766,7 @@ theorem Ne.ne_or_ne {x y : α} (z : α) (h : x ≠ y) : x ≠ z ∨ y ≠ z :=
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:
+-- Porting note: an alternative workaround theorem:
theorem exists_apply_eq (a : α) (b : β) : ∃ f : α → β, f a = b := ⟨fun _ ↦ b, rfl⟩
@[simp] theorem exists_exists_and_eq_and {f : α → β} {p : α → Prop} {q : β → Prop} :
@@ -1285,3 +1285,10 @@ theorem dite_prop_iff_and {Q : P → Prop} {R : ¬P → Prop} [Decidable P] :
#align if_false_left_eq_and if_false_left
end ite
+
+theorem not_beq_of_ne [BEq α] [LawfulBEq α] {a b : α} (ne : a ≠ b) : ¬(a == b) :=
+ fun h => ne (eq_of_beq h)
+
+theorem beq_eq_decide [BEq α] [LawfulBEq α] {a b : α} : (a == b) = decide (a = b) := by
+ rw [← beq_iff_eq a b]
+ cases a == b <;> simp
Init.CCLemmas
(#10696)
Those lemmas were weird and unused, except the last few about transitivity of =
and ≠
, which I moved to Logic.Basic
@@ -529,6 +529,14 @@ theorem ball_mem_comm {α β} [Membership α β] {s : β} {p : α → α → Pro
#align ne_of_apply_ne ne_of_apply_ne
+lemma ne_of_eq_of_ne (h₁ : a = b) (h₂ : b ≠ c) : a ≠ c := h₁.symm ▸ h₂
+lemma ne_of_ne_of_eq (h₁ : a ≠ b) (h₂ : b = c) : a ≠ c := h₂ ▸ h₁
+
+alias Eq.trans_ne := ne_of_eq_of_ne
+alias Ne.trans_eq := ne_of_ne_of_eq
+#align eq.trans_ne Eq.trans_ne
+#align ne.trans_eq Ne.trans_eq
+
theorem eq_equivalence : Equivalence (@Eq α) :=
⟨Eq.refl, @Eq.symm _, @Eq.trans _⟩
#align eq_equivalence eq_equivalence
Init.IteSimp
(#10708)
These lemmas can easily go in Logic.Basic
. Also rename and restate them to better match the convention.
@@ -10,6 +10,7 @@ import Std.Util.LibraryNote
import Std.Tactic.Lint.Basic
#align_import logic.basic from "leanprover-community/mathlib"@"3365b20c2ffa7c35e47e5209b89ba9abdddf3ffe"
+#align_import init.ite_simp from "leanprover-community/lean"@"4a03bdeb31b3688c31d02d7ff8e0ff2e5d6174db"
/-!
# Basic logic properties
@@ -1113,7 +1114,7 @@ end BoundedQuantifiers
section ite
-variable {σ : α → Sort*} (f : α → β) {P Q : Prop} [Decidable P] [Decidable Q]
+variable {σ : α → Sort*} (f : α → β) {P Q R : Prop} [Decidable P] [Decidable Q]
{a b c : α} {A : P → α} {B : ¬P → α}
theorem dite_eq_iff : dite P A B = c ↔ (∃ h, A h = c) ∨ ∃ h, B h = c := by
@@ -1246,7 +1247,7 @@ theorem ite_ite_comm (h : P → ¬Q) :
dite_dite_comm P Q h
#align ite_ite_comm ite_ite_comm
-variable {R : Prop}
+variable {P Q}
theorem ite_prop_iff_or : (if P then Q else R) ↔ (P ∧ Q ∨ ¬ P ∧ R) := by
by_cases p : P <;> simp [p]
@@ -1263,4 +1264,16 @@ theorem dite_prop_iff_and {Q : P → Prop} {R : ¬P → Prop} [Decidable P] :
dite P Q R ↔ (∀ h, Q h) ∧ (∀ h, R h) := by
by_cases h : P <;> simp [h, forall_prop_of_false, forall_prop_of_true]
+@[simp] lemma if_true_right : (if P then Q else True) ↔ ¬P ∨ Q := by by_cases P <;> simp [*]
+#align if_true_right_eq_or if_true_right
+
+@[simp] lemma if_true_left : (if P then True else Q) ↔ P ∨ Q := by by_cases P <;> simp [*]
+#align if_true_left_eq_or if_true_left
+
+@[simp] lemma if_false_right : (if P then Q else False) ↔ P ∧ Q := by by_cases P <;> simp [*]
+#align if_false_right_eq_and if_false_right
+
+@[simp] lemma if_false_left : (if P then False else Q) ↔ ¬P ∧ Q := by by_cases P <;> simp [*]
+#align if_false_left_eq_and if_false_left
+
end ite
@@ -776,6 +776,12 @@ theorem exists_apply_eq (a : α) (b : β) : ∃ f : α → β, f a = b := ⟨fun
⟨fun ⟨_, ⟨a, ha, b, hb, hab⟩, hc⟩ ↦ ⟨a, ha, b, hb, hab.symm ▸ hc⟩,
fun ⟨a, ha, b, hb, hab⟩ ↦ ⟨f a b, ⟨a, ha, b, hb, rfl⟩, hab⟩⟩
+@[simp] theorem exists_exists_exists_and_eq {α β γ : Type*}
+ {f : α → β → γ} {p : γ → Prop} :
+ (∃ c, (∃ a, ∃ b, f a b = c) ∧ p c) ↔ ∃ a, ∃ b, p (f a b) :=
+ ⟨fun ⟨_, ⟨a, b, hab⟩, hc⟩ ↦ ⟨a, b, hab.symm ▸ hc⟩,
+ fun ⟨a, b, hab⟩ ↦ ⟨f a b, ⟨a, b, rfl⟩, hab⟩⟩
+
@[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
@@ -439,7 +439,7 @@ theorem imp_or {a b c : Prop} : a → b ∨ c ↔ (a → b) ∨ (a → c) := Dec
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
+theorem not_imp : ¬(a → b) ↔ a ∧ ¬b := Decidable.not_imp_iff_and_not
#align not_imp not_imp
theorem peirce (a b : Prop) : ((a → b) → a) → a := Decidable.peirce _ _
@@ -474,12 +474,12 @@ theorem not_and_not_right : ¬(a ∧ ¬b) ↔ a → b := Decidable.not_and_not_r
/-! ### De Morgan's laws -/
-#align decidable.not_and_distrib Decidable.not_and
-#align decidable.not_and_distrib' Decidable.not_and'
+#align decidable.not_and_distrib Decidable.not_and_iff_or_not_not
+#align decidable.not_and_distrib' Decidable.not_and_iff_or_not_not'
/-- One of **de Morgan's laws**: the negation of a conjunction is logically equivalent to the
disjunction of the negations. -/
-theorem not_and_or : ¬(a ∧ b) ↔ ¬a ∨ ¬b := Decidable.not_and
+theorem not_and_or : ¬(a ∧ b) ↔ ¬a ∨ ¬b := Decidable.not_and_iff_or_not_not
#align not_and_distrib not_and_or
#align not_or_distrib not_or
Co-authored-by: Scott Morrison <scott.morrison@gmail.com> Co-authored-by: Eric Wieser <wieser.eric@gmail.com> Co-authored-by: Joachim Breitner <mail@joachim-breitner.de>
@@ -302,7 +302,7 @@ lemma Iff.ne_right {α β : Sort*} {a b : α} {c d : β} : (a ≠ b ↔ c = d)
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⟩
+instance : Std.Commutative Xor' := ⟨xor_comm⟩
@[simp] theorem xor_self (a : Prop) : Xor' a a = False := by simp [Xor']
#align xor_self xor_self
Mathlib.Tactic.*
(#9816)
import Mathlib.Tactic.Basic
in Mathlib.Init.Logic
and Mathlib.Logic.Basic
.import
s.@@ -6,7 +6,6 @@ Authors: Jeremy Avigad, Leonardo de Moura
import Mathlib.Init.Logic
import Mathlib.Init.Function
import Mathlib.Init.Algebra.Classes
-import Mathlib.Tactic.Basic
import Std.Util.LibraryNote
import Std.Tactic.Lint.Basic
@@ -527,23 +527,17 @@ theorem ball_mem_comm {α β} [Membership α β] {s : β} {p : α → α → Pro
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
+-- These were migrated to Std but the `@[simp]` attributes were (mysteriously?) removed.
+attribute [simp] eq_mp_eq_cast eq_mpr_eq_cast
-@[simp] theorem eq_mpr_eq_cast (h : α = β) : Eq.mpr h = cast h.symm := rfl
+#align eq_mp_eq_cast eq_mp_eq_cast
#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
@@ -569,12 +563,7 @@ 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
theorem Eq.rec_eq_cast {α : Sort _} {P : α → Sort _} {x y : α} (h : x = y) (z : P x) :
@@ -598,36 +587,16 @@ theorem heq_rec_iff_heq {C : α → Sort*} {x : β} {y : C a} {e : a = b} :
HEq x (e ▸ y) ↔ HEq x y := by subst e; rfl
#align heq_rec_iff_heq heq_rec_iff_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 congr_arg₂ := congrArg₂
#align congr_arg2 congr_arg₂
variable {β : α → Sort*} {γ : ∀ a, β a → Sort*} {δ : ∀ a b, γ a b → Sort*}
-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
Co-authored-by: Scott Morrison <scott.morrison@gmail.com> Co-authored-by: Mario Carneiro <di.gama@gmail.com>
@@ -1217,8 +1217,6 @@ protected theorem Ne.ite_ne_right_iff (h : a ≠ b) : ite P a b ≠ b ↔ P :=
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 :=
Forward port https://github.com/leanprover-community/mathlib/pull/19203
@@ -10,7 +10,7 @@ import Mathlib.Tactic.Basic
import Std.Util.LibraryNote
import Std.Tactic.Lint.Basic
-#align_import logic.basic from "leanprover-community/mathlib"@"48fb5b5280e7c81672afc9524185ae994553ebf4"
+#align_import logic.basic from "leanprover-community/mathlib"@"3365b20c2ffa7c35e47e5209b89ba9abdddf3ffe"
/-!
# Basic logic properties
@@ -281,12 +281,15 @@ theorem Iff.not_right (h : ¬a ↔ b) : a ↔ ¬b := not_not.symm.trans h.not
protected lemma Iff.ne {α β : Sort*} {a b : α} {c d : β} : (a = b ↔ c = d) → (a ≠ b ↔ c ≠ d) :=
Iff.not
+#align iff.ne Iff.ne
lemma Iff.ne_left {α β : Sort*} {a b : α} {c d : β} : (a = b ↔ c ≠ d) → (a ≠ b ↔ c = d) :=
Iff.not_left
+#align iff.ne_left Iff.ne_left
lemma Iff.ne_right {α β : Sort*} {a b : α} {c d : β} : (a ≠ b ↔ c = d) → (a = b ↔ c ≠ d) :=
Iff.not_right
+#align iff.ne_right Iff.ne_right
/-! ### Declarations about `Xor'` -/
@@ -475,7 +475,7 @@ theorem not_and_not_right : ¬(a ∧ ¬b) ↔ a → b := Decidable.not_and_not_r
#align decidable.not_and_distrib Decidable.not_and
#align decidable.not_and_distrib' Decidable.not_and'
-/-- One of de Morgan's laws: the negation of a conjunction is logically equivalent to the
+/-- One of **de Morgan's laws**: the negation of a conjunction is logically equivalent to the
disjunction of the negations. -/
theorem not_and_or : ¬(a ∧ b) ↔ ¬a ∨ ¬b := Decidable.not_and
#align not_and_distrib not_and_or
@@ -279,6 +279,15 @@ 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
+protected lemma Iff.ne {α β : Sort*} {a b : α} {c d : β} : (a = b ↔ c = d) → (a ≠ b ↔ c ≠ d) :=
+ Iff.not
+
+lemma Iff.ne_left {α β : Sort*} {a b : α} {c d : β} : (a = b ↔ c ≠ d) → (a ≠ b ↔ c = d) :=
+ Iff.not_left
+
+lemma Iff.ne_right {α β : Sort*} {a b : α} {c d : β} : (a ≠ b ↔ c = d) → (a = b ↔ c ≠ d) :=
+ Iff.not_right
+
/-! ### Declarations about `Xor'` -/
@[simp] theorem xor_true : Xor' True = Not := by
imp_false
a simp-lemma (#8744)
¬ p
to p → False
, which is probably why we didn't miss this simp-lemma too badly.simp
pull out most connectives into Filter.eventually
/Filter.frequently
, but pushed negations inside. This loops together with this new simp-lemma, and the fact that there is a unification hint . So I removed some simp attributes in Filter.Basic
.@@ -39,7 +39,7 @@ section Miscellany
-- And.decidable Or.decidable Decidable.false Xor.decidable Iff.decidable Decidable.true
-- Implies.decidable Not.decidable Ne.decidable Bool.decidableEq Decidable.toBool
-attribute [simp] cast_eq cast_heq
+attribute [simp] cast_eq cast_heq imp_false
/-- An identity function with its main argument implicit. This will be printed as `hidden` even
if it is applied to a large term, so it can be used for elision,
@@ -790,6 +790,12 @@ theorem exists_apply_eq (a : α) (b : β) : ∃ f : α → β, f a = b := ⟨fun
⟨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_exists_and_exists_and_eq_and {α β γ : Type*}
+ {f : α → β → γ} {p : α → Prop} {q : β → Prop} {r : γ → Prop} :
+ (∃ c, (∃ a, p a ∧ ∃ b, q b ∧ f a b = c) ∧ r c) ↔ ∃ a, p a ∧ ∃ b, q b ∧ r (f a b) :=
+ ⟨fun ⟨_, ⟨a, ha, b, hb, hab⟩, hc⟩ ↦ ⟨a, ha, b, hb, hab.symm ▸ hc⟩,
+ fun ⟨a, ha, b, hb, hab⟩ ↦ ⟨f a b, ⟨a, ha, b, hb, rfl⟩, hab⟩⟩
+
@[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
@@ -682,7 +682,6 @@ theorem exists_swap {p : α → β → Prop} : (∃ x y, p x y) ↔ ∃ y x, p x
#align forall_exists_index forall_exists_index
#align exists_imp_distrib exists_imp
-alias ⟨_, not_exists_of_forall_not⟩ := exists_imp
#align not_exists_of_forall_not not_exists_of_forall_not
#align Exists.some Exists.choose
This is the supremum of
along with some minor fixes from failures on nightly-testing as Mathlib master
is merged into it.
Note that some PRs for changes that are already compatible with the current toolchain and will be necessary have already been split out: #8380.
I am hopeful that in future we will be able to progressively merge adaptation PRs into a bump/v4.X.0
branch, so we never end up with a "big merge" like this. However one of these adaptation PRs (#8056) predates my new scheme for combined CI, and it wasn't possible to keep that PR viable in the meantime.
In particular this includes adjustments for the Lean PRs
We can get rid of all the
local macro_rules | `($x ^ $y) => `(HPow.hPow $x $y) -- Porting note: See issue [lean4#2220](https://github.com/leanprover/lean4/pull/2220)
macros across Mathlib (and in any projects that want to write natural number powers of reals).
Changes the default behaviour of simp
to (config := {decide := false})
. This makes simp
(and consequentially norm_num
) less powerful, but also more consistent, and less likely to blow up in long failures. This requires a variety of changes: changing some previously by simp
or norm_num
to decide
or rfl
, or adding (config := {decide := true})
.
This changed the behaviour of simp
so that simp [f]
will only unfold "fully applied" occurrences of f
. The old behaviour can be recovered with simp (config := { unfoldPartialApp := true })
. We may in future add a syntax for this, e.g. simp [!f]
; please provide feedback! In the meantime, we have made the following changes:
(config := { unfoldPartialApp := true })
in some places, to recover the old behaviour@[eqns]
to manually adjust the equation lemmas for a particular definition, recovering the old behaviour just for that definition. See #8371, where we do this for Function.comp
and Function.flip
.This change in Lean may require further changes down the line (e.g. adding the !f
syntax, and/or upstreaming the special treatment for Function.comp
and Function.flip
, and/or removing this special treatment). Please keep an open and skeptical mind about these changes!
Co-authored-by: leanprover-community-mathlib4-bot <leanprover-community-mathlib4-bot@users.noreply.github.com> Co-authored-by: Scott Morrison <scott.morrison@gmail.com> Co-authored-by: Eric Wieser <wieser.eric@gmail.com> Co-authored-by: Mauricio Collares <mauricio@collares.org>
@@ -281,7 +281,8 @@ theorem Iff.not_right (h : ¬a ↔ b) : a ↔ ¬b := not_not.symm.trans h.not
/-! ### Declarations about `Xor'` -/
-@[simp] theorem xor_true : Xor' True = Not := by simp [Xor']
+@[simp] theorem xor_true : Xor' True = Not := by
+ simp (config := { unfoldPartialApp := true }) [Xor']
#align xor_true xor_true
@[simp] theorem xor_false : Xor' False = id := by ext; simp [Xor']
@@ -885,8 +885,6 @@ theorem exists_unique_prop_of_true {p : Prop} {q : p → Prop} (h : p) : (∃! 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' :=
@@ -1145,18 +1143,9 @@ theorem dite_eq_iff' : dite P A B = c ↔ (∀ h, A h = c) ∧ ∀ h, B h = c :=
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
#align ite_eq_left_iff ite_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
theorem dite_ne_left_iff : dite P (fun _ ↦ a) B ≠ a ↔ ∃ h, a ≠ B h := by
@@ -356,11 +356,9 @@ theorem Or.imp3 (had : a → d) (hbe : b → e) (hcf : c → f) : a ∨ b ∨ c
#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
+export Classical (or_iff_not_imp_left or_iff_not_imp_right)
+#align or_iff_not_imp_left Classical.or_iff_not_imp_left
+#align or_iff_not_imp_right Classical.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
@@ -7,7 +7,6 @@ import Mathlib.Init.Logic
import Mathlib.Init.Function
import Mathlib.Init.Algebra.Classes
import Mathlib.Tactic.Basic
-import Mathlib.Tactic.LeftRight
import Std.Util.LibraryNote
import Std.Tactic.Lint.Basic
@@ -690,33 +690,20 @@ alias ⟨_, not_exists_of_forall_not⟩ := exists_imp
#align Exists.some Exists.choose
#align Exists.some_spec Exists.choose_spec
--- See Note [decidable namespace]
-protected theorem Decidable.not_forall {p : α → Prop} [Decidable (∃ x, ¬p x)]
- [∀ 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
+export Classical (not_forall)
+#align not_forall Classical.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
+export Classical (not_exists_not)
+#align not_exists_not Classical.not_exists_not
lemma forall_or_exists_not (P : α → Prop) : (∀ a, P a) ∨ ∃ a, ¬ P a := by
rw [← not_forall]; exact em _
@@ -756,8 +743,6 @@ theorem forall₃_true_iff {β : α → Sort*} {γ : ∀ a, β a → Sort*} :
-- 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 α] :
@@ -824,21 +809,13 @@ 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']
Two pairs of the form foo
and foo'
, where foo'
is the simp lemma (and hence used in many simp only
s) and foo
is not used at all.
Swap the primes, so that when it is time (now!) to upstream the lemma we actually use, it doesn't need to have a prime...
Co-authored-by: Scott Morrison <scott.morrison@gmail.com>
@@ -820,21 +820,21 @@ theorem exists_apply_eq (a : α) (b : β) : ∃ f : α → β, f a = b := ⟨fun
@[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} :
+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
+#align forall_apply_eq_imp_iff forall_apply_eq_imp_iff'
-@[simp] theorem forall_apply_eq_imp_iff' {f : α → β} {p : β → Prop} :
+@[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'
+#align forall_apply_eq_imp_iff' forall_apply_eq_imp_iff
-theorem forall_eq_apply_imp_iff {f : α → β} {p : β → Prop} :
+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
+#align forall_eq_apply_imp_iff forall_eq_apply_imp_iff'
-@[simp] theorem forall_eq_apply_imp_iff' {f : α → β} {p : β → Prop} :
+@[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'
+#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) :=
@@ -718,6 +718,12 @@ protected theorem Decidable.not_exists_not [∀ x, Decidable (p x)] : (¬∃ x,
theorem not_exists_not : (¬∃ x, ¬p x) ↔ ∀ x, p x := Decidable.not_exists_not
#align not_exists_not not_exists_not
+lemma forall_or_exists_not (P : α → Prop) : (∀ a, P a) ∨ ∃ a, ¬ P a := by
+ rw [← not_forall]; exact em _
+
+lemma exists_or_forall_not (P : α → Prop) : (∃ a, P a) ∨ ∀ a, ¬ P a := by
+ rw [← not_exists]; exact em _
+
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)⟩
And fix some names in comments where this revealed issues
@@ -984,7 +984,7 @@ namespace Classical
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`. -/
+/-- Any prop `p` is decidable classically. A shorthand for `Classical.propDecidable`. -/
noncomputable def dec (p : Prop) : Decidable p := by infer_instance
#align classical.dec Classical.dec
@@ -383,6 +383,10 @@ theorem imp_iff_or_not : b → a ↔ a ∨ ¬b := Decidable.imp_iff_or_not
theorem not_imp_not : ¬a → ¬b ↔ b → a := Decidable.not_imp_not
#align not_imp_not not_imp_not
+@[simp]
+theorem imp_and_neg_imp_iff (p q : Prop) : (p → q) ∧ (¬p → q) ↔ q := by
+ rw [imp_iff_or_not, imp_iff_or_not, not_not, ← or_and_left, not_and_self_iff, or_false_iff]
+
/-- 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
@@ -563,6 +567,9 @@ 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
+theorem Eq.rec_eq_cast {α : Sort _} {P : α → Sort _} {x y : α} (h : x = y) (z : P x) :
+ h ▸ z = cast (congr_arg P h) z := by induction h; rfl
+
--Porting note: new theorem. More general version of `eqRec_heq`
theorem eqRec_heq' {α : Sort u_1} {a' : α} {motive : (a : α) → a' = a → Sort u}
(p : motive a' (rfl : a' = a')) {a : α} (t : a' = a) :
@@ -1280,4 +1280,21 @@ theorem ite_ite_comm (h : P → ¬Q) :
dite_dite_comm P Q h
#align ite_ite_comm ite_ite_comm
+variable {R : Prop}
+
+theorem ite_prop_iff_or : (if P then Q else R) ↔ (P ∧ Q ∨ ¬ P ∧ R) := by
+ by_cases p : P <;> simp [p]
+
+theorem dite_prop_iff_or {Q : P → Prop} {R : ¬P → Prop} [Decidable P] :
+ dite P Q R ↔ (∃ p, Q p) ∨ (∃ p, R p) := by
+ by_cases h : P <;> simp [h, exists_prop_of_false, exists_prop_of_true]
+
+-- TODO make this a simp lemma in a future PR
+theorem ite_prop_iff_and : (if P then Q else R) ↔ ((P → Q) ∧ (¬ P → R)) := by
+ by_cases p : P <;> simp [p]
+
+theorem dite_prop_iff_and {Q : P → Prop} {R : ¬P → Prop} [Decidable P] :
+ dite P Q R ↔ (∀ h, Q h) ∧ (∀ h, R h) := by
+ by_cases h : P <;> simp [h, forall_prop_of_false, forall_prop_of_true]
+
end ite
@@ -162,7 +162,7 @@ instance : IsRefl Prop Iff := ⟨Iff.refl⟩
instance : IsTrans Prop Iff := ⟨fun _ _ _ ↦ Iff.trans⟩
-alias imp_congr ← Iff.imp
+alias Iff.imp := imp_congr
#align iff.imp Iff.imp
@[simp] theorem eq_true_eq_id : Eq True = id := by
@@ -185,13 +185,13 @@ protected theorem Function.mt : (a → b) → ¬b → ¬a := mt
/-! ### Declarations about `not` -/
-alias Decidable.em ← dec_em
+alias dec_em := Decidable.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
+alias em := Classical.em
#align em em
theorem em' (p : Prop) : ¬p ∨ p := (em p).symm
@@ -220,7 +220,7 @@ 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
+alias by_contra := by_contradiction
#align by_contra by_contra
library_note "decidable namespace"/--
@@ -256,7 +256,7 @@ theorem not_ne_iff : ¬a ≠ b ↔ a = b := not_not
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
+alias Not.decidable_imp_symm := Decidable.not_imp_symm
#align not.decidable_imp_symm Not.decidable_imp_symm
theorem Not.imp_symm : (¬a → b) → ¬b → a := Not.decidable_imp_symm
@@ -271,7 +271,7 @@ theorem not_imp_comm : ¬a → b ↔ ¬b → a := Decidable.not_imp_comm
theorem Imp.swap : a → b → c ↔ b → a → c := ⟨Function.swap, Function.swap⟩
#align imp.swap Imp.swap
-alias not_congr ← Iff.not
+alias Iff.not := not_congr
#align iff.not Iff.not
theorem Iff.not_left (h : a ↔ ¬b) : ¬a ↔ b := h.not.trans not_not
@@ -310,14 +310,14 @@ protected theorem Xor'.or (h : Xor' a b) : a ∨ b := h.imp And.left And.left
/-! ### Declarations about `and` -/
-alias and_congr ← Iff.and
+alias Iff.and := and_congr
#align iff.and Iff.and
#align and_congr_left and_congr_leftₓ -- reorder implicits
#align and_congr_right' and_congr_right'ₓ -- reorder implicits
#align and.right_comm and_right_comm
#align and_and_distrib_left and_and_left
#align and_and_distrib_right and_and_right
-alias and_rotate ↔ And.rotate _
+alias ⟨And.rotate, _⟩ := and_rotate
#align and.rotate And.rotate
#align and.congr_right_iff and_congr_right_iff
#align and.congr_left_iff and_congr_left_iffₓ -- reorder implicits
@@ -327,12 +327,12 @@ theorem and_symm_left (a b : α) (p : Prop) : a = b ∧ p ↔ b = a ∧ p := by
/-! ### Declarations about `or` -/
-alias or_congr ← Iff.or
+alias Iff.or := or_congr
#align iff.or Iff.or
#align or_congr_left' or_congr_left
#align or_congr_right' or_congr_rightₓ -- reorder implicits
#align or.right_comm or_right_comm
-alias or_rotate ↔ Or.rotate _
+alias ⟨Or.rotate, _⟩ := or_rotate
#align or.rotate Or.rotate
@[deprecated Or.imp]
@@ -409,7 +409,7 @@ theorem or_congr_right' (h : ¬a → (b ↔ c)) : a ∨ b ↔ a ∨ c := Decidab
/-! Declarations about `iff` -/
-alias iff_congr ← Iff.iff
+alias Iff.iff := iff_congr
#align iff.iff Iff.iff
-- @[simp] -- FIXME simp ignores proof rewrites
@@ -494,8 +494,8 @@ end Propositional
/-! ### Declarations about equality -/
-alias ne_of_mem_of_not_mem ← Membership.mem.ne_of_not_mem
-alias ne_of_mem_of_not_mem' ← Membership.mem.ne_of_not_mem'
+alias Membership.mem.ne_of_not_mem := ne_of_mem_of_not_mem
+alias Membership.mem.ne_of_not_mem' := ne_of_mem_of_not_mem'
#align has_mem.mem.ne_of_not_mem Membership.mem.ne_of_not_mem
#align has_mem.mem.ne_of_not_mem' Membership.mem.ne_of_not_mem'
@@ -591,7 +591,7 @@ theorem Eq.congr_left {x y z : α} (h : x = y) : x = z ↔ y = z := by rw [h]
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₂
+alias congr_arg₂ := congrArg₂
#align congr_arg2 congr_arg₂
variable {β : α → Sort*} {γ : ∀ a, β a → Sort*} {δ : ∀ a b, γ a b → Sort*}
@@ -677,7 +677,7 @@ theorem exists_swap {p : α → β → Prop} : (∃ x y, p x y) ↔ ∃ y x, p x
#align forall_exists_index forall_exists_index
#align exists_imp_distrib exists_imp
-alias exists_imp ↔ _ not_exists_of_forall_not
+alias ⟨_, not_exists_of_forall_not⟩ := exists_imp
#align not_exists_of_forall_not not_exists_of_forall_not
#align Exists.some Exists.choose
Autoimplicits are highly controversial and also defeat the performance-improving work in #6474.
The intent of this PR is to make autoImplicit
opt-in on a per-file basis, by disabling it in the lakefile and enabling it again with set_option autoImplicit true
in the few files that rely on it.
That also keeps this PR small, as opposed to attempting to "fix" files to not need it any more.
I claim that many of the uses of autoImplicit
in these files are accidental; situations such as:
variables
are in scope, but pasting the lemma in the wrong sectionHaving set_option autoImplicit false
as the default prevents these types of mistake being made in the 90% of files where autoImplicit
s are not used at all, and causes them to be caught by CI during review.
I think there were various points during the port where we encouraged porters to delete the universes u v
lines; I think having autoparams for universe variables only would cover a lot of the cases we actually use them, while avoiding any real shortcomings.
A Zulip poll (after combining overlapping votes accordingly) was in favor of this change with 5:5:18
as the no:dontcare:yes
vote ratio.
While this PR was being reviewed, a handful of files gained some more likely-accidental autoImplicits. In these places, set_option autoImplicit true
has been placed locally within a section, rather than at the top of the file.
@@ -24,6 +24,8 @@ Theorems that require decidability hypotheses are in the namespace `Decidable`.
Classical versions are in the namespace `Classical`.
-/
+set_option autoImplicit true
+
open Function
attribute [local instance 10] Classical.propDecidable
Type _
and Sort _
(#6499)
We remove all possible occurences of Type _
and Sort _
in favor of Type*
and Sort*
.
This has nice performance benefits.
@@ -43,14 +43,14 @@ attribute [simp] cast_eq cast_heq
/-- An identity function with its main argument implicit. This will be printed as `hidden` even
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
+@[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)
#align decidable_eq_of_subsingleton decidableEq_of_subsingleton
-instance (α : Sort _) [Subsingleton α] (p : α → Prop) : Subsingleton (Subtype p) :=
+instance (α : Sort*) [Subsingleton α] (p : α → Prop) : Subsingleton (Subtype p) :=
⟨fun ⟨x, _⟩ ⟨y, _⟩ ↦ by cases Subsingleton.elim x y; rfl⟩
#align pempty PEmpty
@@ -60,7 +60,7 @@ theorem congr_heq {α β γ : Sort _} {f : α → γ} {g : β → γ} {x : α} {
cases h₂; cases h₁; rfl
#align congr_heq congr_heq
-theorem congr_arg_heq {α} {β : α → Sort _} (f : ∀ a, β a) :
+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
@@ -73,11 +73,11 @@ theorem ULift.down_injective {α : Sort _} : Function.Injective (@ULift.down α)
⟨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 α)
+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 :=
+@[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
@@ -89,7 +89,7 @@ theorem PLift.down_injective {α : Sort _} : Function.Injective (@PLift.down α)
⟨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 :=
+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
@@ -128,8 +128,8 @@ theorem fact_iff {p : Prop} : Fact p ↔ p := ⟨fun h ↦ h.1, fun h ↦ ⟨h
#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₂)
+@[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₂
@@ -137,12 +137,12 @@ theorem fact_iff {p : Prop} : Fact p ↔ p := ⟨fun h ↦ h.1, fun h ↦ ⟨h
-- /-- If `x : α . tac_name` then `x.out : α`. These are definitionally equal, but this can
-- nevertheless be useful for various reasons, e.g. to apply further projection notation or in an
-- argument to `simp`. -/
--- def autoParam'.out {α : Sort _} {n : Name} (x : autoParam' α n) : α := x
+-- def autoParam'.out {α : Sort*} {n : Name} (x : autoParam' α n) : α := x
-- /-- If `x : α := d` then `x.out : α`. These are definitionally equal, but this can
-- nevertheless be useful for various reasons, e.g. to apply further projection notation or in an
-- argument to `simp`. -/
--- def optParam.out {α : Sort _} {d : α} (x : α := d) : α := x
+-- def optParam.out {α : Sort*} {d : α} (x : α := d) : α := x
end Miscellany
@@ -511,7 +511,7 @@ theorem ball_mem_comm {α β} [Membership α β] {s : β} {p : α → α → Pro
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 :=
+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
@@ -567,15 +567,15 @@ theorem eqRec_heq' {α : Sort u_1} {a' : α} {motive : (a : α) → a' = a → S
HEq (@Eq.rec α a' motive p a t) p :=
by subst t; rfl
-theorem rec_heq_of_heq {C : α → Sort _} {x : C a} {y : β} (e : a = b) (h : HEq x y) :
+theorem rec_heq_of_heq {C : α → Sort*} {x : C a} {y : β} (e : a = b) (h : HEq x y) :
HEq (e ▸ x) y := by subst e; exact h
#align rec_heq_of_heq rec_heq_of_heq
-theorem rec_heq_iff_heq {C : α → Sort _} {x : C a} {y : β} {e : a = b} :
+theorem rec_heq_iff_heq {C : α → Sort*} {x : C a} {y : β} {e : a = b} :
HEq (e ▸ x) y ↔ HEq x y := by subst e; rfl
#align rec_heq_iff_heq rec_heq_iff_heq
-theorem heq_rec_iff_heq {C : α → Sort _} {x : β} {y : C a} {e : a = b} :
+theorem heq_rec_iff_heq {C : α → Sort*} {x : β} {y : C a} {e : a = b} :
HEq x (e ▸ y) ↔ HEq x y := by subst e; rfl
#align heq_rec_iff_heq heq_rec_iff_heq
@@ -592,7 +592,7 @@ theorem Eq.congr_right {x y z : α} (h : x = y) : z = x ↔ z = y := by rw [h]
alias congrArg₂ ← congr_arg₂
#align congr_arg2 congr_arg₂
-variable {β : α → Sort _} {γ : ∀ a, β a → Sort _} {δ : ∀ a b, γ a b → Sort _}
+variable {β : α → Sort*} {γ : ∀ a, β a → Sort*} {δ : ∀ a b, γ a b → Sort*}
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 _) _
@@ -619,8 +619,8 @@ end Equality
section Quantifiers
section Dependent
-variable {β : α → Sort _} {γ : ∀ a, β a → Sort _} {δ : ∀ a b, γ a b → Sort _}
- {ε : ∀ a b c, δ a b c → Sort _}
+variable {β : α → Sort*} {γ : ∀ a, β a → Sort*} {δ : ∀ a b, γ a b → Sort*}
+ {ε : ∀ a b c, δ a b c → Sort*}
theorem pi_congr {β' : α → Sort _} (h : ∀ a, β a = β' a) : (∀ a, β a) = ∀ a, β' a :=
(funext h : β = β') ▸ rfl
@@ -651,20 +651,20 @@ theorem Exists₃.imp {p q : ∀ a b, γ a b → Prop} (h : ∀ a b c, p a b c
end Dependent
-variable {κ : ι → Sort _} {p q : α → Prop}
+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} :
+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 :=
+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
@@ -726,11 +726,11 @@ theorem forall_true_iff' (h : ∀ a, p a ↔ True) : (∀ a, p a) ↔ True :=
#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
+theorem forall₂_true_iff {β : α → Sort*} : (∀ a, β a → True) ↔ True := by simp
#align forall_2_true_iff forall₂_true_iff
-- This is not marked `@[simp]` because `implies_true : (α → True) = True` works
-theorem forall₃_true_iff {β : α → Sort _} {γ : ∀ a, β a → Sort _} :
+theorem forall₃_true_iff {β : α → Sort*} {γ : ∀ a, β a → Sort*} :
(∀ (a) (b : β a), γ a b → True) ↔ True := by simp
#align forall_3_true_iff forall₃_true_iff
@@ -831,7 +831,7 @@ theorem forall_eq_apply_imp_iff {f : α → β} {p : β → Prop} :
#align exists_comm exists_comm
-theorem exists₂_comm {κ₁ : ι₁ → Sort _} {κ₂ : ι₂ → Sort _} {p : ∀ i₁, κ₁ i₁ → ∀ i₂, κ₂ i₂ → Prop} :
+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
@@ -940,7 +940,7 @@ theorem forall_prop_congr' {p p' : Prop} {q q' : p → Prop} (hq : ∀ h, q h
forall_prop_of_true _
#align forall_true_left forall_true_left
-theorem ExistsUnique.elim₂ {α : Sort _} {p : α → Sort _} [∀ x, Subsingleton (p x)]
+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)
(h₁ : ∀ (x) (h : p x), q x h → (∀ (y) (hy : p y), q y hy → y = x) → b) : b := by
simp only [exists_unique_iff_exists] at h₂
@@ -948,19 +948,19 @@ theorem ExistsUnique.elim₂ {α : Sort _} {p : α → Sort _} [∀ x, Subsingle
exact fun x ⟨hxp, hxq⟩ H ↦ h₁ x hxp hxq fun y hyp hyq ↦ H y ⟨hyp, hyq⟩
#align exists_unique.elim2 ExistsUnique.elim₂
-theorem ExistsUnique.intro₂ {α : Sort _} {p : α → Sort _} [∀ x, Subsingleton (p x)]
+theorem ExistsUnique.intro₂ {α : Sort*} {p : α → Sort*} [∀ x, Subsingleton (p x)]
{q : ∀ (x : α) (_ : p x), Prop} (w : α) (hp : p w) (hq : q w hp)
(H : ∀ (y) (hy : p y), q y hy → y = w) : ∃! (x : _) (hx : p x), q x hx := by
simp only [exists_unique_iff_exists]
exact ExistsUnique.intro w ⟨hp, hq⟩ fun y ⟨hyp, hyq⟩ ↦ H y hyp hyq
#align exists_unique.intro2 ExistsUnique.intro₂
-theorem ExistsUnique.exists₂ {α : Sort _} {p : α → Sort _} {q : ∀ (x : α) (_ : p x), Prop}
+theorem ExistsUnique.exists₂ {α : Sort*} {p : α → Sort*} {q : ∀ (x : α) (_ : p x), Prop}
(h : ∃! (x : _) (hx : p x), q x hx) : ∃ (x : _) (hx : p x), q x hx :=
h.exists.imp fun _ hx ↦ hx.exists
#align exists_unique.exists2 ExistsUnique.exists₂
-theorem ExistsUnique.unique₂ {α : Sort _} {p : α → Sort _} [∀ x, Subsingleton (p x)]
+theorem ExistsUnique.unique₂ {α : Sort*} {p : α → Sort*} [∀ x, Subsingleton (p x)]
{q : ∀ (x : α) (_ : p x), Prop} (h : ∃! (x : _) (hx : p x), q x hx) {y₁ y₂ : α}
(hpy₁ : p y₁) (hqy₁ : q y₁ hpy₁) (hpy₂ : p y₂) (hqy₂ : q y₂ hpy₂) : y₁ = y₂ := by
simp only [exists_unique_iff_exists] at h
@@ -998,22 +998,22 @@ 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)
+theorem some_spec₂ {α : Sort*} {p : α → Prop} {h : ∃ a, p a} (q : α → Prop)
(hpq : ∀ a, p a → q a) : q (choose h) := hpq _ <| choose_spec _
#align classical.some_spec2 Classical.some_spec₂
/-- A version of `Classical.indefiniteDescription` which is definitionally equal to a pair -/
-noncomputable def subtype_of_exists {α : Type _} {P : α → Prop} (h : ∃ x, P x) : { x // P x } :=
+noncomputable def subtype_of_exists {α : Type*} {P : α → Prop} (h : ∃ x, P x) : { x // P x } :=
⟨Classical.choose h, Classical.choose_spec h⟩
#align classical.subtype_of_exists Classical.subtype_of_exists
/-- A version of `byContradiction` that uses types instead of propositions. -/
-protected noncomputable def byContradiction' {α : Sort _} (H : ¬(α → False)) : α :=
+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 α → α :=
+def choice_of_byContradiction' {α : Sort*} (contra : ¬(α → False) → α) : Nonempty α → α :=
fun H ↦ contra H.elim
#align classical.choice_of_by_contradiction' Classical.choice_of_byContradiction'
@@ -1134,7 +1134,7 @@ end BoundedQuantifiers
section ite
-variable {σ : α → Sort _} (f : α → β) {P Q : Prop} [Decidable P] [Decidable Q]
+variable {σ : α → Sort*} (f : α → β) {P Q : Prop} [Decidable P] [Decidable Q]
{a b c : α} {A : P → α} {B : ¬P → α}
theorem dite_eq_iff : dite P A B = c ↔ (∃ h, A h = c) ∨ ∃ h, B h = c := by
@@ -2,11 +2,6 @@
Copyright (c) 2016 Jeremy Avigad. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Jeremy Avigad, Leonardo de Moura
-
-! This file was ported from Lean 3 source module logic.basic
-! leanprover-community/mathlib commit 48fb5b5280e7c81672afc9524185ae994553ebf4
-! Please do not edit these lines, except to modify the commit id
-! if you have ported upstream changes.
-/
import Mathlib.Init.Logic
import Mathlib.Init.Function
@@ -16,6 +11,8 @@ import Mathlib.Tactic.LeftRight
import Std.Util.LibraryNote
import Std.Tactic.Lint.Basic
+#align_import logic.basic from "leanprover-community/mathlib"@"48fb5b5280e7c81672afc9524185ae994553ebf4"
+
/-!
# Basic logic properties
register_simp_attr
s to 1 file (#5681)
There are slight differences between mathlib3
and mathlib4
(different set of attributes, different lemmas are in core/std), so I redid the same refactor instead of forward-porting changes.
mathlib3 PR: leanprover-community/mathlib#19223
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Jeremy Avigad, Leonardo de Moura
! This file was ported from Lean 3 source module logic.basic
-! leanprover-community/mathlib commit d2d8742b0c21426362a9dacebc6005db895ca963
+! leanprover-community/mathlib commit 48fb5b5280e7c81672afc9524185ae994553ebf4
! Please do not edit these lines, except to modify the commit id
! if you have ported upstream changes.
-/
@@ -718,6 +718,7 @@ theorem forall_imp_iff_exists_imp [ha : Nonempty α] : (∀ x, p x) → b ↔
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
+@[mfld_simps]
theorem forall_true_iff : (α → True) ↔ True := imp_true_iff _
#align forall_true_iff forall_true_iff
@@ -273,11 +273,13 @@ theorem Imp.swap : a → b → c ↔ b → a → c := ⟨Function.swap, Function
#align imp.swap Imp.swap
alias not_congr ← Iff.not
+#align iff.not Iff.not
+
theorem Iff.not_left (h : a ↔ ¬b) : ¬a ↔ b := h.not.trans not_not
+#align iff.not_left Iff.not_left
+
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'` -/
@@ -293,29 +295,33 @@ theorem xor_comm (a b) : Xor' a b = Xor' b a := by simp [Xor', and_comm, or_comm
instance : IsCommutative Prop Xor' := ⟨xor_comm⟩
@[simp] theorem xor_self (a : Prop) : Xor' a a = False := by simp [Xor']
+#align xor_self xor_self
+
@[simp] theorem xor_not_left : Xor' (¬a) b ↔ (a ↔ b) := by by_cases a <;> simp [*]
+#align xor_not_left xor_not_left
+
@[simp] theorem xor_not_right : Xor' a (¬b) ↔ (a ↔ b) := by by_cases a <;> simp [*]
+#align xor_not_right xor_not_right
+
theorem xor_not_not : Xor' (¬a) (¬b) ↔ Xor' a b := by simp [Xor', or_comm, and_comm]
+#align xor_not_not xor_not_not
+
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` -/
alias and_congr ← Iff.and
+#align iff.and Iff.and
#align and_congr_left and_congr_leftₓ -- reorder implicits
#align and_congr_right' and_congr_right'ₓ -- reorder implicits
#align and.right_comm and_right_comm
#align and_and_distrib_left and_and_left
#align and_and_distrib_right and_and_right
alias and_rotate ↔ And.rotate _
+#align 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]
@@ -323,12 +329,12 @@ theorem and_symm_left (a b : α) (p : Prop) : a = b ∧ p ↔ b = a ∧ p := by
/-! ### Declarations about `or` -/
alias or_congr ← Iff.or
+#align iff.or Iff.or
#align or_congr_left' or_congr_left
#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₁
@@ -477,11 +483,13 @@ theorem and_iff_not_or_not : a ∧ b ↔ ¬(¬a ∨ ¬b) := Decidable.and_iff_no
#align not_xor not_xor
theorem xor_iff_not_iff (P Q : Prop) : Xor' P Q ↔ ¬ (P ↔ Q) := (not_xor P Q).not_right
+#align xor_iff_not_iff xor_iff_not_iff
+
theorem xor_iff_iff_not : Xor' a b ↔ (a ↔ ¬b) := by simp only [← @xor_not_right a, not_not]
+#align xor_iff_iff_not xor_iff_iff_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
@@ -1159,9 +1167,10 @@ theorem ite_eq_iff' : ite P a b = c ↔ (P → a = c) ∧ (¬P → b = c) := dit
#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
+#align ite_eq_left_iff ite_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]
Heavily based on the code from the sphere eversion project.
Co-authored-by: @fpvandoorn
@@ -871,6 +871,13 @@ theorem Exists.snd {b : Prop} {p : b → Prop} : ∀ h : Exists p, p h.fst
| ⟨_, h⟩ => h
#align Exists.snd Exists.snd
+theorem Prop.exists_iff {p : Prop → Prop} : (∃ h, p h) ↔ p False ∨ p True :=
+ ⟨fun ⟨h₁, h₂⟩ ↦ by_cases (fun H : h₁ ↦ .inr <| by simpa only [H] using h₂)
+ (fun H ↦ .inl <| by simpa only [H] using h₂), fun h ↦ h.elim (.intro _) (.intro _)⟩
+
+theorem Prop.forall_iff {p : Prop → Prop} : (∀ h, p h) ↔ p False ∧ p True :=
+ ⟨fun H ↦ ⟨H _, H _⟩, fun ⟨h₁, h₂⟩ h ↦ by by_cases H : h <;> simpa only [H]⟩
+
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
@@ -1133,7 +1133,7 @@ theorem ite_eq_iff : ite P a b = c ↔ P ∧ a = c ∨ ¬P ∧ b = c :=
#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)
+ eq_comm.trans <| ite_eq_iff.trans <| (Iff.rfl.and eq_comm).or (Iff.rfl.and eq_comm)
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 ↦
Co-authored-by: Scott Morrison <scott.morrison@gmail.com> Co-authored-by: Mario Carneiro <di.gama@gmail.com> Co-authored-by: Floris van Doorn <fpvdoorn@gmail.com> Co-authored-by: Jeremy Tan Jie Rui <reddeloostw@gmail.com> Co-authored-by: Alex J Best <alex.j.best@gmail.com>
@@ -823,8 +823,6 @@ theorem forall_eq_apply_imp_iff {f : α → β} {p : β → Prop} :
@[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} :
fix-comments.py
on all files.@@ -279,7 +279,7 @@ theorem Iff.not_right (h : ¬a ↔ b) : a ↔ ¬b := not_not.symm.trans h.not
#align iff.not_left Iff.not_left
#align iff.not Iff.not
-/-! ### Declarations about `xor` -/
+/-! ### Declarations about `Xor'` -/
@[simp] theorem xor_true : Xor' True = Not := by simp [Xor']
#align xor_true xor_true
@@ -905,7 +905,7 @@ theorem exists_prop_congr' {p p' : Prop} {q q' : p → Prop} (hq : ∀ h, q h
propext (exists_prop_congr hq hp)
#align exists_prop_congr' exists_prop_congr'
-/-- See `IsEmpty.exists_iff` for the `false` version. -/
+/-- 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
@@ -924,7 +924,7 @@ theorem forall_prop_congr' {p p' : Prop} {q q' : p → Prop} (hq : ∀ h, q h
propext (forall_prop_congr hq hp)
#align forall_prop_congr' forall_prop_congr'
-/-- See `IsEmpty.forall_iff` for the `false` version. -/
+/-- 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
Other changes:
Data/Dfinsupp/Lex
to use explicit (i : ι)
;Logic.Basic
.@@ -487,8 +487,11 @@ end Propositional
/-! ### Declarations about equality -/
-alias ne_of_mem_of_not_mem ← Membership.Mem.ne_of_not_mem
-alias ne_of_mem_of_not_mem' ← Membership.Mem.ne_of_not_mem'
+alias ne_of_mem_of_not_mem ← Membership.mem.ne_of_not_mem
+alias ne_of_mem_of_not_mem' ← Membership.mem.ne_of_not_mem'
+
+#align has_mem.mem.ne_of_not_mem Membership.mem.ne_of_not_mem
+#align has_mem.mem.ne_of_not_mem' Membership.mem.ne_of_not_mem'
section Equality
The diffs of these files are solely generalising Types α
from Type
to Sort
, but they are already Sort
in mathlib4.
These SHA just need updating from backport leanprover-community/mathlib#18543
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Jeremy Avigad, Leonardo de Moura
! This file was ported from Lean 3 source module logic.basic
-! leanprover-community/mathlib commit 13cd3e89b30352d5b1b7349f5537ea18ba878e40
+! leanprover-community/mathlib commit d2d8742b0c21426362a9dacebc6005db895ca963
! Please do not edit these lines, except to modify the commit id
! if you have ported upstream changes.
-/
logic.basic
@feb99064803fd3108e37c18b0f77d0a8344677a3
..13cd3e89b30352d5b1b7349f5537ea18ba878e40
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Jeremy Avigad, Leonardo de Moura
! This file was ported from Lean 3 source module logic.basic
-! leanprover-community/mathlib commit feb99064803fd3108e37c18b0f77d0a8344677a3
+! leanprover-community/mathlib commit 13cd3e89b30352d5b1b7349f5537ea18ba878e40
! Please do not edit these lines, except to modify the commit id
! if you have ported upstream changes.
-/
@@ -560,9 +560,17 @@ theorem eqRec_heq' {α : Sort u_1} {a' : α} {motive : (a : α) → a' = a → S
by subst t; rfl
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
+ HEq (e ▸ x) y := by subst e; exact h
#align rec_heq_of_heq rec_heq_of_heq
+theorem rec_heq_iff_heq {C : α → Sort _} {x : C a} {y : β} {e : a = b} :
+ HEq (e ▸ x) y ↔ HEq x y := by subst e; rfl
+#align rec_heq_iff_heq rec_heq_iff_heq
+
+theorem heq_rec_iff_heq {C : α → Sort _} {x : β} {y : C a} {e : a = b} :
+ HEq x (e ▸ y) ↔ HEq x y := by subst e; rfl
+#align heq_rec_iff_heq heq_rec_iff_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
Update some SHAs of files that changed in mathlib3.
These 17 files need mainly only updated SHA as they've been only touched by backports or already have been forward-ported.
The relevant changes are:
Co-authored-by: Eric Wieser <wieser.eric@gmail.com>
@@ -4,7 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE.
Authors: Jeremy Avigad, Leonardo de Moura
! This file was ported from Lean 3 source module logic.basic
-! leanprover-community/mathlib commit 1c521b4fb909320eca16b2bb6f8b5b0490b1cb5e
+! leanprover-community/mathlib commit feb99064803fd3108e37c18b0f77d0a8344677a3
! Please do not edit these lines, except to modify the commit id
! if you have ported upstream changes.
-/
@@ -296,8 +296,8 @@ instance : IsCommutative Prop Xor' := ⟨xor_comm⟩
@[simp] theorem xor_not_left : Xor' (¬a) b ↔ (a ↔ b) := by by_cases a <;> simp [*]
@[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
+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
@@ -244,11 +244,6 @@ it is better to use explicitly introduced ones rather than allowing Lean to auto
classical ones, as these may cause instance mismatch errors later.
-/
-/-- The Double Negation Theorem: `¬ ¬ P` is equivalent to `P`.
-The left-to-right direction, double negation elimination (DNE),
-is classically true but not constructively. -/
-add_decl_doc Classical.not_not -- TODO: move to Std
-
export Classical (not_not)
attribute [simp] not_not
#align not_not Classical.not_not
@@ -661,9 +661,6 @@ theorem exists_swap {p : α → β → Prop} : (∃ x y, p x y) ↔ ∃ y x, p x
⟨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
@@ -247,8 +247,11 @@ classical ones, as these may cause instance mismatch errors later.
/-- The Double Negation Theorem: `¬ ¬ P` is equivalent to `P`.
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
+add_decl_doc Classical.not_not -- TODO: move to Std
+
+export Classical (not_not)
+attribute [simp] not_not
+#align not_not Classical.not_not
theorem of_not_not : ¬¬a → a := by_contra
#align of_not_not of_not_not
This PR is the result of a slight variant on the following "algorithm"
_
and make all uppercase letters into lowercase_
and make all uppercase letters into lowercase(original_lean3_name, OriginalLean4Name)
#align
statement just before the next empty line#align
statement to have been inserted too early)@@ -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
List.perm_replicate_append_replicate
(#1509)
This is a forward-port of leanprover-community/mathlib#18126
Also golf a proof.
@@ -613,8 +613,14 @@ theorem exists_unique_const (α) [i : Nonempty α] [Subsingleton α] :
#align exists_and_distrib_left exists_and_left
#align exists_and_distrib_right exists_and_right
-theorem and_forall_ne (a : α) : (p a ∧ ∀ (b) (_ : b ≠ a), p b) ↔ ∀ b, p b := by
- simp only [← @forall_eq _ p a, ← forall_and, ← or_imp, Classical.em, forall_const]
+theorem Decidable.and_forall_ne [DecidableEq α] (a : α) {p : α → Prop} :
+ (p a ∧ ∀ b, b ≠ a → p b) ↔ ∀ b, p b := by
+ simp only [← @forall_eq _ p a, ← forall_and, ← or_imp, Decidable.em, forall_const]
+#align decidable.and_forall_ne Decidable.and_forall_ne
+
+theorem and_forall_ne (a : α) : (p a ∧ ∀ b, b ≠ a → p b) ↔ ∀ b, p b :=
+ Decidable.and_forall_ne a
+#align and_forall_ne and_forall_ne
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
@@ -379,8 +379,8 @@ theorem not_and_not_right : ¬(a ∧ ¬b) ↔ a → b := Decidable.not_and_not_r
/-! ### De Morgan's laws -/
-#align Decidable.not_and_distrib Decidable.not_and
-#align Decidable.not_and_distrib' Decidable.not_and'
+#align decidable.not_and_distrib Decidable.not_and
+#align decidable.not_and_distrib' Decidable.not_and'
/-- One of de Morgan's laws: the negation of a conjunction is logically equivalent to the
disjunction of the negations. -/
@@ -93,13 +93,13 @@ Certain propositions should not be treated as a class globally,
but sometimes it is very convenient to be able to use the type class system
in specific circumstances.
-For example, `Zmod p` is a field if and only if `p` is a prime number.
+For example, `ZMod p` is a field if and only if `p` is a prime number.
In order to be able to find this field instance automatically by type class search,
we have to turn `p.prime` into an instance implicit assumption.
On the other hand, making `Nat.prime` a class would require a major refactoring of the library,
and it is questionable whether making `Nat.prime` a class is desirable at all.
-The compromise is to add the assumption `[Fact p.prime]` to `Zmod.field`.
+The compromise is to add the assumption `[Fact p.prime]` to `ZMod.field`.
In particular, this class is not intended for turning the type class system
into an automated theorem prover for first order logic. -/
@@ -628,6 +628,9 @@ theorem Ne.ne_or_ne {x y : α} (z : α) (h : x ≠ y) : x ≠ z ∨ y ≠ z :=
-- @[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⟩
+-- porting note: an alternative workaround theorem:
+theorem exists_apply_eq (a : α) (b : β) : ∃ f : α → β, f a = b := ⟨fun _ ↦ b, rfl⟩
+
@[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⟩⟩
@@ -456,6 +456,12 @@ theorem heq_of_cast_eq : ∀ (e : α = β) (_ : cast e a = a'), HEq a a'
theorem cast_eq_iff_heq : cast e a = a' ↔ HEq a a' :=
⟨heq_of_cast_eq _, fun h ↦ by cases h; rfl⟩
+--Porting note: new theorem. More general version of `eqRec_heq`
+theorem eqRec_heq' {α : Sort u_1} {a' : α} {motive : (a : α) → a' = a → Sort u}
+ (p : motive a' (rfl : a' = a')) {a : α} (t : a' = a) :
+ HEq (@Eq.rec α a' motive p a t) p :=
+ by subst t; rfl
+
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
Co-authored-by: Scott Morrison <scott.morrison@gmail.com>
@@ -394,7 +394,7 @@ theorem or_iff_not_and_not : a ∨ b ↔ ¬(¬a ∧ ¬b) := Decidable.or_iff_not
theorem and_iff_not_or_not : a ∧ b ↔ ¬(¬a ∨ ¬b) := Decidable.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, iff_self]
+ simp only [not_and, Xor', not_or, not_not, ← iff_iff_implies_and_implies]
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]
@@ -564,7 +564,7 @@ theorem not_forall_not : (¬∀ x, ¬p x) ↔ ∃ x, p x := Decidable.not_forall
-- 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, iff_self]
+ simp only [not_exists, Decidable.not_not]
theorem not_exists_not : (¬∃ x, ¬p x) ↔ ∀ x, p x := Decidable.not_exists_not
@@ -608,7 +608,7 @@ theorem exists_unique_const (α) [i : Nonempty α] [Subsingleton α] :
#align exists_and_distrib_right exists_and_right
theorem and_forall_ne (a : α) : (p a ∧ ∀ (b) (_ : b ≠ a), p b) ↔ ∀ b, p b := by
- simp only [← @forall_eq _ p a, ← forall_and, ← or_imp, Classical.em, forall_const, iff_self]
+ simp only [← @forall_eq _ p a, ← forall_and, ← or_imp, Classical.em, forall_const]
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
@@ -661,7 +661,7 @@ theorem exists_comm {p : α → β → Prop} : (∃ a b, p a b) ↔ ∃ b a, p a
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 ι₁, iff_self]
+ simp only [@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⟩⟩
Fix a lot of wrong casing mostly in the docstrings but also sometimes in def/theorem names. E.g. fin 2 --> Fin 2
, add_monoid_hom --> AddMonoidHom
Remove \n
from to_additive
docstrings that were inserted by mathport.
Move files and directories with Gcd
and Smul
to GCD
and SMul
@@ -196,20 +196,20 @@ alias by_contradiction ← 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.
-The `decidable` namespace contains versions of lemmas from the root namespace that explicitly
+The `Decidable` namespace contains versions of lemmas from the root namespace that explicitly
attempt to avoid the axiom of choice, usually by adding decidability assumptions on the inputs.
You can check if a lemma uses the axiom of choice by using `#print axioms foo` and seeing if
-`classical.choice` appears in the list.
+`Classical.choice` appears in the list.
-/
library_note "decidable arguments"/--
As mathlib is primarily classical,
-if the type signature of a `def` or `lemma` does not require any `decidable` instances to state,
-it is preferable not to introduce any `decidable` instances that are needed in the proof
+if the type signature of a `def` or `lemma` does not require any `Decidable` instances to state,
+it is preferable not to introduce any `Decidable` instances that are needed in the proof
as arguments, but rather to use the `classical` tactic as needed.
-In the other direction, when `decidable` instances do appear in the type signature,
+In the other direction, when `Decidable` instances do appear in the type signature,
it is better to use explicitly introduced ones rather than allowing Lean to automatically infer
classical ones, as these may cause instance mismatch errors later.
-/
@@ -397,8 +397,8 @@ theorem and_iff_not_or_not : a ∧ b ↔ ¬(¬a ∨ ¬b) := Decidable.and_iff_no
simp only [not_and, Xor', not_or, not_not, ← iff_iff_implies_and_implies, iff_self]
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]; rfl
-theorem xor_iff_not_iff' : Xor' a b ↔ (¬a ↔ b) := by simp only [← @xor_not_left _ b, not_not]; rfl
+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]
end Propositional
@@ -846,7 +846,7 @@ theorem bex_congr (H : ∀ x h, P x h ↔ Q x h) : (∃ x h, P x h) ↔ ∃ x h,
exists_congr fun x ↦ exists_congr (H x)
theorem bex_eq_left {a : α} : (∃ (x : _) (_ : x = a), p x) ↔ p a := by
- simp only [exists_prop, exists_eq_left]; rfl
+ simp only [exists_prop, exists_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₁ _ _
@@ -949,7 +949,7 @@ theorem dite_ne_left_iff : dite P (fun _ ↦ a) B ≠ a ↔ ∃ h, a ≠ B h :=
exact exists_congr fun h ↦ by rw [ne_comm]
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]; rfl
+ simp only [Ne.def, dite_eq_right_iff, not_forall]
theorem ite_ne_left_iff : ite P a b ≠ a ↔ ¬P ∧ a ≠ b :=
dite_ne_left_iff.trans <| by simp only; rw [exists_prop]
The script used to do this is included. The yaml file was obtained from https://raw.githubusercontent.com/wiki/leanprover-community/mathlib/mathlib4-port-status.md
@@ -2,6 +2,11 @@
Copyright (c) 2016 Jeremy Avigad. All rights reserved.
Released under Apache 2.0 license as described in the file LICENSE.
Authors: Jeremy Avigad, Leonardo de Moura
+
+! This file was ported from Lean 3 source module logic.basic
+! leanprover-community/mathlib commit 1c521b4fb909320eca16b2bb6f8b5b0490b1cb5e
+! Please do not edit these lines, except to modify the commit id
+! if you have ported upstream changes.
-/
import Mathlib.Init.Logic
import Mathlib.Init.Function
All dependencies are ported!