From 66313702fd153f900ade8ad7279598acb622fb24 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Sun, 6 Nov 2022 20:20:10 +1100 Subject: [PATCH 001/127] port of data.funlike.* --- Mathlib/Data/FunLike/Basic.lean | 209 +++++++++++++++++++++++++++ Mathlib/Data/FunLike/Embedding.lean | 155 ++++++++++++++++++++ Mathlib/Data/FunLike/Equiv.lean | 211 ++++++++++++++++++++++++++++ 3 files changed, 575 insertions(+) create mode 100644 Mathlib/Data/FunLike/Basic.lean create mode 100644 Mathlib/Data/FunLike/Embedding.lean create mode 100644 Mathlib/Data/FunLike/Equiv.lean diff --git a/Mathlib/Data/FunLike/Basic.lean b/Mathlib/Data/FunLike/Basic.lean new file mode 100644 index 0000000000000..5d4117de182fd --- /dev/null +++ b/Mathlib/Data/FunLike/Basic.lean @@ -0,0 +1,209 @@ +/- +Copyright (c) 2021 Anne Baanen. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Anne Baanen +-/ +import Mathlib.Logic.Function.Basic +import Mathlib.Tactic.NormCast + +/-! +# Typeclass for a type `F` with an injective map to `A → B` + +This typeclass is primarily for use by homomorphisms like `monoid_hom` and `linear_map`. + +## Basic usage of `fun_like` + +A typical type of morphisms should be declared as: +``` +structure my_hom (A B : Type*) [my_class A] [my_class B] := +(to_fun : A → B) +(map_op' : ∀ {x y : A}, to_fun (my_class.op x y) = my_class.op (to_fun x) (to_fun y)) + +namespace my_hom + +variables (A B : Type*) [my_class A] [my_class B] + +-- This instance is optional if you follow the "morphism class" design below: +instance : fun_like (my_hom A B) A (λ _, B) := +{ coe := my_hom.to_fun, coe_injective' := λ f g h, by cases f; cases g; congr' } + +/-- Helper instance for when there's too many metavariables to apply +`fun_like.has_coe_to_fun` directly. -/ +instance : has_coe_to_fun (my_hom A B) (λ _, A → B) := fun_like.has_coe_to_fun + +@[simp] lemma to_fun_eq_coe {f : my_hom A B} : f.to_fun = (f : A → B) := rfl + +@[ext] theorem ext {f g : my_hom A B} (h : ∀ x, f x = g x) : f = g := fun_like.ext f g h + +/-- Copy of a `my_hom` with a new `to_fun` equal to the old one. Useful to fix definitional +equalities. -/ +protected def copy (f : my_hom A B) (f' : A → B) (h : f' = ⇑f) : my_hom A B := +{ to_fun := f', + map_op' := h.symm ▸ f.map_op' } + +end my_hom +``` + +This file will then provide a `has_coe_to_fun` instance and various +extensionality and simp lemmas. + +## Morphism classes extending `fun_like` + +The `fun_like` design provides further benefits if you put in a bit more work. +The first step is to extend `fun_like` to create a class of those types satisfying +the axioms of your new type of morphisms. +Continuing the example above: + +``` +section +set_option old_structure_cmd true + +/-- `my_hom_class F A B` states that `F` is a type of `my_class.op`-preserving morphisms. +You should extend this class when you extend `my_hom`. -/ +class my_hom_class (F : Type*) (A B : out_param $ Type*) [my_class A] [my_class B] + extends fun_like F A (λ _, B) := +(map_op : ∀ (f : F) (x y : A), f (my_class.op x y) = my_class.op (f x) (f y)) + +end +@[simp] lemma map_op {F A B : Type*} [my_class A] [my_class B] [my_hom_class F A B] + (f : F) (x y : A) : f (my_class.op x y) = my_class.op (f x) (f y) := +my_hom_class.map_op + +-- You can replace `my_hom.fun_like` with the below instance: +instance : my_hom_class (my_hom A B) A B := +{ coe := my_hom.to_fun, + coe_injective' := λ f g h, by cases f; cases g; congr', + map_op := my_hom.map_op' } + +-- [Insert `has_coe_to_fun`, `to_fun_eq_coe`, `ext` and `copy` here] +``` + +The second step is to add instances of your new `my_hom_class` for all types extending `my_hom`. +Typically, you can just declare a new class analogous to `my_hom_class`: + +``` +structure cooler_hom (A B : Type*) [cool_class A] [cool_class B] + extends my_hom A B := +(map_cool' : to_fun cool_class.cool = cool_class.cool) + +section +set_option old_structure_cmd true + +class cooler_hom_class (F : Type*) (A B : out_param $ Type*) [cool_class A] [cool_class B] + extends my_hom_class F A B := +(map_cool : ∀ (f : F), f cool_class.cool = cool_class.cool) + +end + +@[simp] lemma map_cool {F A B : Type*} [cool_class A] [cool_class B] [cooler_hom_class F A B] + (f : F) : f cool_class.cool = cool_class.cool := +my_hom_class.map_op + +-- You can also replace `my_hom.fun_like` with the below instance: +instance : cool_hom_class (cool_hom A B) A B := +{ coe := cool_hom.to_fun, + coe_injective' := λ f g h, by cases f; cases g; congr', + map_op := cool_hom.map_op', + map_cool := cool_hom.map_cool' } + +-- [Insert `has_coe_to_fun`, `to_fun_eq_coe`, `ext` and `copy` here] +``` + +Then any declaration taking a specific type of morphisms as parameter can instead take the +class you just defined: +``` +-- Compare with: lemma do_something (f : my_hom A B) : sorry := sorry +lemma do_something {F : Type*} [my_hom_class F A B] (f : F) : sorry := sorry +``` + +This means anything set up for `my_hom`s will automatically work for `cool_hom_class`es, +and defining `cool_hom_class` only takes a constant amount of effort, +instead of linearly increasing the work per `my_hom`-related declaration. + +-/ + +-- Porting note: Not sure what to do with this! +-- Porting note: this had priority 10 in mathlib 3. +-- This instance should have low priority, to ensure we follow the chain +-- `FunLike → CoeFun` +-- attribute [instance] coe_fn_trans + +/-- The class `fun_like F α β` expresses that terms of type `F` have an +injective coercion to functions from `α` to `β`. + +This typeclass is used in the definition of the homomorphism typeclasses, +such as `zero_hom_class`, `mul_hom_class`, `monoid_hom_class`, .... +-/ +class FunLike (F : Sort _) (α : outParam (Sort _)) (β : outParam <| α → Sort _) where + coe : F → ∀ a : α, β a + coe_injective' : Function.Injective coe + +section Dependent + +/-! ### `FunLike F α β` where `β` depends on `a : α` -/ + +variable (F α : Sort _) (β : α → Sort _) + +namespace FunLike + +variable {F α β} [i : FunLike F α β] + +-- Give this a priority between `coe_fn_trans` and the default priority +-- `α` and `β` are out_params, so this instance should not be dangerous +-- Porting note: @[nolint dangerous_instance] +instance (priority := 100) : CoeFun F fun _ => ∀ a : α, β a where coe := FunLike.coe + +-- Porting note: the next two lemmas presumably aren't needed. +-- @[simp] +-- theorem coe_eq_coe_fn : (FunLike.coe : F → ∀ a : α, β a) = coeFn := +-- rfl + +-- theorem coe_injective : Function.Injective (coeFn : F → ∀ a : α, β a) := +-- FunLike.coe_injective' + +@[simp] +theorem coe_fn_eq {f g : F} : (f : ∀ a : α, β a) = (g : ∀ a : α, β a) ↔ f = g := + ⟨fun h => FunLike.coe_injective' h, fun h => by cases h <;> rfl⟩ + +theorem ext' {f g : F} (h : (f : ∀ a : α, β a) = (g : ∀ a : α, β a)) : f = g := + FunLike.coe_injective' h + +theorem ext'_iff {f g : F} : f = g ↔ (f : ∀ a : α, β a) = (g : ∀ a : α, β a) := + coe_fn_eq.symm + +theorem ext (f g : F) (h : ∀ x : α, f x = g x) : f = g := + FunLike.coe_injective' (funext h) + +theorem ext_iff {f g : F} : f = g ↔ ∀ x, f x = g x := + coe_fn_eq.symm.trans Function.funext_iff + +protected theorem congr_fun {f g : F} (h₁ : f = g) (x : α) : f x = g x := + congr_fun (congr_arg _ h₁) x + +theorem ne_iff {f g : F} : f ≠ g ↔ ∃ a, f a ≠ g a := + ext_iff.not.trans not_forall + +theorem exists_ne {f g : F} (h : f ≠ g) : ∃ x, f x ≠ g x := + ne_iff.mp h + +end FunLike + +end Dependent + +section NonDependent + +/-! ### `FunLike F α (λ _, β)` where `β` does not depend on `a : α` -/ + +variable {F α β : Sort _} [i : FunLike F α fun _ => β] + +namespace FunLike + +protected theorem congr {f g : F} {x y : α} (h₁ : f = g) (h₂ : x = y) : f x = g y := + congr (congr_arg _ h₁) h₂ + +protected theorem congr_arg (f : F) {x y : α} (h₂ : x = y) : f x = f y := + congr_arg _ h₂ + +end FunLike + +end NonDependent diff --git a/Mathlib/Data/FunLike/Embedding.lean b/Mathlib/Data/FunLike/Embedding.lean new file mode 100644 index 0000000000000..74d09e37b90fc --- /dev/null +++ b/Mathlib/Data/FunLike/Embedding.lean @@ -0,0 +1,155 @@ +/- +Copyright (c) 2021 Anne Baanen. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Anne Baanen +-/ +import Mathlib.Data.FunLike.Basic + +/-! +# Typeclass for a type `F` with an injective map to `A ↪ B` + +This typeclass is primarily for use by embeddings such as `rel_embedding`. + +## Basic usage of `embedding_like` + +A typical type of embedding should be declared as: +``` +structure my_embedding (A B : Type*) [my_class A] [my_class B] := +(to_fun : A → B) +(injective' : function.injective to_fun) +(map_op' : ∀ {x y : A}, to_fun (my_class.op x y) = my_class.op (to_fun x) (to_fun y)) + +namespace my_embedding + +variables (A B : Type*) [my_class A] [my_class B] + +-- This instance is optional if you follow the "Embedding class" design below: +instance : embedding_like (my_embedding A B) A B := +{ coe := my_embedding.to_fun, + coe_injective' := λ f g h, by cases f; cases g; congr', + injective' := my_embedding.injective' } + +/-- Helper instance for when there's too many metavariables to directly +apply `fun_like.to_coe_fn`. -/ +instance : has_coe_to_fun (my_embedding A B) (λ _, A → B) := ⟨my_embedding.to_fun⟩ + +@[simp] lemma to_fun_eq_coe {f : my_embedding A B} : f.to_fun = (f : A → B) := rfl + +@[ext] theorem ext {f g : my_embedding A B} (h : ∀ x, f x = g x) : f = g := fun_like.ext f g h + +/-- Copy of a `my_embedding` with a new `to_fun` equal to the old one. Useful to fix definitional +equalities. -/ +protected def copy (f : my_embedding A B) (f' : A → B) (h : f' = ⇑f) : my_embedding A B := +{ to_fun := f', + injective' := h.symm ▸ f.injective', + map_op' := h.symm ▸ f.map_op' } + +end my_embedding +``` + +This file will then provide a `has_coe_to_fun` instance and various +extensionality and simp lemmas. + +## Embedding classes extending `embedding_like` + +The `embedding_like` design provides further benefits if you put in a bit more work. +The first step is to extend `embedding_like` to create a class of those types satisfying +the axioms of your new type of morphisms. +Continuing the example above: + +``` +section +set_option old_structure_cmd true + +/-- `my_embedding_class F A B` states that `F` is a type of `my_class.op`-preserving embeddings. +You should extend this class when you extend `my_embedding`. -/ +class my_embedding_class (F : Type*) (A B : out_param $ Type*) [my_class A] [my_class B] + extends embedding_like F A B := +(map_op : ∀ (f : F) (x y : A), f (my_class.op x y) = my_class.op (f x) (f y)) + +end + +@[simp] lemma map_op {F A B : Type*} [my_class A] [my_class B] [my_embedding_class F A B] + (f : F) (x y : A) : f (my_class.op x y) = my_class.op (f x) (f y) := +my_embedding_class.map_op + +-- You can replace `my_embedding.embedding_like` with the below instance: +instance : my_embedding_class (my_embedding A B) A B := +{ coe := my_embedding.to_fun, + coe_injective' := λ f g h, by cases f; cases g; congr', + injective' := my_embedding.injective', + map_op := my_embedding.map_op' } + +-- [Insert `has_coe_to_fun`, `to_fun_eq_coe`, `ext` and `copy` here] +``` + +The second step is to add instances of your new `my_embedding_class` for all types extending +`my_embedding`. +Typically, you can just declare a new class analogous to `my_embedding_class`: + +``` +structure cooler_embedding (A B : Type*) [cool_class A] [cool_class B] + extends my_embedding A B := +(map_cool' : to_fun cool_class.cool = cool_class.cool) + +section +set_option old_structure_cmd true + +class cooler_embedding_class (F : Type*) (A B : out_param $ Type*) [cool_class A] [cool_class B] + extends my_embedding_class F A B := +(map_cool : ∀ (f : F), f cool_class.cool = cool_class.cool) + +end + +@[simp] lemma map_cool {F A B : Type*} [cool_class A] [cool_class B] [cooler_embedding_class F A B] + (f : F) : f cool_class.cool = cool_class.cool := +my_embedding_class.map_op + +-- You can also replace `my_embedding.embedding_like` with the below instance: +instance : cool_embedding_class (cool_embedding A B) A B := +{ coe := cool_embedding.to_fun, + coe_injective' := λ f g h, by cases f; cases g; congr', + injective' := my_embedding.injective', + map_op := cool_embedding.map_op', + map_cool := cool_embedding.map_cool' } + +-- [Insert `has_coe_to_fun`, `to_fun_eq_coe`, `ext` and `copy` here] +``` + +Then any declaration taking a specific type of morphisms as parameter can instead take the +class you just defined: +``` +-- Compare with: lemma do_something (f : my_embedding A B) : sorry := sorry +lemma do_something {F : Type*} [my_embedding_class F A B] (f : F) : sorry := sorry +``` + +This means anything set up for `my_embedding`s will automatically work for `cool_embedding_class`es, +and defining `cool_embedding_class` only takes a constant amount of effort, +instead of linearly increasing the work per `my_embedding`-related declaration. + +-/ + + +/-- The class `EmbeddingLike F α β` expresses that terms of type `F` have an +injective coercion to injective functions `α ↪ β`. +-/ +class EmbeddingLike (F : Sort _) (α β : outParam (Sort _)) extends FunLike F α fun _ => β where + injective' : ∀ f : F, @Function.Injective α β (coe f) + +namespace EmbeddingLike + +variable {F α β γ : Sort _} [i : EmbeddingLike F α β] + +protected theorem injective (f : F) : Function.Injective f := + injective' f + +@[simp] +theorem apply_eq_iff_eq (f : F) {x y : α} : f x = f y ↔ x = y := + (EmbeddingLike.injective f).eq_iff + +@[simp] +theorem comp_injective {F : Sort _} [EmbeddingLike F β γ] (f : α → β) (e : F) : + Function.Injective (e ∘ f) ↔ Function.Injective f := + (EmbeddingLike.injective e).of_comp_iff f + +end EmbeddingLike diff --git a/Mathlib/Data/FunLike/Equiv.lean b/Mathlib/Data/FunLike/Equiv.lean new file mode 100644 index 0000000000000..28a232f55d02e --- /dev/null +++ b/Mathlib/Data/FunLike/Equiv.lean @@ -0,0 +1,211 @@ +/- +Copyright (c) 2021 Anne Baanen. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Anne Baanen +-/ +import Mathlib.Data.FunLike.Embedding + +/-! +# Typeclass for a type `F` with an injective map to `A ≃ B` + +This typeclass is primarily for use by isomorphisms like `monoid_equiv` and `linear_equiv`. + +## Basic usage of `equiv_like` + +A typical type of morphisms should be declared as: +``` +structure my_iso (A B : Type*) [my_class A] [my_class B] + extends equiv A B := +(map_op' : ∀ {x y : A}, to_fun (my_class.op x y) = my_class.op (to_fun x) (to_fun y)) + +namespace my_iso + +variables (A B : Type*) [my_class A] [my_class B] + +-- This instance is optional if you follow the "Isomorphism class" design below: +instance : equiv_like (my_iso A B) A (λ _, B) := +{ coe := my_iso.to_equiv.to_fun, + inv := my_iso.to_equiv.inv_fun, + left_inv := my_iso.to_equiv.left_inv, + right_inv := my_iso.to_equiv.right_inv, + coe_injective' := λ f g h, by cases f; cases g; congr' } + +/-- Helper instance for when there's too many metavariables to apply `equiv_like.coe` directly. -/ +instance : has_coe_to_fun (my_iso A B) := to_fun.to_coe_fn + +@[simp] lemma to_fun_eq_coe {f : my_iso A B} : f.to_fun = (f : A → B) := rfl + +@[ext] theorem ext {f g : my_iso A B} (h : ∀ x, f x = g x) : f = g := fun_like.ext f g h + +/-- Copy of a `my_iso` with a new `to_fun` equal to the old one. Useful to fix definitional +equalities. -/ +protected def copy (f : my_iso A B) (f' : A → B) (f_inv : B → A) (h : f' = ⇑f) : my_iso A B := +{ to_fun := f', + inv_fun := f_inv, + left_inv := h.symm ▸ f.left_inv, + right_inv := h.symm ▸ f.right_inv, + map_op' := h.symm ▸ f.map_op' } + +end my_iso +``` + +This file will then provide a `has_coe_to_fun` instance and various +extensionality and simp lemmas. + +## Isomorphism classes extending `equiv_like` + +The `equiv_like` design provides further benefits if you put in a bit more work. +The first step is to extend `equiv_like` to create a class of those types satisfying +the axioms of your new type of isomorphisms. +Continuing the example above: + +``` +section +set_option old_structure_cmd true + +/-- `my_iso_class F A B` states that `F` is a type of `my_class.op`-preserving morphisms. +You should extend this class when you extend `my_iso`. -/ +class my_iso_class (F : Type*) (A B : out_param $ Type*) [my_class A] [my_class B] + extends equiv_like F A (λ _, B), my_hom_class F A B. + +end + +-- You can replace `my_iso.equiv_like` with the below instance: +instance : my_iso_class (my_iso A B) A B := +{ coe := my_iso.to_fun, + inv := my_iso.inv_fun, + left_inv := my_iso.left_inv, + right_inv := my_iso.right_inv, + coe_injective' := λ f g h, by cases f; cases g; congr', + map_op := my_iso.map_op' } + +-- [Insert `has_coe_to_fun`, `to_fun_eq_coe`, `ext` and `copy` here] +``` + +The second step is to add instances of your new `my_iso_class` for all types extending `my_iso`. +Typically, you can just declare a new class analogous to `my_iso_class`: + +``` +structure cooler_iso (A B : Type*) [cool_class A] [cool_class B] + extends my_iso A B := +(map_cool' : to_fun cool_class.cool = cool_class.cool) + +section +set_option old_structure_cmd true + +class cooler_iso_class (F : Type*) (A B : out_param $ Type*) [cool_class A] [cool_class B] + extends my_iso_class F A B := +(map_cool : ∀ (f : F), f cool_class.cool = cool_class.cool) + +end + +@[simp] lemma map_cool {F A B : Type*} [cool_class A] [cool_class B] [cooler_iso_class F A B] + (f : F) : f cool_class.cool = cool_class.cool := +my_iso_class.map_op + +-- You can also replace `my_iso.equiv_like` with the below instance: +instance : cool_iso_class (cool_iso A B) A B := +{ coe := cool_iso.to_fun, + coe_injective' := λ f g h, by cases f; cases g; congr', + map_op := cool_iso.map_op', + map_cool := cool_iso.map_cool' } + +-- [Insert `has_coe_to_fun`, `to_fun_eq_coe`, `ext` and `copy` here] +``` + +Then any declaration taking a specific type of morphisms as parameter can instead take the +class you just defined: +``` +-- Compare with: lemma do_something (f : my_iso A B) : sorry := sorry +lemma do_something {F : Type*} [my_iso_class F A B] (f : F) : sorry := sorry +``` + +This means anything set up for `my_iso`s will automatically work for `cool_iso_class`es, +and defining `cool_iso_class` only takes a constant amount of effort, +instead of linearly increasing the work per `my_iso`-related declaration. + +-/ + + +/-- The class `equiv_like E α β` expresses that terms of type `E` have an +injective coercion to bijections between `α` and `β`. + +This typeclass is used in the definition of the homomorphism typeclasses, +such as `zero_equiv_class`, `mul_equiv_class`, `monoid_equiv_class`, .... +-/ +class EquivLike (E : Sort _) (α β : outParam (Sort _)) where + coe : E → α → β + inv : E → β → α + left_inv : ∀ e, Function.LeftInverse (inv e) (coe e) + right_inv : ∀ e, Function.RightInverse (inv e) (coe e) + -- The `inv` hypothesis makes this easier to prove with `congr'` + coe_injective' : ∀ e g, coe e = coe g → inv e = inv g → e = g + +namespace EquivLike + +variable {E F α β γ : Sort _} [iE : EquivLike E α β] [iF : EquivLike F β γ] + +theorem inv_injective : Function.Injective (EquivLike.inv : E → β → α) := fun e g h => + coe_injective' e g ((right_inv e).eq_right_inverse (h.symm ▸ left_inv g)) h + +instance (priority := 100) toEmbeddingLike : EmbeddingLike E α β where + coe := (coe : E → α → β) + coe_injective' e g h := coe_injective' e g h ((left_inv e).eq_right_inverse (h.symm ▸ right_inv g)) + injective' e := (left_inv e).injective + +protected theorem injective (e : E) : Function.Injective e := + EmbeddingLike.injective e + +protected theorem surjective (e : E) : Function.Surjective e := + (right_inv e).surjective + +protected theorem bijective (e : E) : Function.Bijective (e : α → β) := + ⟨EquivLike.injective e, EquivLike.surjective e⟩ + +theorem apply_eq_iff_eq (f : E) {x y : α} : f x = f y ↔ x = y := + EmbeddingLike.apply_eq_iff_eq f + +@[simp] +theorem injective_comp (e : E) (f : β → γ) : Function.Injective (f ∘ e) ↔ Function.Injective f := + Function.Injective.of_comp_iff' f (EquivLike.bijective e) + +@[simp] +theorem surjective_comp (e : E) (f : β → γ) : Function.Surjective (f ∘ e) ↔ Function.Surjective f := + (EquivLike.surjective e).of_comp_iff f + +@[simp] +theorem bijective_comp (e : E) (f : β → γ) : Function.Bijective (f ∘ e) ↔ Function.Bijective f := + (EquivLike.bijective e).of_comp_iff f + +/-- This lemma is only supposed to be used in the generic context, when working with instances +of classes extending `equiv_like`. +For concrete isomorphism types such as `equiv`, you should use `equiv.symm_apply_apply` +or its equivalent. + +TODO: define a generic form of `equiv.symm`. -/ +@[simp] +theorem inv_apply_apply (e : E) (a : α) : EquivLike.inv e (e a) = a := + left_inv _ _ + +/-- This lemma is only supposed to be used in the generic context, when working with instances +of classes extending `equiv_like`. +For concrete isomorphism types such as `equiv`, you should use `equiv.apply_symm_apply` +or its equivalent. + +TODO: define a generic form of `equiv.symm`. -/ +@[simp] +theorem apply_inv_apply (e : E) (b : β) : e (EquivLike.inv e b) = b := + right_inv _ _ + +theorem comp_injective (f : α → β) (e : F) : Function.Injective (e ∘ f) ↔ Function.Injective f := + EmbeddingLike.comp_injective f e + +@[simp] +theorem comp_surjective (f : α → β) (e : F) : Function.Surjective (e ∘ f) ↔ Function.Surjective f := + Function.Surjective.of_comp_iff' (EquivLike.bijective e) f + +@[simp] +theorem comp_bijective (f : α → β) (e : F) : Function.Bijective (e ∘ f) ↔ Function.Bijective f := + (EquivLike.bijective e).of_comp_iff' f + +end EquivLike From 7700ecf3cacf15f79b221202efc8a82647fb9e13 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Mon, 7 Nov 2022 15:10:26 +1100 Subject: [PATCH 002/127] long line --- Mathlib/Data/FunLike/Equiv.lean | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/Mathlib/Data/FunLike/Equiv.lean b/Mathlib/Data/FunLike/Equiv.lean index 28a232f55d02e..6c4fb549e0618 100644 --- a/Mathlib/Data/FunLike/Equiv.lean +++ b/Mathlib/Data/FunLike/Equiv.lean @@ -150,7 +150,8 @@ theorem inv_injective : Function.Injective (EquivLike.inv : E → β → α) := instance (priority := 100) toEmbeddingLike : EmbeddingLike E α β where coe := (coe : E → α → β) - coe_injective' e g h := coe_injective' e g h ((left_inv e).eq_right_inverse (h.symm ▸ right_inv g)) + coe_injective' e g h := + coe_injective' e g h ((left_inv e).eq_right_inverse (h.symm ▸ right_inv g)) injective' e := (left_inv e).injective protected theorem injective (e : E) : Function.Injective e := From df8ac7fb77df37a2e354a6ff3f5d9676ee2f6f99 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Mon, 7 Nov 2022 15:15:24 +1100 Subject: [PATCH 003/127] docBlame --- Mathlib/Data/FunLike/Basic.lean | 2 ++ Mathlib/Data/FunLike/Embedding.lean | 1 + Mathlib/Data/FunLike/Equiv.lean | 8 +++++++- 3 files changed, 10 insertions(+), 1 deletion(-) diff --git a/Mathlib/Data/FunLike/Basic.lean b/Mathlib/Data/FunLike/Basic.lean index 5d4117de182fd..a894373286718 100644 --- a/Mathlib/Data/FunLike/Basic.lean +++ b/Mathlib/Data/FunLike/Basic.lean @@ -135,7 +135,9 @@ This typeclass is used in the definition of the homomorphism typeclasses, such as `zero_hom_class`, `mul_hom_class`, `monoid_hom_class`, .... -/ class FunLike (F : Sort _) (α : outParam (Sort _)) (β : outParam <| α → Sort _) where + /-- The coercion from `F` to a function. -/ coe : F → ∀ a : α, β a + /-- The coercion to functions must be injective. -/ coe_injective' : Function.Injective coe section Dependent diff --git a/Mathlib/Data/FunLike/Embedding.lean b/Mathlib/Data/FunLike/Embedding.lean index 74d09e37b90fc..0c95299cda48b 100644 --- a/Mathlib/Data/FunLike/Embedding.lean +++ b/Mathlib/Data/FunLike/Embedding.lean @@ -134,6 +134,7 @@ instead of linearly increasing the work per `my_embedding`-related declaration. injective coercion to injective functions `α ↪ β`. -/ class EmbeddingLike (F : Sort _) (α β : outParam (Sort _)) extends FunLike F α fun _ => β where + /-- The coercion to functions must produce injective functions. -/ injective' : ∀ f : F, @Function.Injective α β (coe f) namespace EmbeddingLike diff --git a/Mathlib/Data/FunLike/Equiv.lean b/Mathlib/Data/FunLike/Equiv.lean index 6c4fb549e0618..11f2f5482f19f 100644 --- a/Mathlib/Data/FunLike/Equiv.lean +++ b/Mathlib/Data/FunLike/Equiv.lean @@ -134,12 +134,18 @@ This typeclass is used in the definition of the homomorphism typeclasses, such as `zero_equiv_class`, `mul_equiv_class`, `monoid_equiv_class`, .... -/ class EquivLike (E : Sort _) (α β : outParam (Sort _)) where + /-- The coercion to a function in the forward direction. -/ coe : E → α → β + /-- The coercion to a function in the backwards direction. -/ inv : E → β → α + /-- The coercions are left inverses. -/ left_inv : ∀ e, Function.LeftInverse (inv e) (coe e) + /-- The coercions are right inverses. -/ right_inv : ∀ e, Function.RightInverse (inv e) (coe e) - -- The `inv` hypothesis makes this easier to prove with `congr'` + /-- If two coercions to functions are jointly injective. -/ coe_injective' : ∀ e g, coe e = coe g → inv e = inv g → e = g + -- This is mathematically equivalent to either of the coercions to functions being injective, but + -- the `inv` hypothesis makes this easier to prove with `congr'` namespace EquivLike From 4f074d329d2b7ee4fc14ae9576247388142e80f1 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Mon, 7 Nov 2022 15:27:15 +1100 Subject: [PATCH 004/127] import --- Mathlib.lean | 3 +++ 1 file changed, 3 insertions(+) diff --git a/Mathlib.lean b/Mathlib.lean index c62f9f3b59789..f538295361411 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -29,6 +29,9 @@ import Mathlib.Data.Fin.Basic import Mathlib.Data.Fin.Fin2 import Mathlib.Data.Finset.Basic import Mathlib.Data.Fintype.Basic +import Mathlib.Data.FunLike.Basic +import Mathlib.Data.FunLike.Embedding +import Mathlib.Data.FunLike.Equiv import Mathlib.Data.Int.Basic import Mathlib.Data.Int.Cast import Mathlib.Data.Int.Cast.Defs From 6170aab669cc94747631c51199d13ca2f9d6b018 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Mon, 7 Nov 2022 15:48:22 +1100 Subject: [PATCH 005/127] copy and paste mathport output --- Mathlib/Data/Equiv/Defs.lean | 0 1 file changed, 0 insertions(+), 0 deletions(-) create mode 100644 Mathlib/Data/Equiv/Defs.lean diff --git a/Mathlib/Data/Equiv/Defs.lean b/Mathlib/Data/Equiv/Defs.lean new file mode 100644 index 0000000000000..e69de29bb2d1d From 76bc3e2a6242bc018d5e26b2cf1be63b404cadb6 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Mon, 7 Nov 2022 15:55:12 +1100 Subject: [PATCH 006/127] restore lemmas --- Mathlib/Data/FunLike/Basic.lean | 11 +++++------ 1 file changed, 5 insertions(+), 6 deletions(-) diff --git a/Mathlib/Data/FunLike/Basic.lean b/Mathlib/Data/FunLike/Basic.lean index a894373286718..ff43170b66024 100644 --- a/Mathlib/Data/FunLike/Basic.lean +++ b/Mathlib/Data/FunLike/Basic.lean @@ -155,13 +155,12 @@ variable {F α β} [i : FunLike F α β] -- Porting note: @[nolint dangerous_instance] instance (priority := 100) : CoeFun F fun _ => ∀ a : α, β a where coe := FunLike.coe --- Porting note: the next two lemmas presumably aren't needed. --- @[simp] --- theorem coe_eq_coe_fn : (FunLike.coe : F → ∀ a : α, β a) = coeFn := --- rfl +@[simp] +theorem coe_eq_coe_fn : (FunLike.coe : F → ∀ a : α, β a) = (fun f => ↑f) := + rfl --- theorem coe_injective : Function.Injective (coeFn : F → ∀ a : α, β a) := --- FunLike.coe_injective' +theorem coe_injective : Function.Injective (fun f : F => (f : ∀ a : α, β a)) := + FunLike.coe_injective' @[simp] theorem coe_fn_eq {f g : F} : (f : ∀ a : α, β a) = (g : ∀ a : α, β a) ↔ f = g := From f240c63219a2cb82883fd09ac9abbab23fdf3a91 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Mon, 7 Nov 2022 16:05:25 +1100 Subject: [PATCH 007/127] starting --- Mathlib/Data/Equiv/Defs.lean | 1148 ++++++++++++++++++++++++++++++++++ 1 file changed, 1148 insertions(+) diff --git a/Mathlib/Data/Equiv/Defs.lean b/Mathlib/Data/Equiv/Defs.lean index e69de29bb2d1d..89df0a9b9a52c 100644 --- a/Mathlib/Data/Equiv/Defs.lean +++ b/Mathlib/Data/Equiv/Defs.lean @@ -0,0 +1,1148 @@ +/- +Copyright (c) 2015 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Leonardo de Moura, Mario Carneiro +-/ +import Mathlib.Logic.Unique +import Mathlib.Data.FunLike.Equiv +import Mathlib.Tactic.Relation.Rfl +import Mathlib.Tactic.Relation.Symm +import Mathlib.Tactic.Relation.Trans +import Mathlib.Tactic.Simps.Basic + +/-! +# Equivalence between types + +In this file we define two types: + +* `equiv α β` a.k.a. `α ≃ β`: a bijective map `α → β` bundled with its inverse map; we use this (and + not equality!) to express that various `Type`s or `Sort`s are equivalent. + +* `equiv.perm α`: the group of permutations `α ≃ α`. More lemmas about `equiv.perm` can be found in + `group_theory/perm`. + +Then we define + +* canonical isomorphisms between various types: e.g., + + - `equiv.refl α` is the identity map interpreted as `α ≃ α`; + +* operations on equivalences: e.g., + + - `equiv.symm e : β ≃ α` is the inverse of `e : α ≃ β`; + + - `equiv.trans e₁ e₂ : α ≃ γ` is the composition of `e₁ : α ≃ β` and `e₂ : β ≃ γ` (note the order + of the arguments!); + +* definitions that transfer some instances along an equivalence. By convention, we transfer + instances from right to left. + + - `equiv.inhabited` takes `e : α ≃ β` and `[inhabited β]` and returns `inhabited α`; + - `equiv.unique` takes `e : α ≃ β` and `[unique β]` and returns `unique α`; + - `equiv.decidable_eq` takes `e : α ≃ β` and `[decidable_eq β]` and returns `decidable_eq α`. + + More definitions of this kind can be found in other files. E.g., `data/equiv/transfer_instance` + does it for many algebraic type classes like `group`, `module`, etc. + +Many more such isomorphisms and operations are defined in `logic/equiv/basic`. + +## Tags + +equivalence, congruence, bijective map +-/ + + +open Function + +universe u v w z + +variable {α : Sort u} {β : Sort v} {γ : Sort w} + +/-- `α ≃ β` is the type of functions from `α → β` with a two-sided inverse. -/ +structure Equiv (α : Sort _) (β : Sort _) where + toFun : α → β + invFun : β → α + left_inv : LeftInverse invFun toFun + right_inv : RightInverse invFun toFun + +infixl:25 " ≃ " => Equiv + +instance {F} [EquivLike F α β] : CoeTC F (α ≃ β) := + ⟨fun f => + { toFun := f, + invFun := EquivLike.inv f, + left_inv := EquivLike.left_inv f, + right_inv := EquivLike.right_inv f }⟩ + +/-- `perm α` is the type of bijections from `α` to itself. -/ +@[reducible] +def Equiv.Perm (α : Sort _) := + Equiv α α + +namespace Equiv + +instance : EquivLike (α ≃ β) α β where + coe := toFun + inv := invFun + left_inv := left_inv + right_inv := right_inv + coe_injective' e₁ e₂ h₁ h₂ := by + cases e₁ + cases e₂ + congr + +instance : CoeFun (α ≃ β) fun _ => α → β := + ⟨toFun⟩ + +/-- The map `(r ≃ s) → (r → s)` is injective. -/ +theorem coe_fn_injective : @Function.Injective (α ≃ β) (α → β) (fun e => e) := + FunLike.coe_injective' + +protected theorem coe_inj {e₁ e₂ : α ≃ β} : (e₁ : α → β) = e₂ ↔ e₁ = e₂ := + @FunLike.coe_fn_eq _ _ _ _ e₁ e₂ + +@[ext] +theorem ext {f g : Equiv α β} (H : ∀ x, f x = g x) : f = g := + FunLike.ext f g H + +protected theorem congr_arg {f : Equiv α β} {x x' : α} : x = x' → f x = f x' := + FunLike.congr_arg f + +protected theorem congr_fun {f g : Equiv α β} (h : f = g) (x : α) : f x = g x := + FunLike.congr_fun h x + +theorem ext_iff {f g : Equiv α β} : f = g ↔ ∀ x, f x = g x := + FunLike.ext_iff + +@[ext] +theorem Perm.ext {σ τ : Equiv.Perm α} (H : ∀ x, σ x = τ x) : σ = τ := + Equiv.ext H + +protected theorem Perm.congr_arg {f : Equiv.Perm α} {x x' : α} : x = x' → f x = f x' := + Equiv.congr_arg + +protected theorem Perm.congr_fun {f g : Equiv.Perm α} (h : f = g) (x : α) : f x = g x := + Equiv.congr_fun h x + +theorem Perm.ext_iff {σ τ : Equiv.Perm α} : σ = τ ↔ ∀ x, σ x = τ x := + Equiv.ext_iff + +/-- Any type is equivalent to itself. -/ +@[refl] +protected def refl (α : Sort _) : α ≃ α := + ⟨id, id, fun x => rfl, fun x => rfl⟩ + +instance inhabited' : Inhabited (α ≃ α) := + ⟨Equiv.refl α⟩ + +/-- Inverse of an equivalence `e : α ≃ β`. -/ +-- Porting note: `symm` attribute rejects this lemma because of implicit arguments. +-- @[symm] +protected def symm (e : α ≃ β) : β ≃ α := + ⟨e.invFun, e.toFun, e.right_inv, e.left_inv⟩ + +/-- See Note [custom simps projection] -/ +def Simps.symmApply (e : α ≃ β) : β → α := + e.symm + +initialize_simps_projections Equiv (toFun → apply, invFun → symmApply) + +/-- Composition of equivalences `e₁ : α ≃ β` and `e₂ : β ≃ γ`. -/ +-- Porting note: `trans` attribute rejects this lemma because of implicit arguments. +-- @[trans] +protected def trans (e₁ : α ≃ β) (e₂ : β ≃ γ) : α ≃ γ := + ⟨e₂ ∘ e₁, e₁.symm ∘ e₂.symm, e₂.left_inv.comp e₁.left_inv, e₂.right_inv.comp e₁.right_inv⟩ + +@[simp] +theorem to_fun_as_coe (e : α ≃ β) : e.toFun = e := + rfl + +@[simp] +theorem inv_fun_as_coe (e : α ≃ β) : e.invFun = e.symm := + rfl + +protected theorem injective (e : α ≃ β) : Injective e := + EquivLike.injective e + +protected theorem surjective (e : α ≃ β) : Surjective e := + EquivLike.surjective e + +protected theorem bijective (e : α ≃ β) : Bijective e := + EquivLike.bijective e + +protected theorem subsingleton (e : α ≃ β) [Subsingleton β] : Subsingleton α := + e.Injective.Subsingleton + +protected theorem subsingleton.symm (e : α ≃ β) [Subsingleton α] : Subsingleton β := + e.symm.Injective.Subsingleton + +theorem subsingleton_congr (e : α ≃ β) : Subsingleton α ↔ Subsingleton β := + ⟨fun h => e.symm.subsingleton, fun h => e.subsingleton⟩ + +instance equiv_subsingleton_cod [Subsingleton β] : Subsingleton (α ≃ β) := + ⟨fun f g => Equiv.ext fun x => Subsingleton.elim _ _⟩ + +instance equiv_subsingleton_dom [Subsingleton α] : Subsingleton (α ≃ β) := + ⟨fun f g => Equiv.ext fun x => @Subsingleton.elim _ (Equiv.subsingleton.symm f) _ _⟩ + +instance permUnique [Subsingleton α] : Unique (Perm α) := + uniqueOfSubsingleton (Equiv.refl α) + +theorem Perm.subsingleton_eq_refl [Subsingleton α] (e : Perm α) : e = Equiv.refl α := + Subsingleton.elim _ _ + +/-- Transfer `decidable_eq` across an equivalence. -/ +protected def decidableEq (e : α ≃ β) [DecidableEq β] : DecidableEq α := + e.Injective.DecidableEq + +theorem nonempty_congr (e : α ≃ β) : Nonempty α ↔ Nonempty β := + Nonempty.congr e e.symm + +protected theorem nonempty (e : α ≃ β) [Nonempty β] : Nonempty α := + e.nonempty_congr.mpr ‹_› + +/-- If `α ≃ β` and `β` is inhabited, then so is `α`. -/ +protected def inhabited [Inhabited β] (e : α ≃ β) : Inhabited α := + ⟨e.symm default⟩ + +/-- If `α ≃ β` and `β` is a singleton type, then so is `α`. -/ +protected def unique [Unique β] (e : α ≃ β) : Unique α := + e.symm.Surjective.unique + +/-- Equivalence between equal types. -/ +protected def cast {α β : Sort _} (h : α = β) : α ≃ β := + ⟨cast h, cast h.symm, fun x => by + cases h + rfl, fun x => by + cases h + rfl⟩ + +@[simp] +theorem coe_fn_symm_mk (f : α → β) (g l r) : ((Equiv.mk f g l r).symm : β → α) = g := + rfl + +@[simp] +theorem coe_refl : ⇑(Equiv.refl α) = id := + rfl + +/-- This cannot be a `simp` lemmas as it incorrectly matches against `e : α ≃ synonym α`, when +`synonym α` is semireducible. This makes a mess of `multiplicative.of_add` etc. -/ +theorem Perm.coe_subsingleton {α : Type _} [Subsingleton α] (e : Perm α) : ⇑e = id := by + rw [perm.subsingleton_eq_refl e, coe_refl] + +theorem refl_apply (x : α) : Equiv.refl α x = x := + rfl + +@[simp] +theorem coe_trans (f : α ≃ β) (g : β ≃ γ) : ⇑(f.trans g) = g ∘ f := + rfl + +theorem trans_apply (f : α ≃ β) (g : β ≃ γ) (a : α) : (f.trans g) a = g (f a) := + rfl + +@[simp] +theorem apply_symm_apply (e : α ≃ β) (x : β) : e (e.symm x) = x := + e.right_inv x + +@[simp] +theorem symm_apply_apply (e : α ≃ β) (x : α) : e.symm (e x) = x := + e.left_inv x + +@[simp] +theorem symm_comp_self (e : α ≃ β) : e.symm ∘ e = id := + funext e.symm_apply_apply + +@[simp] +theorem self_comp_symm (e : α ≃ β) : e ∘ e.symm = id := + funext e.apply_symm_apply + +@[simp] +theorem symm_trans_apply (f : α ≃ β) (g : β ≃ γ) (a : γ) : (f.trans g).symm a = f.symm (g.symm a) := + rfl + +-- The `simp` attribute is needed to make this a `dsimp` lemma. +-- `simp` will always rewrite with `equiv.symm_symm` before this has a chance to fire. +@[simp, nolint simp_nf] +theorem symm_symm_apply (f : α ≃ β) (b : α) : f.symm.symm b = f b := + rfl + +theorem apply_eq_iff_eq (f : α ≃ β) {x y : α} : f x = f y ↔ x = y := + EquivLike.apply_eq_iff_eq f + +/- failed to parenthesize: parenthesize: uncaught backtrack exception +[PrettyPrinter.parenthesize.input] (Command.declaration + (Command.declModifiers [] [] [] [] [] []) + (Command.theorem + "theorem" + (Command.declId `apply_eq_iff_eq_symm_apply []) + (Command.declSig + [(Term.implicitBinder "{" [`α `β] [":" (Term.sort "Sort" [(Level.hole "_")])] "}") + (Term.explicitBinder "(" [`f] [":" (Logic.Equiv.Defs.«term_≃_» `α " ≃ " `β)] [] ")") + (Term.implicitBinder "{" [`x] [":" `α] "}") + (Term.implicitBinder "{" [`y] [":" `β] "}")] + (Term.typeSpec + ":" + («term_↔_» + («term_=_» (Term.app `f [`x]) "=" `y) + "↔" + («term_=_» `x "=" (Term.app (Term.proj `f "." `symm) [`y]))))) + (Command.declValSimple + ":=" + (Term.byTactic + "by" + (Tactic.tacticSeq + (Tactic.tacticSeq1Indented + [(Mathlib.Tactic.Conv.convLHS + "conv_lhs" + [] + [] + "=>" + (Tactic.Conv.convSeq + (Tactic.Conv.convSeq1Indented + [(Tactic.Conv.convRw__ + "rw" + [] + (Tactic.rwRuleSeq "[" [(Tactic.rwRule ["←"] (Term.app `apply_symm_apply [`f `y]))] "]"))]))) + [] + (Tactic.rwSeq "rw" [] (Tactic.rwRuleSeq "[" [(Tactic.rwRule [] `apply_eq_iff_eq)] "]") [])]))) + []) + [] + [])) +[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Command.theorem', expected 'Lean.Parser.Command.abbrev' +[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Command.theorem', expected 'Lean.Parser.Command.def' +[PrettyPrinter.parenthesize] parenthesizing (cont := (none, [anonymous])) + (Term.byTactic + "by" + (Tactic.tacticSeq + (Tactic.tacticSeq1Indented + [(Mathlib.Tactic.Conv.convLHS + "conv_lhs" + [] + [] + "=>" + (Tactic.Conv.convSeq + (Tactic.Conv.convSeq1Indented + [(Tactic.Conv.convRw__ + "rw" + [] + (Tactic.rwRuleSeq "[" [(Tactic.rwRule ["←"] (Term.app `apply_symm_apply [`f `y]))] "]"))]))) + [] + (Tactic.rwSeq "rw" [] (Tactic.rwRuleSeq "[" [(Tactic.rwRule [] `apply_eq_iff_eq)] "]") [])]))) +[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Tactic.tacticSeq1Indented', expected 'Lean.Parser.Tactic.tacticSeqBracketed' +[PrettyPrinter.parenthesize] parenthesizing (cont := (none, [anonymous])) + (Tactic.rwSeq "rw" [] (Tactic.rwRuleSeq "[" [(Tactic.rwRule [] `apply_eq_iff_eq)] "]") []) +[PrettyPrinter.parenthesize] parenthesizing (cont := (none, [anonymous])) + `apply_eq_iff_eq +[PrettyPrinter.parenthesize] ...precedences are 0 >? 1024, (none, [anonymous]) <=? (none, [anonymous]) +[PrettyPrinter.parenthesize] ...precedences are 0 >? 1022 +[PrettyPrinter.parenthesize] parenthesizing (cont := (none, [anonymous])) + (Mathlib.Tactic.Conv.convLHS + "conv_lhs" + [] + [] + "=>" + (Tactic.Conv.convSeq + (Tactic.Conv.convSeq1Indented + [(Tactic.Conv.convRw__ + "rw" + [] + (Tactic.rwRuleSeq "[" [(Tactic.rwRule ["←"] (Term.app `apply_symm_apply [`f `y]))] "]"))]))) +[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Tactic.Conv.convSeq1Indented', expected 'Lean.Parser.Tactic.Conv.convSeqBracketed' +[PrettyPrinter.parenthesize] parenthesizing (cont := (none, [anonymous])) + (Term.app `apply_symm_apply [`f `y]) +[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'ident', expected 'Lean.Parser.Term.namedArgument' +[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'ident', expected 'Lean.Parser.Term.ellipsis' +[PrettyPrinter.parenthesize] parenthesizing (cont := (none, [anonymous])) + `y +[PrettyPrinter.parenthesize] ...precedences are 1023 >? 1024, (none, [anonymous]) <=? (none, [anonymous]) +[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'ident', expected 'Lean.Parser.Term.namedArgument' +[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'ident', expected 'Lean.Parser.Term.ellipsis' +[PrettyPrinter.parenthesize] parenthesizing (cont := (some 1024, term)) + `f +[PrettyPrinter.parenthesize] ...precedences are 1023 >? 1024, (none, [anonymous]) <=? (some 1024, term) +[PrettyPrinter.parenthesize] parenthesizing (cont := (some 1022, term)) + `apply_symm_apply +[PrettyPrinter.parenthesize] ...precedences are 1024 >? 1024, (none, [anonymous]) <=? (some 1022, term) +[PrettyPrinter.parenthesize] ...precedences are 0 >? 1022, (some 1023, term) <=? (none, [anonymous]) +[PrettyPrinter.parenthesize.backtrack] unexpected node kind '«←»', expected 'patternIgnore' +[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Command.declValSimple', expected 'Lean.Parser.Command.declValEqns' +[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Command.declValSimple', expected 'Lean.Parser.Command.whereStructInst' +[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Command.theorem', expected 'Lean.Parser.Command.opaque' +[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Command.theorem', expected 'Lean.Parser.Command.instance' +[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Command.theorem', expected 'Lean.Parser.Command.axiom' +[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Command.theorem', expected 'Lean.Parser.Command.example' +[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Command.theorem', expected 'Lean.Parser.Command.inductive' +[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Command.theorem', expected 'Lean.Parser.Command.classInductive' +[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Command.theorem', expected 'Lean.Parser.Command.structure'-/-- failed to format: format: uncaught backtrack exception +theorem + apply_eq_iff_eq_symm_apply + { α β : Sort _ } ( f : α ≃ β ) { x : α } { y : β } : f x = y ↔ x = f . symm y + := by conv_lhs => rw [ ← apply_symm_apply f y ] rw [ apply_eq_iff_eq ] + +@[simp] +theorem cast_apply {α β} (h : α = β) (x : α) : Equiv.cast h x = cast h x := + rfl + +@[simp] +theorem cast_symm {α β} (h : α = β) : (Equiv.cast h).symm = Equiv.cast h.symm := + rfl + +@[simp] +theorem cast_refl {α} (h : α = α := rfl) : Equiv.cast h = Equiv.refl α := + rfl + +@[simp] +theorem cast_trans {α β γ} (h : α = β) (h2 : β = γ) : (Equiv.cast h).trans (Equiv.cast h2) = Equiv.cast (h.trans h2) := + ext fun x => by + substs h h2 + rfl + +theorem cast_eq_iff_heq {α β} (h : α = β) {a : α} {b : β} : Equiv.cast h a = b ↔ HEq a b := by + subst h + simp + +theorem symm_apply_eq {α β} (e : α ≃ β) {x y} : e.symm x = y ↔ x = e y := + ⟨fun H => by simp [H.symm], fun H => by simp [H]⟩ + +theorem eq_symm_apply {α β} (e : α ≃ β) {x y} : y = e.symm x ↔ e y = x := + (eq_comm.trans e.symm_apply_eq).trans eq_comm + +@[simp] +theorem symm_symm (e : α ≃ β) : e.symm.symm = e := by + cases e + rfl + +@[simp] +theorem trans_refl (e : α ≃ β) : e.trans (Equiv.refl β) = e := by + cases e + rfl + +@[simp] +theorem refl_symm : (Equiv.refl α).symm = Equiv.refl α := + rfl + +@[simp] +theorem refl_trans (e : α ≃ β) : (Equiv.refl α).trans e = e := by + cases e + rfl + +@[simp] +theorem symm_trans_self (e : α ≃ β) : e.symm.trans e = Equiv.refl β := + ext (by simp) + +@[simp] +theorem self_trans_symm (e : α ≃ β) : e.trans e.symm = Equiv.refl α := + ext (by simp) + +theorem trans_assoc {δ} (ab : α ≃ β) (bc : β ≃ γ) (cd : γ ≃ δ) : (ab.trans bc).trans cd = ab.trans (bc.trans cd) := + Equiv.ext fun a => rfl + +theorem left_inverse_symm (f : Equiv α β) : LeftInverse f.symm f := + f.left_inv + +theorem right_inverse_symm (f : Equiv α β) : Function.RightInverse f.symm f := + f.right_inv + +theorem injective_comp (e : α ≃ β) (f : β → γ) : Injective (f ∘ e) ↔ Injective f := + EquivLike.injective_comp e f + +theorem comp_injective (f : α → β) (e : β ≃ γ) : Injective (e ∘ f) ↔ Injective f := + EquivLike.comp_injective f e + +theorem surjective_comp (e : α ≃ β) (f : β → γ) : Surjective (f ∘ e) ↔ Surjective f := + EquivLike.surjective_comp e f + +theorem comp_surjective (f : α → β) (e : β ≃ γ) : Surjective (e ∘ f) ↔ Surjective f := + EquivLike.comp_surjective f e + +theorem bijective_comp (e : α ≃ β) (f : β → γ) : Bijective (f ∘ e) ↔ Bijective f := + EquivLike.bijective_comp e f + +theorem comp_bijective (f : α → β) (e : β ≃ γ) : Bijective (e ∘ f) ↔ Bijective f := + EquivLike.comp_bijective f e + +/-- If `α` is equivalent to `β` and `γ` is equivalent to `δ`, then the type of equivalences `α ≃ γ` +is equivalent to the type of equivalences `β ≃ δ`. -/ +def equivCongr {δ} (ab : α ≃ β) (cd : γ ≃ δ) : α ≃ γ ≃ (β ≃ δ) := + ⟨fun ac => (ab.symm.trans ac).trans cd, fun bd => ab.trans <| bd.trans <| cd.symm, fun ac => by + ext x + simp, fun ac => by + ext x + simp⟩ + +@[simp] +theorem equiv_congr_refl {α β} : (Equiv.refl α).equivCongr (Equiv.refl β) = Equiv.refl (α ≃ β) := by + ext + rfl + +@[simp] +theorem equiv_congr_symm {δ} (ab : α ≃ β) (cd : γ ≃ δ) : (ab.equivCongr cd).symm = ab.symm.equivCongr cd.symm := by + ext + rfl + +@[simp] +theorem equiv_congr_trans {δ ε ζ} (ab : α ≃ β) (de : δ ≃ ε) (bc : β ≃ γ) (ef : ε ≃ ζ) : + (ab.equivCongr de).trans (bc.equivCongr ef) = (ab.trans bc).equivCongr (de.trans ef) := by + ext + rfl + +@[simp] +theorem equiv_congr_refl_left {α β γ} (bg : β ≃ γ) (e : α ≃ β) : (Equiv.refl α).equivCongr bg e = e.trans bg := + rfl + +@[simp] +theorem equiv_congr_refl_right {α β} (ab e : α ≃ β) : ab.equivCongr (Equiv.refl β) e = ab.symm.trans e := + rfl + +@[simp] +theorem equiv_congr_apply_apply {δ} (ab : α ≃ β) (cd : γ ≃ δ) (e : α ≃ γ) (x) : + ab.equivCongr cd e x = cd (e (ab.symm x)) := + rfl + +section PermCongr + +variable {α' β' : Type _} (e : α' ≃ β') + +/-- If `α` is equivalent to `β`, then `perm α` is equivalent to `perm β`. -/ +def permCongr : Perm α' ≃ Perm β' := + equivCongr e e + +theorem perm_congr_def (p : Equiv.Perm α') : e.permCongr p = (e.symm.trans p).trans e := + rfl + +@[simp] +theorem perm_congr_refl : e.permCongr (Equiv.refl _) = Equiv.refl _ := by simp [perm_congr_def] + +@[simp] +theorem perm_congr_symm : e.permCongr.symm = e.symm.permCongr := + rfl + +@[simp] +theorem perm_congr_apply (p : Equiv.Perm α') (x) : e.permCongr p x = e (p (e.symm x)) := + rfl + +theorem perm_congr_symm_apply (p : Equiv.Perm β') (x) : e.permCongr.symm p x = e.symm (p (e x)) := + rfl + +theorem perm_congr_trans (p p' : Equiv.Perm α') : (e.permCongr p).trans (e.permCongr p') = e.permCongr (p.trans p') := + by + ext + simp + +end PermCongr + +/-- Two empty types are equivalent. -/ +def equivOfIsEmpty (α β : Sort _) [IsEmpty α] [IsEmpty β] : α ≃ β := + ⟨isEmptyElim, isEmptyElim, isEmptyElim, isEmptyElim⟩ + +/-- If `α` is an empty type, then it is equivalent to the `empty` type. -/ +def equivEmpty (α : Sort u) [IsEmpty α] : α ≃ Empty := + equivOfIsEmpty α _ + +/-- If `α` is an empty type, then it is equivalent to the `pempty` type in any universe. -/ +def equivPempty (α : Sort v) [IsEmpty α] : α ≃ PEmpty.{u} := + equivOfIsEmpty α _ + +/-- `α` is equivalent to an empty type iff `α` is empty. -/ +def equivEmptyEquiv (α : Sort u) : α ≃ Empty ≃ IsEmpty α := + ⟨fun e => Function.is_empty e, @equivEmpty α, fun e => ext fun x => (e x).elim, fun p => rfl⟩ + +/-- The `Sort` of proofs of a false proposition is equivalent to `pempty`. -/ +def propEquivPempty {p : Prop} (h : ¬p) : p ≃ PEmpty := + @equivPempty p <| IsEmpty.prop_iff.2 h + +/-- If both `α` and `β` have a unique element, then `α ≃ β`. -/ +def equivOfUnique (α β : Sort _) [Unique α] [Unique β] : α ≃ β where + toFun := default + invFun := default + left_inv _ := Subsingleton.elim _ _ + right_inv _ := Subsingleton.elim _ _ + +/-- If `α` has a unique element, then it is equivalent to any `punit`. -/ +def equivPunit (α : Sort _) [Unique α] : α ≃ PUnit.{v} := + equivOfUnique α _ + +/-- The `Sort` of proofs of a true proposition is equivalent to `punit`. -/ +def propEquivPunit {p : Prop} (h : p) : p ≃ PUnit := + @equivPunit p <| uniqueProp h + +/-- `ulift α` is equivalent to `α`. -/ +@[simps (config := { fullyApplied := false }) apply symmApply] +protected def ulift {α : Type v} : ULift.{u} α ≃ α := + ⟨ULift.down, ULift.up, ULift.up_down, fun a => rfl⟩ + +/-- `plift α` is equivalent to `α`. -/ +@[simps (config := { fullyApplied := false }) apply symmApply] +protected def plift : PLift α ≃ α := + ⟨PLift.down, PLift.up, PLift.up_down, PLift.down_up⟩ + +/-- equivalence of propositions is the same as iff -/ +def ofIff {P Q : Prop} (h : P ↔ Q) : P ≃ Q where + toFun := h.mp + invFun := h.mpr + left_inv x := rfl + right_inv y := rfl + +/-- If `α₁` is equivalent to `α₂` and `β₁` is equivalent to `β₂`, then the type of maps `α₁ → β₁` +is equivalent to the type of maps `α₂ → β₂`. -/ +@[congr, simps apply] +def arrowCongr {α₁ β₁ α₂ β₂ : Sort _} (e₁ : α₁ ≃ α₂) (e₂ : β₁ ≃ β₂) : (α₁ → β₁) ≃ (α₂ → β₂) where + toFun f := e₂ ∘ f ∘ e₁.symm + invFun f := e₂.symm ∘ f ∘ e₁ + left_inv f := funext fun x => by simp + right_inv f := funext fun x => by simp + +theorem arrow_congr_comp {α₁ β₁ γ₁ α₂ β₂ γ₂ : Sort _} (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) (ec : γ₁ ≃ γ₂) (f : α₁ → β₁) + (g : β₁ → γ₁) : arrowCongr ea ec (g ∘ f) = arrowCongr eb ec g ∘ arrowCongr ea eb f := by + ext + simp only [comp, arrow_congr_apply, eb.symm_apply_apply] + +@[simp] +theorem arrow_congr_refl {α β : Sort _} : arrowCongr (Equiv.refl α) (Equiv.refl β) = Equiv.refl (α → β) := + rfl + +@[simp] +theorem arrow_congr_trans {α₁ β₁ α₂ β₂ α₃ β₃ : Sort _} (e₁ : α₁ ≃ α₂) (e₁' : β₁ ≃ β₂) (e₂ : α₂ ≃ α₃) (e₂' : β₂ ≃ β₃) : + arrowCongr (e₁.trans e₂) (e₁'.trans e₂') = (arrowCongr e₁ e₁').trans (arrowCongr e₂ e₂') := + rfl + +@[simp] +theorem arrow_congr_symm {α₁ β₁ α₂ β₂ : Sort _} (e₁ : α₁ ≃ α₂) (e₂ : β₁ ≃ β₂) : + (arrowCongr e₁ e₂).symm = arrowCongr e₁.symm e₂.symm := + rfl + +/-- A version of `equiv.arrow_congr` in `Type`, rather than `Sort`. + +The `equiv_rw` tactic is not able to use the default `Sort` level `equiv.arrow_congr`, +because Lean's universe rules will not unify `?l_1` with `imax (1 ?m_1)`. +-/ +@[congr, simps apply] +def arrowCongr' {α₁ β₁ α₂ β₂ : Type _} (hα : α₁ ≃ α₂) (hβ : β₁ ≃ β₂) : (α₁ → β₁) ≃ (α₂ → β₂) := + Equiv.arrowCongr hα hβ + +@[simp] +theorem arrow_congr'_refl {α β : Type _} : arrowCongr' (Equiv.refl α) (Equiv.refl β) = Equiv.refl (α → β) := + rfl + +@[simp] +theorem arrow_congr'_trans {α₁ β₁ α₂ β₂ α₃ β₃ : Type _} (e₁ : α₁ ≃ α₂) (e₁' : β₁ ≃ β₂) (e₂ : α₂ ≃ α₃) (e₂' : β₂ ≃ β₃) : + arrowCongr' (e₁.trans e₂) (e₁'.trans e₂') = (arrowCongr' e₁ e₁').trans (arrowCongr' e₂ e₂') := + rfl + +@[simp] +theorem arrow_congr'_symm {α₁ β₁ α₂ β₂ : Type _} (e₁ : α₁ ≃ α₂) (e₂ : β₁ ≃ β₂) : + (arrowCongr' e₁ e₂).symm = arrowCongr' e₁.symm e₂.symm := + rfl + +/-- Conjugate a map `f : α → α` by an equivalence `α ≃ β`. -/ +@[simps apply] +def conj (e : α ≃ β) : (α → α) ≃ (β → β) := + arrowCongr e e + +@[simp] +theorem conj_refl : conj (Equiv.refl α) = Equiv.refl (α → α) := + rfl + +@[simp] +theorem conj_symm (e : α ≃ β) : e.conj.symm = e.symm.conj := + rfl + +@[simp] +theorem conj_trans (e₁ : α ≃ β) (e₂ : β ≃ γ) : (e₁.trans e₂).conj = e₁.conj.trans e₂.conj := + rfl + +-- This should not be a simp lemma as long as `(∘)` is reducible: +-- when `(∘)` is reducible, Lean can unify `f₁ ∘ f₂` with any `g` using +-- `f₁ := g` and `f₂ := λ x, x`. This causes nontermination. +theorem conj_comp (e : α ≃ β) (f₁ f₂ : α → α) : e.conj (f₁ ∘ f₂) = e.conj f₁ ∘ e.conj f₂ := by apply arrow_congr_comp + +theorem eq_comp_symm {α β γ} (e : α ≃ β) (f : β → γ) (g : α → γ) : f = g ∘ e.symm ↔ f ∘ e = g := + (e.arrowCongr (Equiv.refl γ)).symm_apply_eq.symm + +theorem comp_symm_eq {α β γ} (e : α ≃ β) (f : β → γ) (g : α → γ) : g ∘ e.symm = f ↔ g = f ∘ e := + (e.arrowCongr (Equiv.refl γ)).eq_symm_apply.symm + +theorem eq_symm_comp {α β γ} (e : α ≃ β) (f : γ → α) (g : γ → β) : f = e.symm ∘ g ↔ e ∘ f = g := + ((Equiv.refl γ).arrowCongr e).eq_symm_apply + +theorem symm_comp_eq {α β γ} (e : α ≃ β) (f : γ → α) (g : γ → β) : e.symm ∘ g = f ↔ g = e ∘ f := + ((Equiv.refl γ).arrowCongr e).symm_apply_eq + +/-- `punit` sorts in any two universes are equivalent. -/ +def punitEquivPunit : PUnit.{v} ≃ PUnit.{w} := + ⟨fun _ => PUnit.unit, fun _ => PUnit.unit, fun u => by + cases u + rfl, fun u => by + cases u + rfl⟩ + +/-- `Prop` is noncomputably equivalent to `bool`. -/ +noncomputable def propEquivBool : Prop ≃ Bool := + ⟨fun p => @decide p (Classical.propDecidable _), fun b => b, fun p => by simp, fun b => by simp⟩ + +section + +/-- The sort of maps to `punit.{v}` is equivalent to `punit.{w}`. -/ +def arrowPunitEquivPunit (α : Sort _) : (α → PUnit.{v}) ≃ PUnit.{w} := + ⟨fun f => PUnit.unit, fun u f => PUnit.unit, fun f => by + funext x + cases f x + rfl, fun u => by + cases u + rfl⟩ + +/-- If `α` is `subsingleton` and `a : α`, then the type of dependent functions `Π (i : α), β +i` is equivalent to `β i`. -/ +@[simps] +def piSubsingleton {α} (β : α → Sort _) [Subsingleton α] (a : α) : (∀ a', β a') ≃ β a where + toFun := eval a + invFun x b := cast (congr_arg β <| Subsingleton.elim a b) x + left_inv f := + funext fun b => by + rw [Subsingleton.elim b a] + rfl + right_inv b := rfl + +/-- If `α` has a unique term, then the type of function `α → β` is equivalent to `β`. -/ +@[simps (config := { fullyApplied := false })] +def funUnique (α β) [Unique α] : (α → β) ≃ β := + piSubsingleton _ default + +/-- The sort of maps from `punit` is equivalent to the codomain. -/ +def punitArrowEquiv (α : Sort _) : (PUnit.{u} → α) ≃ α := + funUnique _ _ + +/-- The sort of maps from `true` is equivalent to the codomain. -/ +def trueArrowEquiv (α : Sort _) : (True → α) ≃ α := + funUnique _ _ + +/-- The sort of maps from a type that `is_empty` is equivalent to `punit`. -/ +def arrowPunitOfIsEmpty (α β : Sort _) [IsEmpty α] : (α → β) ≃ PUnit.{u} := + ⟨fun f => PUnit.unit, fun u => isEmptyElim, fun f => funext isEmptyElim, fun u => by + cases u + rfl⟩ + +/-- The sort of maps from `empty` is equivalent to `punit`. -/ +def emptyArrowEquivPunit (α : Sort _) : (Empty → α) ≃ PUnit.{u} := + arrowPunitOfIsEmpty _ _ + +/-- The sort of maps from `pempty` is equivalent to `punit`. -/ +def pemptyArrowEquivPunit (α : Sort _) : (PEmpty → α) ≃ PUnit.{u} := + arrowPunitOfIsEmpty _ _ + +/-- The sort of maps from `false` is equivalent to `punit`. -/ +def falseArrowEquivPunit (α : Sort _) : (False → α) ≃ PUnit.{u} := + arrowPunitOfIsEmpty _ _ + +end + +section + +/-- A `psigma`-type is equivalent to the corresponding `sigma`-type. -/ +@[simps apply symmApply] +def psigmaEquivSigma {α} (β : α → Type _) : (Σ'i, β i) ≃ Σi, β i := + ⟨fun a => ⟨a.1, a.2⟩, fun a => ⟨a.1, a.2⟩, fun ⟨a, b⟩ => rfl, fun ⟨a, b⟩ => rfl⟩ + +/-- A `psigma`-type is equivalent to the corresponding `sigma`-type. -/ +@[simps apply symmApply] +def psigmaEquivSigmaPlift {α} (β : α → Sort _) : (Σ'i, β i) ≃ Σi : PLift α, PLift (β i.down) := + ⟨fun a => ⟨PLift.up a.1, PLift.up a.2⟩, fun a => ⟨a.1.down, a.2.down⟩, fun ⟨a, b⟩ => rfl, fun ⟨⟨a⟩, ⟨b⟩⟩ => rfl⟩ + +/-- A family of equivalences `Π a, β₁ a ≃ β₂ a` generates an equivalence between `Σ' a, β₁ a` and +`Σ' a, β₂ a`. -/ +@[simps apply] +def psigmaCongrRight {α} {β₁ β₂ : α → Sort _} (F : ∀ a, β₁ a ≃ β₂ a) : (Σ'a, β₁ a) ≃ Σ'a, β₂ a := + ⟨fun a => ⟨a.1, F a.1 a.2⟩, fun a => ⟨a.1, (F a.1).symm a.2⟩, fun ⟨a, b⟩ => + congr_arg (PSigma.mk a) <| symm_apply_apply (F a) b, fun ⟨a, b⟩ => + congr_arg (PSigma.mk a) <| apply_symm_apply (F a) b⟩ + +@[simp] +theorem psigma_congr_right_trans {α} {β₁ β₂ β₃ : α → Sort _} (F : ∀ a, β₁ a ≃ β₂ a) (G : ∀ a, β₂ a ≃ β₃ a) : + (psigmaCongrRight F).trans (psigmaCongrRight G) = psigmaCongrRight fun a => (F a).trans (G a) := by + ext1 x + cases x + rfl + +@[simp] +theorem psigma_congr_right_symm {α} {β₁ β₂ : α → Sort _} (F : ∀ a, β₁ a ≃ β₂ a) : + (psigmaCongrRight F).symm = psigmaCongrRight fun a => (F a).symm := by + ext1 x + cases x + rfl + +@[simp] +theorem psigma_congr_right_refl {α} {β : α → Sort _} : + (psigmaCongrRight fun a => Equiv.refl (β a)) = Equiv.refl (Σ'a, β a) := by + ext1 x + cases x + rfl + +/-- A family of equivalences `Π a, β₁ a ≃ β₂ a` generates an equivalence between `Σ a, β₁ a` and +`Σ a, β₂ a`. -/ +@[simps apply] +def sigmaCongrRight {α} {β₁ β₂ : α → Type _} (F : ∀ a, β₁ a ≃ β₂ a) : (Σa, β₁ a) ≃ Σa, β₂ a := + ⟨fun a => ⟨a.1, F a.1 a.2⟩, fun a => ⟨a.1, (F a.1).symm a.2⟩, fun ⟨a, b⟩ => + congr_arg (Sigma.mk a) <| symm_apply_apply (F a) b, fun ⟨a, b⟩ => + congr_arg (Sigma.mk a) <| apply_symm_apply (F a) b⟩ + +@[simp] +theorem sigma_congr_right_trans {α} {β₁ β₂ β₃ : α → Type _} (F : ∀ a, β₁ a ≃ β₂ a) (G : ∀ a, β₂ a ≃ β₃ a) : + (sigmaCongrRight F).trans (sigmaCongrRight G) = sigmaCongrRight fun a => (F a).trans (G a) := by + ext1 x + cases x + rfl + +@[simp] +theorem sigma_congr_right_symm {α} {β₁ β₂ : α → Type _} (F : ∀ a, β₁ a ≃ β₂ a) : + (sigmaCongrRight F).symm = sigmaCongrRight fun a => (F a).symm := by + ext1 x + cases x + rfl + +@[simp] +theorem sigma_congr_right_refl {α} {β : α → Type _} : + (sigmaCongrRight fun a => Equiv.refl (β a)) = Equiv.refl (Σa, β a) := by + ext1 x + cases x + rfl + +/-- A `psigma` with `Prop` fibers is equivalent to the subtype. -/ +def psigmaEquivSubtype {α : Type v} (P : α → Prop) : (Σ'i, P i) ≃ Subtype P where + toFun x := ⟨x.1, x.2⟩ + invFun x := ⟨x.1, x.2⟩ + left_inv x := by + cases x + rfl + right_inv x := by + cases x + rfl + +/-- A `sigma` with `plift` fibers is equivalent to the subtype. -/ +def sigmaPliftEquivSubtype {α : Type v} (P : α → Prop) : (Σi, PLift (P i)) ≃ Subtype P := + ((psigmaEquivSigma _).symm.trans (psigmaCongrRight fun a => Equiv.plift)).trans (psigmaEquivSubtype P) + +/-- A `sigma` with `λ i, ulift (plift (P i))` fibers is equivalent to `{ x // P x }`. +Variant of `sigma_plift_equiv_subtype`. +-/ +def sigmaUliftPliftEquivSubtype {α : Type v} (P : α → Prop) : (Σi, ULift (PLift (P i))) ≃ Subtype P := + (sigmaCongrRight fun a => Equiv.ulift).trans (sigmaPliftEquivSubtype P) + +namespace Perm + +/-- A family of permutations `Π a, perm (β a)` generates a permuation `perm (Σ a, β₁ a)`. -/ +@[reducible] +def sigmaCongrRight {α} {β : α → Sort _} (F : ∀ a, Perm (β a)) : Perm (Σa, β a) := + Equiv.sigmaCongrRight F + +@[simp] +theorem sigma_congr_right_trans {α} {β : α → Sort _} (F : ∀ a, Perm (β a)) (G : ∀ a, Perm (β a)) : + (sigmaCongrRight F).trans (sigmaCongrRight G) = sigmaCongrRight fun a => (F a).trans (G a) := + Equiv.sigma_congr_right_trans F G + +@[simp] +theorem sigma_congr_right_symm {α} {β : α → Sort _} (F : ∀ a, Perm (β a)) : + (sigmaCongrRight F).symm = sigmaCongrRight fun a => (F a).symm := + Equiv.sigma_congr_right_symm F + +@[simp] +theorem sigma_congr_right_refl {α} {β : α → Sort _} : + (sigmaCongrRight fun a => Equiv.refl (β a)) = Equiv.refl (Σa, β a) := + Equiv.sigma_congr_right_refl + +end Perm + +/-- An equivalence `f : α₁ ≃ α₂` generates an equivalence between `Σ a, β (f a)` and `Σ a, β a`. -/ +@[simps apply] +def sigmaCongrLeft {α₁ α₂} {β : α₂ → Sort _} (e : α₁ ≃ α₂) : (Σa : α₁, β (e a)) ≃ Σa : α₂, β a := + ⟨fun a => ⟨e a.1, a.2⟩, fun a => ⟨e.symm a.1, @Eq.ndrec β a.2 (e.right_inv a.1).symm⟩, fun ⟨a, b⟩ => + match (motive := ∀ (a') (h : a' = a), @Sigma.mk _ (β ∘ e) _ (@Eq.ndrec β b (congr_arg e h.symm)) = ⟨a, b⟩) + e.symm (e a), e.left_inv a with + | _, rfl => rfl, + fun ⟨a, b⟩ => + match (motive := ∀ (a') (h : a' = a), Sigma.mk a' (@Eq.ndrec β b h.symm) = ⟨a, b⟩) e (e.symm a), _ with + | _, rfl => rfl⟩ + +/-- Transporting a sigma type through an equivalence of the base -/ +def sigmaCongrLeft' {α₁ α₂} {β : α₁ → Sort _} (f : α₁ ≃ α₂) : (Σa : α₁, β a) ≃ Σa : α₂, β (f.symm a) := + (sigmaCongrLeft f.symm).symm + +/-- Transporting a sigma type through an equivalence of the base and a family of equivalences +of matching fibers -/ +def sigmaCongr {α₁ α₂} {β₁ : α₁ → Sort _} {β₂ : α₂ → Sort _} (f : α₁ ≃ α₂) (F : ∀ a, β₁ a ≃ β₂ (f a)) : + Sigma β₁ ≃ Sigma β₂ := + (sigmaCongrRight F).trans (sigmaCongrLeft f) + +/-- `sigma` type with a constant fiber is equivalent to the product. -/ +@[simps apply symmApply] +def sigmaEquivProd (α β : Type _) : (Σ_ : α, β) ≃ α × β := + ⟨fun a => ⟨a.1, a.2⟩, fun a => ⟨a.1, a.2⟩, fun ⟨a, b⟩ => rfl, fun ⟨a, b⟩ => rfl⟩ + +/-- If each fiber of a `sigma` type is equivalent to a fixed type, then the sigma type +is equivalent to the product. -/ +def sigmaEquivProdOfEquiv {α β} {β₁ : α → Sort _} (F : ∀ a, β₁ a ≃ β) : Sigma β₁ ≃ α × β := + (sigmaCongrRight F).trans (sigmaEquivProd α β) + +/-- Dependent product of types is associative up to an equivalence. -/ +def sigmaAssoc {α : Type _} {β : α → Type _} (γ : ∀ a : α, β a → Type _) : + (Σab : Σa : α, β a, γ ab.1 ab.2) ≃ Σa : α, Σb : β a, γ a b where + toFun x := ⟨x.1.1, ⟨x.1.2, x.2⟩⟩ + invFun x := ⟨⟨x.1, x.2.1⟩, x.2.2⟩ + left_inv := fun ⟨⟨a, b⟩, c⟩ => rfl + right_inv := fun ⟨a, ⟨b, c⟩⟩ => rfl + +end + +protected theorem exists_unique_congr {p : α → Prop} {q : β → Prop} (f : α ≃ β) (h : ∀ {x}, p x ↔ q (f x)) : + (∃! x, p x) ↔ ∃! y, q y := by + constructor + · rintro ⟨a, ha₁, ha₂⟩ + exact ⟨f a, h.1 ha₁, fun b hb => f.symm_apply_eq.1 (ha₂ (f.symm b) (h.2 (by simpa using hb)))⟩ + + · rintro ⟨b, hb₁, hb₂⟩ + exact ⟨f.symm b, h.2 (by simpa using hb₁), fun y hy => (eq_symm_apply f).2 (hb₂ _ (h.1 hy))⟩ + + +protected theorem exists_unique_congr_left' {p : α → Prop} (f : α ≃ β) : (∃! x, p x) ↔ ∃! y, p (f.symm y) := + Equiv.exists_unique_congr f fun x => by simp + +protected theorem exists_unique_congr_left {p : β → Prop} (f : α ≃ β) : (∃! x, p (f x)) ↔ ∃! y, p y := + (Equiv.exists_unique_congr_left' f.symm).symm + +/- failed to parenthesize: parenthesize: uncaught backtrack exception +[PrettyPrinter.parenthesize.input] (Command.declaration + (Command.declModifiers [] [] [(Command.protected "protected")] [] [] []) + (Command.theorem + "theorem" + (Command.declId `forall_congr []) + (Command.declSig + [(Term.implicitBinder "{" [`p] [":" (Term.arrow `α "→" (Term.prop "Prop"))] "}") + (Term.implicitBinder "{" [`q] [":" (Term.arrow `β "→" (Term.prop "Prop"))] "}") + (Term.explicitBinder "(" [`f] [":" (Logic.Equiv.Defs.«term_≃_» `α " ≃ " `β)] [] ")") + (Term.explicitBinder + "(" + [`h] + [":" + (Term.forall + "∀" + [(Term.implicitBinder "{" [`x] [] "}")] + [] + "," + («term_↔_» (Term.app `p [`x]) "↔" (Term.app `q [(Term.app `f [`x])])))] + [] + ")")] + (Term.typeSpec + ":" + («term_↔_» + (Term.forall "∀" [`x] [] "," (Term.app `p [`x])) + "↔" + (Term.forall "∀" [`y] [] "," (Term.app `q [`y]))))) + (Command.declValSimple + ":=" + (Term.byTactic + "by" + (Tactic.tacticSeq + (Tactic.tacticSeq1Indented + [(Tactic.«tactic_<;>_» (Tactic.constructor "constructor") "<;>" (Tactic.intro "intro" [`h₂ `x])) + [] + («tactic___;_» + (cdotTk (patternIgnore (token.«·» "·"))) + [(group + (Tactic.rwSeq "rw" [] (Tactic.rwRuleSeq "[" [(Tactic.rwRule ["←"] (Term.app `f.right_inv [`x]))] "]") []) + []) + (group (Tactic.apply "apply" `h.mp) []) + (group (Tactic.apply "apply" `h₂) [])]) + [] + (Tactic.apply "apply" `h.mpr) + [] + (Tactic.apply "apply" `h₂)]))) + []) + [] + [])) +[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Command.theorem', expected 'Lean.Parser.Command.abbrev' +[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Command.theorem', expected 'Lean.Parser.Command.def' +[PrettyPrinter.parenthesize] parenthesizing (cont := (none, [anonymous])) + (Term.byTactic + "by" + (Tactic.tacticSeq + (Tactic.tacticSeq1Indented + [(Tactic.«tactic_<;>_» (Tactic.constructor "constructor") "<;>" (Tactic.intro "intro" [`h₂ `x])) + [] + («tactic___;_» + (cdotTk (patternIgnore (token.«·» "·"))) + [(group + (Tactic.rwSeq "rw" [] (Tactic.rwRuleSeq "[" [(Tactic.rwRule ["←"] (Term.app `f.right_inv [`x]))] "]") []) + []) + (group (Tactic.apply "apply" `h.mp) []) + (group (Tactic.apply "apply" `h₂) [])]) + [] + (Tactic.apply "apply" `h.mpr) + [] + (Tactic.apply "apply" `h₂)]))) +[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Tactic.tacticSeq1Indented', expected 'Lean.Parser.Tactic.tacticSeqBracketed' +[PrettyPrinter.parenthesize] parenthesizing (cont := (none, [anonymous])) + (Tactic.apply "apply" `h₂) +[PrettyPrinter.parenthesize] parenthesizing (cont := (none, [anonymous])) + `h₂ +[PrettyPrinter.parenthesize] ...precedences are 0 >? 1024, (none, [anonymous]) <=? (none, [anonymous]) +[PrettyPrinter.parenthesize] ...precedences are 0 >? 1022 +[PrettyPrinter.parenthesize] parenthesizing (cont := (none, [anonymous])) + (Tactic.apply "apply" `h.mpr) +[PrettyPrinter.parenthesize] parenthesizing (cont := (none, [anonymous])) + `h.mpr +[PrettyPrinter.parenthesize] ...precedences are 0 >? 1024, (none, [anonymous]) <=? (none, [anonymous]) +[PrettyPrinter.parenthesize] ...precedences are 0 >? 1022 +[PrettyPrinter.parenthesize] parenthesizing (cont := (none, [anonymous])) + («tactic___;_» + (cdotTk (patternIgnore (token.«·» "·"))) + [(group + (Tactic.rwSeq "rw" [] (Tactic.rwRuleSeq "[" [(Tactic.rwRule ["←"] (Term.app `f.right_inv [`x]))] "]") []) + []) + (group (Tactic.apply "apply" `h.mp) []) + (group (Tactic.apply "apply" `h₂) [])]) +[PrettyPrinter.parenthesize] parenthesizing (cont := (none, [anonymous])) + (Tactic.apply "apply" `h₂) +[PrettyPrinter.parenthesize] parenthesizing (cont := (none, [anonymous])) + `h₂ +[PrettyPrinter.parenthesize] ...precedences are 0 >? 1024, (none, [anonymous]) <=? (none, [anonymous]) +[PrettyPrinter.parenthesize] ...precedences are 0 >? 1022 +[PrettyPrinter.parenthesize] parenthesizing (cont := (some 1022, tactic)) + (Tactic.apply "apply" `h.mp) +[PrettyPrinter.parenthesize] parenthesizing (cont := (none, [anonymous])) + `h.mp +[PrettyPrinter.parenthesize] ...precedences are 0 >? 1024, (none, [anonymous]) <=? (none, [anonymous]) +[PrettyPrinter.parenthesize] ...precedences are 0 >? 1022 +[PrettyPrinter.parenthesize] parenthesizing (cont := (some 1022, tactic)) + (Tactic.rwSeq "rw" [] (Tactic.rwRuleSeq "[" [(Tactic.rwRule ["←"] (Term.app `f.right_inv [`x]))] "]") []) +[PrettyPrinter.parenthesize] parenthesizing (cont := (none, [anonymous])) + (Term.app `f.right_inv [`x]) +[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'ident', expected 'Lean.Parser.Term.namedArgument' +[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'ident', expected 'Lean.Parser.Term.ellipsis' +[PrettyPrinter.parenthesize] parenthesizing (cont := (none, [anonymous])) + `x +[PrettyPrinter.parenthesize] ...precedences are 1023 >? 1024, (none, [anonymous]) <=? (none, [anonymous]) +[PrettyPrinter.parenthesize] parenthesizing (cont := (some 1022, term)) + `f.right_inv +[PrettyPrinter.parenthesize] ...precedences are 1024 >? 1024, (none, [anonymous]) <=? (some 1022, term) +[PrettyPrinter.parenthesize] ...precedences are 0 >? 1022, (some 1023, term) <=? (none, [anonymous]) +[PrettyPrinter.parenthesize.backtrack] unexpected node kind '«←»', expected 'patternIgnore' +[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Command.declValSimple', expected 'Lean.Parser.Command.declValEqns' +[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Command.declValSimple', expected 'Lean.Parser.Command.whereStructInst' +[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Command.theorem', expected 'Lean.Parser.Command.opaque' +[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Command.theorem', expected 'Lean.Parser.Command.instance' +[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Command.theorem', expected 'Lean.Parser.Command.axiom' +[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Command.theorem', expected 'Lean.Parser.Command.example' +[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Command.theorem', expected 'Lean.Parser.Command.inductive' +[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Command.theorem', expected 'Lean.Parser.Command.classInductive' +[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Command.theorem', expected 'Lean.Parser.Command.structure'-/-- failed to format: format: uncaught backtrack exception +protected + theorem + forall_congr + { p : α → Prop } { q : β → Prop } ( f : α ≃ β ) ( h : ∀ { x } , p x ↔ q f x ) : ∀ x , p x ↔ ∀ y , q y + := by constructor <;> intro h₂ x · rw [ ← f.right_inv x ] apply h.mp apply h₂ apply h.mpr apply h₂ + +protected theorem forall_congr' {p : α → Prop} {q : β → Prop} (f : α ≃ β) (h : ∀ {x}, p (f.symm x) ↔ q x) : + (∀ x, p x) ↔ ∀ y, q y := + (Equiv.forall_congr f.symm fun x => h.symm).symm + +-- We next build some higher arity versions of `equiv.forall_congr`. +-- Although they appear to just be repeated applications of `equiv.forall_congr`, +-- unification of metavariables works better with these versions. +-- In particular, they are necessary in `equiv_rw`. +-- (Stopping at ternary functions seems reasonable: at least in 1-categorical mathematics, +-- it's rare to have axioms involving more than 3 elements at once.) +universe ua1 ua2 ub1 ub2 ug1 ug2 + +variable {α₁ : Sort ua1} {α₂ : Sort ua2} {β₁ : Sort ub1} {β₂ : Sort ub2} {γ₁ : Sort ug1} {γ₂ : Sort ug2} + +protected theorem forall₂_congr {p : α₁ → β₁ → Prop} {q : α₂ → β₂ → Prop} (eα : α₁ ≃ α₂) (eβ : β₁ ≃ β₂) + (h : ∀ {x y}, p x y ↔ q (eα x) (eβ y)) : (∀ x y, p x y) ↔ ∀ x y, q x y := by + apply Equiv.forall_congr + intros + apply Equiv.forall_congr + intros + apply h + +protected theorem forall₂_congr' {p : α₁ → β₁ → Prop} {q : α₂ → β₂ → Prop} (eα : α₁ ≃ α₂) (eβ : β₁ ≃ β₂) + (h : ∀ {x y}, p (eα.symm x) (eβ.symm y) ↔ q x y) : (∀ x y, p x y) ↔ ∀ x y, q x y := + (Equiv.forall₂_congr eα.symm eβ.symm fun x y => h.symm).symm + +protected theorem forall₃_congr {p : α₁ → β₁ → γ₁ → Prop} {q : α₂ → β₂ → γ₂ → Prop} (eα : α₁ ≃ α₂) (eβ : β₁ ≃ β₂) + (eγ : γ₁ ≃ γ₂) (h : ∀ {x y z}, p x y z ↔ q (eα x) (eβ y) (eγ z)) : (∀ x y z, p x y z) ↔ ∀ x y z, q x y z := by + apply Equiv.forall₂_congr + intros + apply Equiv.forall_congr + intros + apply h + +protected theorem forall₃_congr' {p : α₁ → β₁ → γ₁ → Prop} {q : α₂ → β₂ → γ₂ → Prop} (eα : α₁ ≃ α₂) (eβ : β₁ ≃ β₂) + (eγ : γ₁ ≃ γ₂) (h : ∀ {x y z}, p (eα.symm x) (eβ.symm y) (eγ.symm z) ↔ q x y z) : + (∀ x y z, p x y z) ↔ ∀ x y z, q x y z := + (Equiv.forall₃_congr eα.symm eβ.symm eγ.symm fun x y z => h.symm).symm + +protected theorem forall_congr_left' {p : α → Prop} (f : α ≃ β) : (∀ x, p x) ↔ ∀ y, p (f.symm y) := + Equiv.forall_congr f fun x => by simp + +protected theorem forall_congr_left {p : β → Prop} (f : α ≃ β) : (∀ x, p (f x)) ↔ ∀ y, p y := + (Equiv.forall_congr_left' f.symm).symm + +protected theorem exists_congr_left {α β} (f : α ≃ β) {p : α → Prop} : (∃ a, p a) ↔ ∃ b, p (f.symm b) := + ⟨fun ⟨a, h⟩ => ⟨f a, by simpa using h⟩, fun ⟨b, h⟩ => ⟨_, h⟩⟩ + +end Equiv + +namespace Quot + +/-- An equivalence `e : α ≃ β` generates an equivalence between quotient spaces, +if `ra a₁ a₂ ↔ rb (e a₁) (e a₂). -/ +protected def congr {ra : α → α → Prop} {rb : β → β → Prop} (e : α ≃ β) (eq : ∀ a₁ a₂, ra a₁ a₂ ↔ rb (e a₁) (e a₂)) : + Quot ra ≃ Quot rb where + toFun := Quot.map e fun a₁ a₂ => (Eq a₁ a₂).1 + invFun := + Quot.map e.symm fun b₁ b₂ h => + (Eq (e.symm b₁) (e.symm b₂)).2 ((e.apply_symm_apply b₁).symm ▸ (e.apply_symm_apply b₂).symm ▸ h) + left_inv := by + rintro ⟨a⟩ + dsimp only [Quot.map] + simp only [Equiv.symm_apply_apply] + right_inv := by + rintro ⟨a⟩ + dsimp only [Quot.map] + simp only [Equiv.apply_symm_apply] + +@[simp] +theorem congr_mk {ra : α → α → Prop} {rb : β → β → Prop} (e : α ≃ β) (eq : ∀ a₁ a₂ : α, ra a₁ a₂ ↔ rb (e a₁) (e a₂)) + (a : α) : Quot.congr e Eq (Quot.mk ra a) = Quot.mk rb (e a) := + rfl + +/-- Quotients are congruent on equivalences under equality of their relation. +An alternative is just to use rewriting with `eq`, but then computational proofs get stuck. -/ +protected def congrRight {r r' : α → α → Prop} (eq : ∀ a₁ a₂, r a₁ a₂ ↔ r' a₁ a₂) : Quot r ≃ Quot r' := + Quot.congr (Equiv.refl α) Eq + +/-- An equivalence `e : α ≃ β` generates an equivalence between the quotient space of `α` +by a relation `ra` and the quotient space of `β` by the image of this relation under `e`. -/ +protected def congrLeft {r : α → α → Prop} (e : α ≃ β) : Quot r ≃ Quot fun b b' => r (e.symm b) (e.symm b') := + @Quot.congr α β r (fun b b' => r (e.symm b) (e.symm b')) e fun a₁ a₂ => by simp only [e.symm_apply_apply] + +end Quot + +namespace Quotient + +/-- An equivalence `e : α ≃ β` generates an equivalence between quotient spaces, +if `ra a₁ a₂ ↔ rb (e a₁) (e a₂). -/ +protected def congr {ra : Setoid α} {rb : Setoid β} (e : α ≃ β) + (eq : ∀ a₁ a₂, @Setoid.R α ra a₁ a₂ ↔ @Setoid.R β rb (e a₁) (e a₂)) : Quotient ra ≃ Quotient rb := + Quot.congr e Eq + +@[simp] +theorem congr_mk {ra : Setoid α} {rb : Setoid β} (e : α ≃ β) (eq : ∀ a₁ a₂ : α, Setoid.R a₁ a₂ ↔ Setoid.R (e a₁) (e a₂)) + (a : α) : Quotient.congr e Eq (Quotient.mk a) = Quotient.mk (e a) := + rfl + +/-- Quotients are congruent on equivalences under equality of their relation. +An alternative is just to use rewriting with `eq`, but then computational proofs get stuck. -/ +protected def congrRight {r r' : Setoid α} (eq : ∀ a₁ a₂, @Setoid.R α r a₁ a₂ ↔ @Setoid.R α r' a₁ a₂) : + Quotient r ≃ Quotient r' := + Quot.congrRight Eq + +end Quotient From 246a5c8be88d6d349eb69673ba7afe9a33b31a43 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Mon, 7 Nov 2022 16:08:41 +1100 Subject: [PATCH 008/127] rename --- Mathlib/{Data => Logic}/Equiv/Defs.lean | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) rename Mathlib/{Data => Logic}/Equiv/Defs.lean (99%) diff --git a/Mathlib/Data/Equiv/Defs.lean b/Mathlib/Logic/Equiv/Defs.lean similarity index 99% rename from Mathlib/Data/Equiv/Defs.lean rename to Mathlib/Logic/Equiv/Defs.lean index 89df0a9b9a52c..3005c7a4f01c5 100644 --- a/Mathlib/Data/Equiv/Defs.lean +++ b/Mathlib/Logic/Equiv/Defs.lean @@ -171,10 +171,10 @@ protected theorem bijective (e : α ≃ β) : Bijective e := EquivLike.bijective e protected theorem subsingleton (e : α ≃ β) [Subsingleton β] : Subsingleton α := - e.Injective.Subsingleton + e.injective.subsingleton protected theorem subsingleton.symm (e : α ≃ β) [Subsingleton α] : Subsingleton β := - e.symm.Injective.Subsingleton + e.symm.injective.subsingleton theorem subsingleton_congr (e : α ≃ β) : Subsingleton α ↔ Subsingleton β := ⟨fun h => e.symm.subsingleton, fun h => e.subsingleton⟩ @@ -193,7 +193,7 @@ theorem Perm.subsingleton_eq_refl [Subsingleton α] (e : Perm α) : e = Equiv.re /-- Transfer `decidable_eq` across an equivalence. -/ protected def decidableEq (e : α ≃ β) [DecidableEq β] : DecidableEq α := - e.Injective.DecidableEq + e.injective.decidableEq theorem nonempty_congr (e : α ≃ β) : Nonempty α ↔ Nonempty β := Nonempty.congr e e.symm @@ -207,7 +207,7 @@ protected def inhabited [Inhabited β] (e : α ≃ β) : Inhabited α := /-- If `α ≃ β` and `β` is a singleton type, then so is `α`. -/ protected def unique [Unique β] (e : α ≃ β) : Unique α := - e.symm.Surjective.unique + e.symm.surjective.unique /-- Equivalence between equal types. -/ protected def cast {α β : Sort _} (h : α = β) : α ≃ β := From b21f8e896f685c5c60a0619408fee3e594f6572f Mon Sep 17 00:00:00 2001 From: Vierkantor Date: Tue, 8 Nov 2022 11:39:56 +0000 Subject: [PATCH 009/127] =?UTF-8?q?Fix:=20`Function.Surjective.unique`=20w?= =?UTF-8?q?as=20inferring=20`=CE=B1=20:=20Prop`=20when=20it=20can=20be=20a?= =?UTF-8?q?ny=20`Sort`?= MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit --- Mathlib/Logic/Unique.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Logic/Unique.lean b/Mathlib/Logic/Unique.lean index 5d8f732d4348e..568b3d703eed6 100644 --- a/Mathlib/Logic/Unique.lean +++ b/Mathlib/Logic/Unique.lean @@ -205,7 +205,7 @@ protected theorem Surjective.subsingleton [Subsingleton α] (hf : Surjective f) /-- If the domain of a surjective function is a singleton, then the codomain is a singleton as well. -/ -protected def Surjective.unique (hf : Surjective f) [Unique α] : Unique β := +protected def Surjective.unique (f : α → β) (hf : Surjective f) [Unique.{u} α] : Unique β := @Unique.mk' _ ⟨f default⟩ hf.subsingleton /-- If `α` is inhabited and admits an injective map to a subsingleton type, then `α` is `Unique`. -/ From d88f4d8cf2b39a69cd9f13cf1dc61fb987c61f90 Mon Sep 17 00:00:00 2001 From: Vierkantor Date: Tue, 8 Nov 2022 14:50:06 +0000 Subject: [PATCH 010/127] Builds until the Quot section --- Mathlib/Logic/Equiv/Defs.lean | 412 +++++++++------------------------- Mathlib/Logic/Unique.lean | 2 +- 2 files changed, 112 insertions(+), 302 deletions(-) diff --git a/Mathlib/Logic/Equiv/Defs.lean b/Mathlib/Logic/Equiv/Defs.lean index 3005c7a4f01c5..fd02493c0cfb9 100644 --- a/Mathlib/Logic/Equiv/Defs.lean +++ b/Mathlib/Logic/Equiv/Defs.lean @@ -3,12 +3,15 @@ Copyright (c) 2015 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Mario Carneiro -/ -import Mathlib.Logic.Unique import Mathlib.Data.FunLike.Equiv +import Mathlib.Init.Data.Bool.Lemmas +import Mathlib.Logic.Unique +import Mathlib.Tactic.Conv import Mathlib.Tactic.Relation.Rfl import Mathlib.Tactic.Relation.Symm import Mathlib.Tactic.Relation.Trans import Mathlib.Tactic.Simps.Basic +import Mathlib.Tactic.Substs /-! # Equivalence between types @@ -130,7 +133,7 @@ theorem Perm.ext_iff {σ τ : Equiv.Perm α} : σ = τ ↔ ∀ x, σ x = τ x := /-- Any type is equivalent to itself. -/ @[refl] protected def refl (α : Sort _) : α ≃ α := - ⟨id, id, fun x => rfl, fun x => rfl⟩ + ⟨id, id, fun _ => rfl, fun _ => rfl⟩ instance inhabited' : Inhabited (α ≃ α) := ⟨Equiv.refl α⟩ @@ -177,13 +180,13 @@ protected theorem subsingleton.symm (e : α ≃ β) [Subsingleton α] : Subsingl e.symm.injective.subsingleton theorem subsingleton_congr (e : α ≃ β) : Subsingleton α ↔ Subsingleton β := - ⟨fun h => e.symm.subsingleton, fun h => e.subsingleton⟩ + ⟨fun _ => e.symm.subsingleton, fun _ => e.subsingleton⟩ instance equiv_subsingleton_cod [Subsingleton β] : Subsingleton (α ≃ β) := - ⟨fun f g => Equiv.ext fun x => Subsingleton.elim _ _⟩ + ⟨fun _ _ => Equiv.ext fun _ => Subsingleton.elim _ _⟩ instance equiv_subsingleton_dom [Subsingleton α] : Subsingleton (α ≃ β) := - ⟨fun f g => Equiv.ext fun x => @Subsingleton.elim _ (Equiv.subsingleton.symm f) _ _⟩ + ⟨fun f _ => Equiv.ext fun _ => @Subsingleton.elim _ (Equiv.subsingleton.symm f) _ _⟩ instance permUnique [Subsingleton α] : Unique (Perm α) := uniqueOfSubsingleton (Equiv.refl α) @@ -222,21 +225,27 @@ theorem coe_fn_symm_mk (f : α → β) (g l r) : ((Equiv.mk f g l r).symm : β rfl @[simp] -theorem coe_refl : ⇑(Equiv.refl α) = id := +theorem coe_refl : (Equiv.refl α : α → α) = id := rfl /-- This cannot be a `simp` lemmas as it incorrectly matches against `e : α ≃ synonym α`, when `synonym α` is semireducible. This makes a mess of `multiplicative.of_add` etc. -/ -theorem Perm.coe_subsingleton {α : Type _} [Subsingleton α] (e : Perm α) : ⇑e = id := by - rw [perm.subsingleton_eq_refl e, coe_refl] +theorem Perm.coe_subsingleton {α : Type _} [Subsingleton α] (e : Perm α) : (e : α → α) = id := by + rw [Perm.subsingleton_eq_refl e, coe_refl] +-- porting note: marking this as `@[simp]` because `simp` doesn't fire on `coe_refl` +-- in an expression such as `Equiv.refl a x` +@[simp] theorem refl_apply (x : α) : Equiv.refl α x = x := rfl @[simp] -theorem coe_trans (f : α ≃ β) (g : β ≃ γ) : ⇑(f.trans g) = g ∘ f := +theorem coe_trans (f : α ≃ β) (g : β ≃ γ) : (f.trans g : α → γ) = g ∘ f := rfl +-- porting note: marking this as `@[simp]` because `simp` doesn't fire on `coe_trans` +-- in an expression such as `Equiv.trans f g x` +@[simp] theorem trans_apply (f : α ≃ β) (g : β ≃ γ) (a : α) : (f.trans g) a = g (f a) := rfl @@ -262,122 +271,16 @@ theorem symm_trans_apply (f : α ≃ β) (g : β ≃ γ) (a : γ) : (f.trans g). -- The `simp` attribute is needed to make this a `dsimp` lemma. -- `simp` will always rewrite with `equiv.symm_symm` before this has a chance to fire. -@[simp, nolint simp_nf] +@[simp, nolint simpNF] theorem symm_symm_apply (f : α ≃ β) (b : α) : f.symm.symm b = f b := rfl theorem apply_eq_iff_eq (f : α ≃ β) {x y : α} : f x = f y ↔ x = y := EquivLike.apply_eq_iff_eq f -/- failed to parenthesize: parenthesize: uncaught backtrack exception -[PrettyPrinter.parenthesize.input] (Command.declaration - (Command.declModifiers [] [] [] [] [] []) - (Command.theorem - "theorem" - (Command.declId `apply_eq_iff_eq_symm_apply []) - (Command.declSig - [(Term.implicitBinder "{" [`α `β] [":" (Term.sort "Sort" [(Level.hole "_")])] "}") - (Term.explicitBinder "(" [`f] [":" (Logic.Equiv.Defs.«term_≃_» `α " ≃ " `β)] [] ")") - (Term.implicitBinder "{" [`x] [":" `α] "}") - (Term.implicitBinder "{" [`y] [":" `β] "}")] - (Term.typeSpec - ":" - («term_↔_» - («term_=_» (Term.app `f [`x]) "=" `y) - "↔" - («term_=_» `x "=" (Term.app (Term.proj `f "." `symm) [`y]))))) - (Command.declValSimple - ":=" - (Term.byTactic - "by" - (Tactic.tacticSeq - (Tactic.tacticSeq1Indented - [(Mathlib.Tactic.Conv.convLHS - "conv_lhs" - [] - [] - "=>" - (Tactic.Conv.convSeq - (Tactic.Conv.convSeq1Indented - [(Tactic.Conv.convRw__ - "rw" - [] - (Tactic.rwRuleSeq "[" [(Tactic.rwRule ["←"] (Term.app `apply_symm_apply [`f `y]))] "]"))]))) - [] - (Tactic.rwSeq "rw" [] (Tactic.rwRuleSeq "[" [(Tactic.rwRule [] `apply_eq_iff_eq)] "]") [])]))) - []) - [] - [])) -[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Command.theorem', expected 'Lean.Parser.Command.abbrev' -[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Command.theorem', expected 'Lean.Parser.Command.def' -[PrettyPrinter.parenthesize] parenthesizing (cont := (none, [anonymous])) - (Term.byTactic - "by" - (Tactic.tacticSeq - (Tactic.tacticSeq1Indented - [(Mathlib.Tactic.Conv.convLHS - "conv_lhs" - [] - [] - "=>" - (Tactic.Conv.convSeq - (Tactic.Conv.convSeq1Indented - [(Tactic.Conv.convRw__ - "rw" - [] - (Tactic.rwRuleSeq "[" [(Tactic.rwRule ["←"] (Term.app `apply_symm_apply [`f `y]))] "]"))]))) - [] - (Tactic.rwSeq "rw" [] (Tactic.rwRuleSeq "[" [(Tactic.rwRule [] `apply_eq_iff_eq)] "]") [])]))) -[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Tactic.tacticSeq1Indented', expected 'Lean.Parser.Tactic.tacticSeqBracketed' -[PrettyPrinter.parenthesize] parenthesizing (cont := (none, [anonymous])) - (Tactic.rwSeq "rw" [] (Tactic.rwRuleSeq "[" [(Tactic.rwRule [] `apply_eq_iff_eq)] "]") []) -[PrettyPrinter.parenthesize] parenthesizing (cont := (none, [anonymous])) - `apply_eq_iff_eq -[PrettyPrinter.parenthesize] ...precedences are 0 >? 1024, (none, [anonymous]) <=? (none, [anonymous]) -[PrettyPrinter.parenthesize] ...precedences are 0 >? 1022 -[PrettyPrinter.parenthesize] parenthesizing (cont := (none, [anonymous])) - (Mathlib.Tactic.Conv.convLHS - "conv_lhs" - [] - [] - "=>" - (Tactic.Conv.convSeq - (Tactic.Conv.convSeq1Indented - [(Tactic.Conv.convRw__ - "rw" - [] - (Tactic.rwRuleSeq "[" [(Tactic.rwRule ["←"] (Term.app `apply_symm_apply [`f `y]))] "]"))]))) -[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Tactic.Conv.convSeq1Indented', expected 'Lean.Parser.Tactic.Conv.convSeqBracketed' -[PrettyPrinter.parenthesize] parenthesizing (cont := (none, [anonymous])) - (Term.app `apply_symm_apply [`f `y]) -[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'ident', expected 'Lean.Parser.Term.namedArgument' -[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'ident', expected 'Lean.Parser.Term.ellipsis' -[PrettyPrinter.parenthesize] parenthesizing (cont := (none, [anonymous])) - `y -[PrettyPrinter.parenthesize] ...precedences are 1023 >? 1024, (none, [anonymous]) <=? (none, [anonymous]) -[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'ident', expected 'Lean.Parser.Term.namedArgument' -[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'ident', expected 'Lean.Parser.Term.ellipsis' -[PrettyPrinter.parenthesize] parenthesizing (cont := (some 1024, term)) - `f -[PrettyPrinter.parenthesize] ...precedences are 1023 >? 1024, (none, [anonymous]) <=? (some 1024, term) -[PrettyPrinter.parenthesize] parenthesizing (cont := (some 1022, term)) - `apply_symm_apply -[PrettyPrinter.parenthesize] ...precedences are 1024 >? 1024, (none, [anonymous]) <=? (some 1022, term) -[PrettyPrinter.parenthesize] ...precedences are 0 >? 1022, (some 1023, term) <=? (none, [anonymous]) -[PrettyPrinter.parenthesize.backtrack] unexpected node kind '«←»', expected 'patternIgnore' -[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Command.declValSimple', expected 'Lean.Parser.Command.declValEqns' -[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Command.declValSimple', expected 'Lean.Parser.Command.whereStructInst' -[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Command.theorem', expected 'Lean.Parser.Command.opaque' -[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Command.theorem', expected 'Lean.Parser.Command.instance' -[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Command.theorem', expected 'Lean.Parser.Command.axiom' -[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Command.theorem', expected 'Lean.Parser.Command.example' -[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Command.theorem', expected 'Lean.Parser.Command.inductive' -[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Command.theorem', expected 'Lean.Parser.Command.classInductive' -[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Command.theorem', expected 'Lean.Parser.Command.structure'-/-- failed to format: format: uncaught backtrack exception -theorem - apply_eq_iff_eq_symm_apply - { α β : Sort _ } ( f : α ≃ β ) { x : α } { y : β } : f x = y ↔ x = f . symm y - := by conv_lhs => rw [ ← apply_symm_apply f y ] rw [ apply_eq_iff_eq ] +theorem apply_eq_iff_eq_symm_apply (f : α ≃ β) : f x = y ↔ x = f.symm y + := by conv_lhs => rw [ ← apply_symm_apply f y ] + rw [ apply_eq_iff_eq ] @[simp] theorem cast_apply {α β} (h : α = β) (x : α) : Equiv.cast h x = cast h x := @@ -392,14 +295,14 @@ theorem cast_refl {α} (h : α = α := rfl) : Equiv.cast h = Equiv.refl α := rfl @[simp] -theorem cast_trans {α β γ} (h : α = β) (h2 : β = γ) : (Equiv.cast h).trans (Equiv.cast h2) = Equiv.cast (h.trans h2) := - ext fun x => by - substs h h2 - rfl +theorem cast_trans {α β γ} (h : α = β) (h2 : β = γ) : + (Equiv.cast h).trans (Equiv.cast h2) = Equiv.cast (h.trans h2) := + ext fun x => by substs h h2 + rfl -theorem cast_eq_iff_heq {α β} (h : α = β) {a : α} {b : β} : Equiv.cast h a = b ↔ HEq a b := by - subst h - simp +theorem cast_eq_iff_heq {α β} (h : α = β) {a : α} {b : β} : Equiv.cast h a = b ↔ HEq a b := +by subst h + simp [coe_refl] theorem symm_apply_eq {α β} (e : α ≃ β) {x y} : e.symm x = y ↔ x = e y := ⟨fun H => by simp [H.symm], fun H => by simp [H]⟩ @@ -435,7 +338,7 @@ theorem self_trans_symm (e : α ≃ β) : e.trans e.symm = Equiv.refl α := ext (by simp) theorem trans_assoc {δ} (ab : α ≃ β) (bc : β ≃ γ) (cd : γ ≃ δ) : (ab.trans bc).trans cd = ab.trans (bc.trans cd) := - Equiv.ext fun a => rfl + Equiv.ext fun _ => rfl theorem left_inverse_symm (f : Equiv α β) : LeftInverse f.symm f := f.left_inv @@ -545,31 +448,33 @@ def equivPempty (α : Sort v) [IsEmpty α] : α ≃ PEmpty.{u} := /-- `α` is equivalent to an empty type iff `α` is empty. -/ def equivEmptyEquiv (α : Sort u) : α ≃ Empty ≃ IsEmpty α := - ⟨fun e => Function.is_empty e, @equivEmpty α, fun e => ext fun x => (e x).elim, fun p => rfl⟩ + ⟨fun e => Function.is_empty e, @equivEmpty α, fun e => ext fun x => (e x).elim, fun _ => rfl⟩ /-- The `Sort` of proofs of a false proposition is equivalent to `pempty`. -/ def propEquivPempty {p : Prop} (h : ¬p) : p ≃ PEmpty := @equivPempty p <| IsEmpty.prop_iff.2 h /-- If both `α` and `β` have a unique element, then `α ≃ β`. -/ -def equivOfUnique (α β : Sort _) [Unique α] [Unique β] : α ≃ β where +def equivOfUnique (α β : Sort _) [Unique.{u} α] [Unique.{v} β] : α ≃ β where toFun := default invFun := default left_inv _ := Subsingleton.elim _ _ right_inv _ := Subsingleton.elim _ _ /-- If `α` has a unique element, then it is equivalent to any `punit`. -/ -def equivPunit (α : Sort _) [Unique α] : α ≃ PUnit.{v} := +def equivPunit (α : Sort u) [Unique α] : α ≃ PUnit.{v} := equivOfUnique α _ /-- The `Sort` of proofs of a true proposition is equivalent to `punit`. -/ -def propEquivPunit {p : Prop} (h : p) : p ≃ PUnit := - @equivPunit p <| uniqueProp h +def propEquivPunit {p : Prop} (h : p) : p ≃ PUnit.{0} := +-- porting note: the following causes an error "unknown free variable [anonymous]" +-- @equivPunit p <| uniqueProp h +⟨default, λ _ => h, λ _ => Subsingleton.elim _ _, λ _ => Subsingleton.elim _ _⟩ /-- `ulift α` is equivalent to `α`. -/ @[simps (config := { fullyApplied := false }) apply symmApply] protected def ulift {α : Type v} : ULift.{u} α ≃ α := - ⟨ULift.down, ULift.up, ULift.up_down, fun a => rfl⟩ + ⟨ULift.down, ULift.up, ULift.up_down, fun _ => rfl⟩ /-- `plift α` is equivalent to `α`. -/ @[simps (config := { fullyApplied := false }) apply symmApply] @@ -580,22 +485,25 @@ protected def plift : PLift α ≃ α := def ofIff {P Q : Prop} (h : P ↔ Q) : P ≃ Q where toFun := h.mp invFun := h.mpr - left_inv x := rfl - right_inv y := rfl + left_inv _ := rfl + right_inv _ := rfl /-- If `α₁` is equivalent to `α₂` and `β₁` is equivalent to `β₂`, then the type of maps `α₁ → β₁` is equivalent to the type of maps `α₂ → β₂`. -/ -@[congr, simps apply] +-- porting note: removing `congr` attribute +@[simps apply] def arrowCongr {α₁ β₁ α₂ β₂ : Sort _} (e₁ : α₁ ≃ α₂) (e₂ : β₁ ≃ β₂) : (α₁ → β₁) ≃ (α₂ → β₂) where toFun f := e₂ ∘ f ∘ e₁.symm invFun f := e₂.symm ∘ f ∘ e₁ left_inv f := funext fun x => by simp right_inv f := funext fun x => by simp +#align equiv.arrow_congr_apply Equiv.arrowCongr_apply + theorem arrow_congr_comp {α₁ β₁ γ₁ α₂ β₂ γ₂ : Sort _} (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) (ec : γ₁ ≃ γ₂) (f : α₁ → β₁) (g : β₁ → γ₁) : arrowCongr ea ec (g ∘ f) = arrowCongr eb ec g ∘ arrowCongr ea eb f := by ext - simp only [comp, arrow_congr_apply, eb.symm_apply_apply] + simp only [comp, arrowCongr_apply, eb.symm_apply_apply] @[simp] theorem arrow_congr_refl {α β : Sort _} : arrowCongr (Equiv.refl α) (Equiv.refl β) = Equiv.refl (α → β) := @@ -616,7 +524,8 @@ theorem arrow_congr_symm {α₁ β₁ α₂ β₂ : Sort _} (e₁ : α₁ ≃ α The `equiv_rw` tactic is not able to use the default `Sort` level `equiv.arrow_congr`, because Lean's universe rules will not unify `?l_1` with `imax (1 ?m_1)`. -/ -@[congr, simps apply] +-- porting note: removing `congr` attribute +@[simps apply] def arrowCongr' {α₁ β₁ α₂ β₂ : Type _} (hα : α₁ ≃ α₂) (hβ : β₁ ≃ β₂) : (α₁ → β₁) ≃ (α₂ → β₂) := Equiv.arrowCongr hα hβ @@ -676,20 +585,31 @@ def punitEquivPunit : PUnit.{v} ≃ PUnit.{w} := cases u rfl⟩ +@[simp] lemma Bool.decide_eq_true (b : Bool) {h} : @decide (b = true) h = b := +by cases b + . simp + . simp + /-- `Prop` is noncomputably equivalent to `bool`. -/ -noncomputable def propEquivBool : Prop ≃ Bool := - ⟨fun p => @decide p (Classical.propDecidable _), fun b => b, fun p => by simp, fun b => by simp⟩ +noncomputable def propEquivBool : Prop ≃ Bool where + toFun p := @decide p (Classical.propDecidable _) + invFun b := b + left_inv p := by simp [@Bool.decide_iff p (Classical.propDecidable _)] + right_inv b := by simp section /-- The sort of maps to `punit.{v}` is equivalent to `punit.{w}`. -/ -def arrowPunitEquivPunit (α : Sort _) : (α → PUnit.{v}) ≃ PUnit.{w} := - ⟨fun f => PUnit.unit, fun u f => PUnit.unit, fun f => by +def arrowPunitEquivPunit (α : Sort _) : (α → PUnit.{v}) ≃ PUnit.{w} where + toFun _ := PUnit.unit + invFun _ _ := PUnit.unit + left_inv f := by funext x cases f x - rfl, fun u => by + rfl + right_inv u := by cases u - rfl⟩ + rfl /-- If `α` is `subsingleton` and `a : α`, then the type of dependent functions `Π (i : α), β i` is equivalent to `β i`. -/ @@ -705,22 +625,25 @@ def piSubsingleton {α} (β : α → Sort _) [Subsingleton α] (a : α) : (∀ a /-- If `α` has a unique term, then the type of function `α → β` is equivalent to `β`. -/ @[simps (config := { fullyApplied := false })] -def funUnique (α β) [Unique α] : (α → β) ≃ β := +def funUnique (α β) [Unique.{u} α] : (α → β) ≃ β := piSubsingleton _ default /-- The sort of maps from `punit` is equivalent to the codomain. -/ def punitArrowEquiv (α : Sort _) : (PUnit.{u} → α) ≃ α := - funUnique _ _ + funUnique PUnit.{u} α /-- The sort of maps from `true` is equivalent to the codomain. -/ def trueArrowEquiv (α : Sort _) : (True → α) ≃ α := funUnique _ _ /-- The sort of maps from a type that `is_empty` is equivalent to `punit`. -/ -def arrowPunitOfIsEmpty (α β : Sort _) [IsEmpty α] : (α → β) ≃ PUnit.{u} := - ⟨fun f => PUnit.unit, fun u => isEmptyElim, fun f => funext isEmptyElim, fun u => by +def arrowPunitOfIsEmpty (α β : Sort _) [IsEmpty α] : (α → β) ≃ PUnit.{u} where + toFun _ := PUnit.unit + invFun _ := isEmptyElim + left_inv _ := funext isEmptyElim + right_inv u := by cases u - rfl⟩ + rfl /-- The sort of maps from `empty` is equivalent to `punit`. -/ def emptyArrowEquivPunit (α : Sort _) : (Empty → α) ≃ PUnit.{u} := @@ -740,13 +663,19 @@ section /-- A `psigma`-type is equivalent to the corresponding `sigma`-type. -/ @[simps apply symmApply] -def psigmaEquivSigma {α} (β : α → Type _) : (Σ'i, β i) ≃ Σi, β i := - ⟨fun a => ⟨a.1, a.2⟩, fun a => ⟨a.1, a.2⟩, fun ⟨a, b⟩ => rfl, fun ⟨a, b⟩ => rfl⟩ +def psigmaEquivSigma {α} (β : α → Type _) : (Σ'i, β i) ≃ Σi, β i where + toFun a := ⟨a.1, a.2⟩ + invFun a := ⟨a.1, a.2⟩ + left_inv := λ ⟨_, _⟩ => rfl + right_inv := λ ⟨_, _⟩ => rfl /-- A `psigma`-type is equivalent to the corresponding `sigma`-type. -/ @[simps apply symmApply] -def psigmaEquivSigmaPlift {α} (β : α → Sort _) : (Σ'i, β i) ≃ Σi : PLift α, PLift (β i.down) := - ⟨fun a => ⟨PLift.up a.1, PLift.up a.2⟩, fun a => ⟨a.1.down, a.2.down⟩, fun ⟨a, b⟩ => rfl, fun ⟨⟨a⟩, ⟨b⟩⟩ => rfl⟩ +def psigmaEquivSigmaPlift {α} (β : α → Sort _) : (Σ'i, β i) ≃ Σi : PLift α, PLift (β i.down) where + toFun a := ⟨PLift.up a.1, PLift.up a.2⟩ + invFun a := ⟨a.1.down, a.2.down⟩ + left_inv := λ ⟨_, _⟩ => rfl + right_inv := λ ⟨_, _⟩ => rfl /-- A family of equivalences `Π a, β₁ a ≃ β₂ a` generates an equivalence between `Σ' a, β₁ a` and `Σ' a, β₂ a`. -/ @@ -819,13 +748,13 @@ def psigmaEquivSubtype {α : Type v} (P : α → Prop) : (Σ'i, P i) ≃ Subtype /-- A `sigma` with `plift` fibers is equivalent to the subtype. -/ def sigmaPliftEquivSubtype {α : Type v} (P : α → Prop) : (Σi, PLift (P i)) ≃ Subtype P := - ((psigmaEquivSigma _).symm.trans (psigmaCongrRight fun a => Equiv.plift)).trans (psigmaEquivSubtype P) + ((psigmaEquivSigma _).symm.trans (psigmaCongrRight fun _ => Equiv.plift)).trans (psigmaEquivSubtype P) /-- A `sigma` with `λ i, ulift (plift (P i))` fibers is equivalent to `{ x // P x }`. Variant of `sigma_plift_equiv_subtype`. -/ def sigmaUliftPliftEquivSubtype {α : Type v} (P : α → Prop) : (Σi, ULift (PLift (P i))) ≃ Subtype P := - (sigmaCongrRight fun a => Equiv.ulift).trans (sigmaPliftEquivSubtype P) + (sigmaCongrRight fun _ => Equiv.ulift).trans (sigmaPliftEquivSubtype P) namespace Perm @@ -853,14 +782,18 @@ end Perm /-- An equivalence `f : α₁ ≃ α₂` generates an equivalence between `Σ a, β (f a)` and `Σ a, β a`. -/ @[simps apply] -def sigmaCongrLeft {α₁ α₂} {β : α₂ → Sort _} (e : α₁ ≃ α₂) : (Σa : α₁, β (e a)) ≃ Σa : α₂, β a := - ⟨fun a => ⟨e a.1, a.2⟩, fun a => ⟨e.symm a.1, @Eq.ndrec β a.2 (e.right_inv a.1).symm⟩, fun ⟨a, b⟩ => - match (motive := ∀ (a') (h : a' = a), @Sigma.mk _ (β ∘ e) _ (@Eq.ndrec β b (congr_arg e h.symm)) = ⟨a, b⟩) +def sigmaCongrLeft {α₁ α₂} {β : α₂ → Sort _} (e : α₁ ≃ α₂) : (Σa : α₁, β (e a)) ≃ Σa : α₂, β a where + toFun a := ⟨e a.1, a.2⟩ + invFun a := ⟨e.symm a.1, (e.right_inv a.1).symm ▸ a.2⟩ + -- porting note: this was a pretty gnarly match already, and it got worse after porting + left_inv := fun ⟨a, b⟩ => + match (motive := ∀ (a') (h : a' = a), @Sigma.mk _ (β ∘ e) _ ((congr_arg e h.symm) ▸ b) = ⟨a, b⟩) e.symm (e a), e.left_inv a with - | _, rfl => rfl, - fun ⟨a, b⟩ => - match (motive := ∀ (a') (h : a' = a), Sigma.mk a' (@Eq.ndrec β b h.symm) = ⟨a, b⟩) e (e.symm a), _ with - | _, rfl => rfl⟩ + | _, rfl => rfl + right_inv := fun ⟨a, b⟩ => + match (motive := ∀ (a') (h : a' = a), Sigma.mk a' (h.symm ▸ b) = ⟨a, b⟩) + e (e.symm a), e.apply_symm_apply _ with + | _, rfl => rfl /-- Transporting a sigma type through an equivalence of the base -/ def sigmaCongrLeft' {α₁ α₂} {β : α₁ → Sort _} (f : α₁ ≃ α₂) : (Σa : α₁, β a) ≃ Σa : α₂, β (f.symm a) := @@ -875,7 +808,7 @@ def sigmaCongr {α₁ α₂} {β₁ : α₁ → Sort _} {β₂ : α₂ → Sort /-- `sigma` type with a constant fiber is equivalent to the product. -/ @[simps apply symmApply] def sigmaEquivProd (α β : Type _) : (Σ_ : α, β) ≃ α × β := - ⟨fun a => ⟨a.1, a.2⟩, fun a => ⟨a.1, a.2⟩, fun ⟨a, b⟩ => rfl, fun ⟨a, b⟩ => rfl⟩ + ⟨fun a => ⟨a.1, a.2⟩, fun a => ⟨a.1, a.2⟩, fun ⟨_, _⟩ => rfl, fun ⟨_, _⟩ => rfl⟩ /-- If each fiber of a `sigma` type is equivalent to a fixed type, then the sigma type is equivalent to the product. -/ @@ -887,8 +820,8 @@ def sigmaAssoc {α : Type _} {β : α → Type _} (γ : ∀ a : α, β a → Typ (Σab : Σa : α, β a, γ ab.1 ab.2) ≃ Σa : α, Σb : β a, γ a b where toFun x := ⟨x.1.1, ⟨x.1.2, x.2⟩⟩ invFun x := ⟨⟨x.1, x.2.1⟩, x.2.2⟩ - left_inv := fun ⟨⟨a, b⟩, c⟩ => rfl - right_inv := fun ⟨a, ⟨b, c⟩⟩ => rfl + left_inv := fun ⟨⟨_, _⟩, _⟩ => rfl + right_inv := fun ⟨_, ⟨_, _⟩⟩ => rfl end @@ -903,146 +836,23 @@ protected theorem exists_unique_congr {p : α → Prop} {q : β → Prop} (f : protected theorem exists_unique_congr_left' {p : α → Prop} (f : α ≃ β) : (∃! x, p x) ↔ ∃! y, p (f.symm y) := - Equiv.exists_unique_congr f fun x => by simp + Equiv.exists_unique_congr f fun {_} => by simp protected theorem exists_unique_congr_left {p : β → Prop} (f : α ≃ β) : (∃! x, p (f x)) ↔ ∃! y, p y := (Equiv.exists_unique_congr_left' f.symm).symm -/- failed to parenthesize: parenthesize: uncaught backtrack exception -[PrettyPrinter.parenthesize.input] (Command.declaration - (Command.declModifiers [] [] [(Command.protected "protected")] [] [] []) - (Command.theorem - "theorem" - (Command.declId `forall_congr []) - (Command.declSig - [(Term.implicitBinder "{" [`p] [":" (Term.arrow `α "→" (Term.prop "Prop"))] "}") - (Term.implicitBinder "{" [`q] [":" (Term.arrow `β "→" (Term.prop "Prop"))] "}") - (Term.explicitBinder "(" [`f] [":" (Logic.Equiv.Defs.«term_≃_» `α " ≃ " `β)] [] ")") - (Term.explicitBinder - "(" - [`h] - [":" - (Term.forall - "∀" - [(Term.implicitBinder "{" [`x] [] "}")] - [] - "," - («term_↔_» (Term.app `p [`x]) "↔" (Term.app `q [(Term.app `f [`x])])))] - [] - ")")] - (Term.typeSpec - ":" - («term_↔_» - (Term.forall "∀" [`x] [] "," (Term.app `p [`x])) - "↔" - (Term.forall "∀" [`y] [] "," (Term.app `q [`y]))))) - (Command.declValSimple - ":=" - (Term.byTactic - "by" - (Tactic.tacticSeq - (Tactic.tacticSeq1Indented - [(Tactic.«tactic_<;>_» (Tactic.constructor "constructor") "<;>" (Tactic.intro "intro" [`h₂ `x])) - [] - («tactic___;_» - (cdotTk (patternIgnore (token.«·» "·"))) - [(group - (Tactic.rwSeq "rw" [] (Tactic.rwRuleSeq "[" [(Tactic.rwRule ["←"] (Term.app `f.right_inv [`x]))] "]") []) - []) - (group (Tactic.apply "apply" `h.mp) []) - (group (Tactic.apply "apply" `h₂) [])]) - [] - (Tactic.apply "apply" `h.mpr) - [] - (Tactic.apply "apply" `h₂)]))) - []) - [] - [])) -[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Command.theorem', expected 'Lean.Parser.Command.abbrev' -[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Command.theorem', expected 'Lean.Parser.Command.def' -[PrettyPrinter.parenthesize] parenthesizing (cont := (none, [anonymous])) - (Term.byTactic - "by" - (Tactic.tacticSeq - (Tactic.tacticSeq1Indented - [(Tactic.«tactic_<;>_» (Tactic.constructor "constructor") "<;>" (Tactic.intro "intro" [`h₂ `x])) - [] - («tactic___;_» - (cdotTk (patternIgnore (token.«·» "·"))) - [(group - (Tactic.rwSeq "rw" [] (Tactic.rwRuleSeq "[" [(Tactic.rwRule ["←"] (Term.app `f.right_inv [`x]))] "]") []) - []) - (group (Tactic.apply "apply" `h.mp) []) - (group (Tactic.apply "apply" `h₂) [])]) - [] - (Tactic.apply "apply" `h.mpr) - [] - (Tactic.apply "apply" `h₂)]))) -[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Tactic.tacticSeq1Indented', expected 'Lean.Parser.Tactic.tacticSeqBracketed' -[PrettyPrinter.parenthesize] parenthesizing (cont := (none, [anonymous])) - (Tactic.apply "apply" `h₂) -[PrettyPrinter.parenthesize] parenthesizing (cont := (none, [anonymous])) - `h₂ -[PrettyPrinter.parenthesize] ...precedences are 0 >? 1024, (none, [anonymous]) <=? (none, [anonymous]) -[PrettyPrinter.parenthesize] ...precedences are 0 >? 1022 -[PrettyPrinter.parenthesize] parenthesizing (cont := (none, [anonymous])) - (Tactic.apply "apply" `h.mpr) -[PrettyPrinter.parenthesize] parenthesizing (cont := (none, [anonymous])) - `h.mpr -[PrettyPrinter.parenthesize] ...precedences are 0 >? 1024, (none, [anonymous]) <=? (none, [anonymous]) -[PrettyPrinter.parenthesize] ...precedences are 0 >? 1022 -[PrettyPrinter.parenthesize] parenthesizing (cont := (none, [anonymous])) - («tactic___;_» - (cdotTk (patternIgnore (token.«·» "·"))) - [(group - (Tactic.rwSeq "rw" [] (Tactic.rwRuleSeq "[" [(Tactic.rwRule ["←"] (Term.app `f.right_inv [`x]))] "]") []) - []) - (group (Tactic.apply "apply" `h.mp) []) - (group (Tactic.apply "apply" `h₂) [])]) -[PrettyPrinter.parenthesize] parenthesizing (cont := (none, [anonymous])) - (Tactic.apply "apply" `h₂) -[PrettyPrinter.parenthesize] parenthesizing (cont := (none, [anonymous])) - `h₂ -[PrettyPrinter.parenthesize] ...precedences are 0 >? 1024, (none, [anonymous]) <=? (none, [anonymous]) -[PrettyPrinter.parenthesize] ...precedences are 0 >? 1022 -[PrettyPrinter.parenthesize] parenthesizing (cont := (some 1022, tactic)) - (Tactic.apply "apply" `h.mp) -[PrettyPrinter.parenthesize] parenthesizing (cont := (none, [anonymous])) - `h.mp -[PrettyPrinter.parenthesize] ...precedences are 0 >? 1024, (none, [anonymous]) <=? (none, [anonymous]) -[PrettyPrinter.parenthesize] ...precedences are 0 >? 1022 -[PrettyPrinter.parenthesize] parenthesizing (cont := (some 1022, tactic)) - (Tactic.rwSeq "rw" [] (Tactic.rwRuleSeq "[" [(Tactic.rwRule ["←"] (Term.app `f.right_inv [`x]))] "]") []) -[PrettyPrinter.parenthesize] parenthesizing (cont := (none, [anonymous])) - (Term.app `f.right_inv [`x]) -[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'ident', expected 'Lean.Parser.Term.namedArgument' -[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'ident', expected 'Lean.Parser.Term.ellipsis' -[PrettyPrinter.parenthesize] parenthesizing (cont := (none, [anonymous])) - `x -[PrettyPrinter.parenthesize] ...precedences are 1023 >? 1024, (none, [anonymous]) <=? (none, [anonymous]) -[PrettyPrinter.parenthesize] parenthesizing (cont := (some 1022, term)) - `f.right_inv -[PrettyPrinter.parenthesize] ...precedences are 1024 >? 1024, (none, [anonymous]) <=? (some 1022, term) -[PrettyPrinter.parenthesize] ...precedences are 0 >? 1022, (some 1023, term) <=? (none, [anonymous]) -[PrettyPrinter.parenthesize.backtrack] unexpected node kind '«←»', expected 'patternIgnore' -[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Command.declValSimple', expected 'Lean.Parser.Command.declValEqns' -[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Command.declValSimple', expected 'Lean.Parser.Command.whereStructInst' -[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Command.theorem', expected 'Lean.Parser.Command.opaque' -[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Command.theorem', expected 'Lean.Parser.Command.instance' -[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Command.theorem', expected 'Lean.Parser.Command.axiom' -[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Command.theorem', expected 'Lean.Parser.Command.example' -[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Command.theorem', expected 'Lean.Parser.Command.inductive' -[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Command.theorem', expected 'Lean.Parser.Command.classInductive' -[PrettyPrinter.parenthesize.backtrack] unexpected node kind 'Lean.Parser.Command.theorem', expected 'Lean.Parser.Command.structure'-/-- failed to format: format: uncaught backtrack exception -protected - theorem - forall_congr - { p : α → Prop } { q : β → Prop } ( f : α ≃ β ) ( h : ∀ { x } , p x ↔ q f x ) : ∀ x , p x ↔ ∀ y , q y - := by constructor <;> intro h₂ x · rw [ ← f.right_inv x ] apply h.mp apply h₂ apply h.mpr apply h₂ +protected theorem forall_congr {p : α → Prop} {q : β → Prop} (f : α ≃ β) + (h : ∀ {x}, p x ↔ q (f x)) : (∀ x, p x) ↔ (∀ y, q y) := +by constructor <;> intro h₂ x + . rw [ ← f.right_inv x ] + apply h.mp + apply h₂ + apply h.mpr + apply h₂ protected theorem forall_congr' {p : α → Prop} {q : β → Prop} (f : α ≃ β) (h : ∀ {x}, p (f.symm x) ↔ q x) : (∀ x, p x) ↔ ∀ y, q y := - (Equiv.forall_congr f.symm fun x => h.symm).symm + (Equiv.forall_congr f.symm fun {_} => h.symm).symm -- We next build some higher arity versions of `equiv.forall_congr`. -- Although they appear to just be repeated applications of `equiv.forall_congr`, @@ -1064,7 +874,7 @@ protected theorem forall₂_congr {p : α₁ → β₁ → Prop} {q : α₂ → protected theorem forall₂_congr' {p : α₁ → β₁ → Prop} {q : α₂ → β₂ → Prop} (eα : α₁ ≃ α₂) (eβ : β₁ ≃ β₂) (h : ∀ {x y}, p (eα.symm x) (eβ.symm y) ↔ q x y) : (∀ x y, p x y) ↔ ∀ x y, q x y := - (Equiv.forall₂_congr eα.symm eβ.symm fun x y => h.symm).symm + (Equiv.forall₂_congr eα.symm eβ.symm fun {_ _} => h.symm).symm protected theorem forall₃_congr {p : α₁ → β₁ → γ₁ → Prop} {q : α₂ → β₂ → γ₂ → Prop} (eα : α₁ ≃ α₂) (eβ : β₁ ≃ β₂) (eγ : γ₁ ≃ γ₂) (h : ∀ {x y z}, p x y z ↔ q (eα x) (eβ y) (eγ z)) : (∀ x y z, p x y z) ↔ ∀ x y z, q x y z := by @@ -1077,10 +887,10 @@ protected theorem forall₃_congr {p : α₁ → β₁ → γ₁ → Prop} {q : protected theorem forall₃_congr' {p : α₁ → β₁ → γ₁ → Prop} {q : α₂ → β₂ → γ₂ → Prop} (eα : α₁ ≃ α₂) (eβ : β₁ ≃ β₂) (eγ : γ₁ ≃ γ₂) (h : ∀ {x y z}, p (eα.symm x) (eβ.symm y) (eγ.symm z) ↔ q x y z) : (∀ x y z, p x y z) ↔ ∀ x y z, q x y z := - (Equiv.forall₃_congr eα.symm eβ.symm eγ.symm fun x y z => h.symm).symm + (Equiv.forall₃_congr eα.symm eβ.symm eγ.symm fun {_ _ _} => h.symm).symm protected theorem forall_congr_left' {p : α → Prop} (f : α ≃ β) : (∀ x, p x) ↔ ∀ y, p (f.symm y) := - Equiv.forall_congr f fun x => by simp + Equiv.forall_congr f fun {_} => by simp protected theorem forall_congr_left {p : β → Prop} (f : α ≃ β) : (∀ x, p (f x)) ↔ ∀ y, p y := (Equiv.forall_congr_left' f.symm).symm diff --git a/Mathlib/Logic/Unique.lean b/Mathlib/Logic/Unique.lean index 568b3d703eed6..7591de4cbe0a9 100644 --- a/Mathlib/Logic/Unique.lean +++ b/Mathlib/Logic/Unique.lean @@ -90,7 +90,7 @@ theorem PUnit.default_eq_unit : (default : PUnit) = PUnit.unit := #align punit.default_eq_star PUnit.default_eq_unit /-- Every provable proposition is unique, as all proofs are equal. -/ -def uniqueProp {p : Prop} (h : p) : Unique p where +def uniqueProp {p : Prop} (h : p) : Unique.{0} p where default := h uniq _ := rfl From aa3e0d165565acc87363407be051e3369734a5eb Mon Sep 17 00:00:00 2001 From: Vierkantor Date: Tue, 8 Nov 2022 14:52:33 +0000 Subject: [PATCH 011/127] Move to appropriate file --- Mathlib/Init/Data/Bool/Lemmas.lean | 7 ++++++- Mathlib/Logic/Equiv/Defs.lean | 5 ----- 2 files changed, 6 insertions(+), 6 deletions(-) diff --git a/Mathlib/Init/Data/Bool/Lemmas.lean b/Mathlib/Init/Data/Bool/Lemmas.lean index 0f12f621dde42..6098354e4ba4a 100644 --- a/Mathlib/Init/Data/Bool/Lemmas.lean +++ b/Mathlib/Init/Data/Bool/Lemmas.lean @@ -123,10 +123,15 @@ theorem coe_sort_false : (↥false : Prop) = False := by simp theorem coe_sort_true : (↥true : Prop) = True := by simp #align coe_sort_tt Bool.coe_sort_true +@[simp] theorem decide_eq_true (b : Bool) {h} : @decide (b = true) h = b := +by cases b + . simp + . simp + theorem decide_iff (p : Prop) [d : Decidable p] : decide p = true ↔ p := by simp #align to_bool_iff Bool.decide_iff -theorem decide_true {p : Prop} [Decidable p] : p → decide p := +theorem decide_true (p : Prop) [Decidable p] : p → decide p := (decide_iff p).2 #align to_bool_true Bool.decide_true #align to_bool_tt Bool.decide_true diff --git a/Mathlib/Logic/Equiv/Defs.lean b/Mathlib/Logic/Equiv/Defs.lean index fd02493c0cfb9..c137db50991e0 100644 --- a/Mathlib/Logic/Equiv/Defs.lean +++ b/Mathlib/Logic/Equiv/Defs.lean @@ -585,11 +585,6 @@ def punitEquivPunit : PUnit.{v} ≃ PUnit.{w} := cases u rfl⟩ -@[simp] lemma Bool.decide_eq_true (b : Bool) {h} : @decide (b = true) h = b := -by cases b - . simp - . simp - /-- `Prop` is noncomputably equivalent to `bool`. -/ noncomputable def propEquivBool : Prop ≃ Bool where toFun p := @decide p (Classical.propDecidable _) From 465c82bb6c45f55660f97c9f93daf92eb38d9575 Mon Sep 17 00:00:00 2001 From: Anne Baanen Date: Tue, 8 Nov 2022 15:09:17 +0000 Subject: [PATCH 012/127] Revert wrong implicitness --- Mathlib/Init/Data/Bool/Lemmas.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Init/Data/Bool/Lemmas.lean b/Mathlib/Init/Data/Bool/Lemmas.lean index 6098354e4ba4a..4c58bd6377010 100644 --- a/Mathlib/Init/Data/Bool/Lemmas.lean +++ b/Mathlib/Init/Data/Bool/Lemmas.lean @@ -131,7 +131,7 @@ by cases b theorem decide_iff (p : Prop) [d : Decidable p] : decide p = true ↔ p := by simp #align to_bool_iff Bool.decide_iff -theorem decide_true (p : Prop) [Decidable p] : p → decide p := +theorem decide_true {p : Prop} [Decidable p] : p → decide p := (decide_iff p).2 #align to_bool_true Bool.decide_true #align to_bool_tt Bool.decide_true From 487f8b66c60e55d373b2b7ce6b931d97d27bd226 Mon Sep 17 00:00:00 2001 From: Vierkantor Date: Tue, 8 Nov 2022 15:19:06 +0000 Subject: [PATCH 013/127] `to_fun_as_coe` is useless --- Mathlib/Logic/Equiv/Defs.lean | 3 +++ 1 file changed, 3 insertions(+) diff --git a/Mathlib/Logic/Equiv/Defs.lean b/Mathlib/Logic/Equiv/Defs.lean index c137db50991e0..2cc5b36cd7eb9 100644 --- a/Mathlib/Logic/Equiv/Defs.lean +++ b/Mathlib/Logic/Equiv/Defs.lean @@ -156,9 +156,12 @@ initialize_simps_projections Equiv (toFun → apply, invFun → symmApply) protected def trans (e₁ : α ≃ β) (e₂ : β ≃ γ) : α ≃ γ := ⟨e₂ ∘ e₁, e₁.symm ∘ e₂.symm, e₂.left_inv.comp e₁.left_inv, e₂.right_inv.comp e₁.right_inv⟩ +-- porting note: this lemma is now useless since coercions are eagerly unfolded +/- @[simp] theorem to_fun_as_coe (e : α ≃ β) : e.toFun = e := rfl +-/ @[simp] theorem inv_fun_as_coe (e : α ≃ β) : e.invFun = e.symm := From 1239fed6780333f1d5861639114bcf892e268c1b Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Fri, 11 Nov 2022 12:57:24 +1100 Subject: [PATCH 014/127] long lines --- Mathlib/Logic/Equiv/Defs.lean | 124 +++++++++++++++++++++------------- 1 file changed, 76 insertions(+), 48 deletions(-) diff --git a/Mathlib/Logic/Equiv/Defs.lean b/Mathlib/Logic/Equiv/Defs.lean index 2cc5b36cd7eb9..3270863fc7fef 100644 --- a/Mathlib/Logic/Equiv/Defs.lean +++ b/Mathlib/Logic/Equiv/Defs.lean @@ -340,7 +340,8 @@ theorem symm_trans_self (e : α ≃ β) : e.symm.trans e = Equiv.refl β := theorem self_trans_symm (e : α ≃ β) : e.trans e.symm = Equiv.refl α := ext (by simp) -theorem trans_assoc {δ} (ab : α ≃ β) (bc : β ≃ γ) (cd : γ ≃ δ) : (ab.trans bc).trans cd = ab.trans (bc.trans cd) := +theorem trans_assoc {δ} (ab : α ≃ β) (bc : β ≃ γ) (cd : γ ≃ δ) : + (ab.trans bc).trans cd = ab.trans (bc.trans cd) := Equiv.ext fun _ => rfl theorem left_inverse_symm (f : Equiv α β) : LeftInverse f.symm f := @@ -382,7 +383,8 @@ theorem equiv_congr_refl {α β} : (Equiv.refl α).equivCongr (Equiv.refl β) = rfl @[simp] -theorem equiv_congr_symm {δ} (ab : α ≃ β) (cd : γ ≃ δ) : (ab.equivCongr cd).symm = ab.symm.equivCongr cd.symm := by +theorem equiv_congr_symm {δ} (ab : α ≃ β) (cd : γ ≃ δ) : + (ab.equivCongr cd).symm = ab.symm.equivCongr cd.symm := by ext rfl @@ -393,11 +395,13 @@ theorem equiv_congr_trans {δ ε ζ} (ab : α ≃ β) (de : δ ≃ ε) (bc : β rfl @[simp] -theorem equiv_congr_refl_left {α β γ} (bg : β ≃ γ) (e : α ≃ β) : (Equiv.refl α).equivCongr bg e = e.trans bg := +theorem equiv_congr_refl_left {α β γ} (bg : β ≃ γ) (e : α ≃ β) : + (Equiv.refl α).equivCongr bg e = e.trans bg := rfl @[simp] -theorem equiv_congr_refl_right {α β} (ab e : α ≃ β) : ab.equivCongr (Equiv.refl β) e = ab.symm.trans e := +theorem equiv_congr_refl_right {α β} (ab e : α ≃ β) : + ab.equivCongr (Equiv.refl β) e = ab.symm.trans e := rfl @[simp] @@ -430,8 +434,8 @@ theorem perm_congr_apply (p : Equiv.Perm α') (x) : e.permCongr p x = e (p (e.sy theorem perm_congr_symm_apply (p : Equiv.Perm β') (x) : e.permCongr.symm p x = e.symm (p (e x)) := rfl -theorem perm_congr_trans (p p' : Equiv.Perm α') : (e.permCongr p).trans (e.permCongr p') = e.permCongr (p.trans p') := - by +theorem perm_congr_trans (p p' : Equiv.Perm α') : + (e.permCongr p).trans (e.permCongr p') = e.permCongr (p.trans p') := by ext simp @@ -503,17 +507,20 @@ def arrowCongr {α₁ β₁ α₂ β₂ : Sort _} (e₁ : α₁ ≃ α₂) (e₂ #align equiv.arrow_congr_apply Equiv.arrowCongr_apply -theorem arrow_congr_comp {α₁ β₁ γ₁ α₂ β₂ γ₂ : Sort _} (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) (ec : γ₁ ≃ γ₂) (f : α₁ → β₁) - (g : β₁ → γ₁) : arrowCongr ea ec (g ∘ f) = arrowCongr eb ec g ∘ arrowCongr ea eb f := by +theorem arrow_congr_comp {α₁ β₁ γ₁ α₂ β₂ γ₂ : Sort _} (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) (ec : γ₁ ≃ γ₂) + (f : α₁ → β₁) (g : β₁ → γ₁) : + arrowCongr ea ec (g ∘ f) = arrowCongr eb ec g ∘ arrowCongr ea eb f := by ext simp only [comp, arrowCongr_apply, eb.symm_apply_apply] @[simp] -theorem arrow_congr_refl {α β : Sort _} : arrowCongr (Equiv.refl α) (Equiv.refl β) = Equiv.refl (α → β) := +theorem arrow_congr_refl {α β : Sort _} : + arrowCongr (Equiv.refl α) (Equiv.refl β) = Equiv.refl (α → β) := rfl @[simp] -theorem arrow_congr_trans {α₁ β₁ α₂ β₂ α₃ β₃ : Sort _} (e₁ : α₁ ≃ α₂) (e₁' : β₁ ≃ β₂) (e₂ : α₂ ≃ α₃) (e₂' : β₂ ≃ β₃) : +theorem arrow_congr_trans {α₁ β₁ α₂ β₂ α₃ β₃ : Sort _} + (e₁ : α₁ ≃ α₂) (e₁' : β₁ ≃ β₂) (e₂ : α₂ ≃ α₃) (e₂' : β₂ ≃ β₃) : arrowCongr (e₁.trans e₂) (e₁'.trans e₂') = (arrowCongr e₁ e₁').trans (arrowCongr e₂ e₂') := rfl @@ -533,11 +540,13 @@ def arrowCongr' {α₁ β₁ α₂ β₂ : Type _} (hα : α₁ ≃ α₂) (hβ Equiv.arrowCongr hα hβ @[simp] -theorem arrow_congr'_refl {α β : Type _} : arrowCongr' (Equiv.refl α) (Equiv.refl β) = Equiv.refl (α → β) := +theorem arrow_congr'_refl {α β : Type _} : + arrowCongr' (Equiv.refl α) (Equiv.refl β) = Equiv.refl (α → β) := rfl @[simp] -theorem arrow_congr'_trans {α₁ β₁ α₂ β₂ α₃ β₃ : Type _} (e₁ : α₁ ≃ α₂) (e₁' : β₁ ≃ β₂) (e₂ : α₂ ≃ α₃) (e₂' : β₂ ≃ β₃) : +theorem arrow_congr'_trans {α₁ β₁ α₂ β₂ α₃ β₃ : Type _} + (e₁ : α₁ ≃ α₂) (e₁' : β₁ ≃ β₂) (e₂ : α₂ ≃ α₃) (e₂' : β₂ ≃ β₃) : arrowCongr' (e₁.trans e₂) (e₁'.trans e₂') = (arrowCongr' e₁ e₁').trans (arrowCongr' e₂ e₂') := rfl @@ -566,7 +575,8 @@ theorem conj_trans (e₁ : α ≃ β) (e₂ : β ≃ γ) : (e₁.trans e₂).con -- This should not be a simp lemma as long as `(∘)` is reducible: -- when `(∘)` is reducible, Lean can unify `f₁ ∘ f₂` with any `g` using -- `f₁ := g` and `f₂ := λ x, x`. This causes nontermination. -theorem conj_comp (e : α ≃ β) (f₁ f₂ : α → α) : e.conj (f₁ ∘ f₂) = e.conj f₁ ∘ e.conj f₂ := by apply arrow_congr_comp +theorem conj_comp (e : α ≃ β) (f₁ f₂ : α → α) : e.conj (f₁ ∘ f₂) = e.conj f₁ ∘ e.conj f₂ := by + apply arrow_congr_comp theorem eq_comp_symm {α β γ} (e : α ≃ β) (f : β → γ) (g : α → γ) : f = g ∘ e.symm ↔ f ∘ e = g := (e.arrowCongr (Equiv.refl γ)).symm_apply_eq.symm @@ -684,7 +694,8 @@ def psigmaCongrRight {α} {β₁ β₂ : α → Sort _} (F : ∀ a, β₁ a ≃ congr_arg (PSigma.mk a) <| apply_symm_apply (F a) b⟩ @[simp] -theorem psigma_congr_right_trans {α} {β₁ β₂ β₃ : α → Sort _} (F : ∀ a, β₁ a ≃ β₂ a) (G : ∀ a, β₂ a ≃ β₃ a) : +theorem psigma_congr_right_trans {α} {β₁ β₂ β₃ : α → Sort _} + (F : ∀ a, β₁ a ≃ β₂ a) (G : ∀ a, β₂ a ≃ β₃ a) : (psigmaCongrRight F).trans (psigmaCongrRight G) = psigmaCongrRight fun a => (F a).trans (G a) := by ext1 x cases x @@ -713,7 +724,8 @@ def sigmaCongrRight {α} {β₁ β₂ : α → Type _} (F : ∀ a, β₁ a ≃ congr_arg (Sigma.mk a) <| apply_symm_apply (F a) b⟩ @[simp] -theorem sigma_congr_right_trans {α} {β₁ β₂ β₃ : α → Type _} (F : ∀ a, β₁ a ≃ β₂ a) (G : ∀ a, β₂ a ≃ β₃ a) : +theorem sigma_congr_right_trans {α} {β₁ β₂ β₃ : α → Type _} + (F : ∀ a, β₁ a ≃ β₂ a) (G : ∀ a, β₂ a ≃ β₃ a) : (sigmaCongrRight F).trans (sigmaCongrRight G) = sigmaCongrRight fun a => (F a).trans (G a) := by ext1 x cases x @@ -746,12 +758,14 @@ def psigmaEquivSubtype {α : Type v} (P : α → Prop) : (Σ'i, P i) ≃ Subtype /-- A `sigma` with `plift` fibers is equivalent to the subtype. -/ def sigmaPliftEquivSubtype {α : Type v} (P : α → Prop) : (Σi, PLift (P i)) ≃ Subtype P := - ((psigmaEquivSigma _).symm.trans (psigmaCongrRight fun _ => Equiv.plift)).trans (psigmaEquivSubtype P) + ((psigmaEquivSigma _).symm.trans + (psigmaCongrRight fun _ => Equiv.plift)).trans (psigmaEquivSubtype P) /-- A `sigma` with `λ i, ulift (plift (P i))` fibers is equivalent to `{ x // P x }`. Variant of `sigma_plift_equiv_subtype`. -/ -def sigmaUliftPliftEquivSubtype {α : Type v} (P : α → Prop) : (Σi, ULift (PLift (P i))) ≃ Subtype P := +def sigmaUliftPliftEquivSubtype {α : Type v} (P : α → Prop) : + (Σi, ULift (PLift (P i))) ≃ Subtype P := (sigmaCongrRight fun _ => Equiv.ulift).trans (sigmaPliftEquivSubtype P) namespace Perm @@ -794,13 +808,14 @@ def sigmaCongrLeft {α₁ α₂} {β : α₂ → Sort _} (e : α₁ ≃ α₂) : | _, rfl => rfl /-- Transporting a sigma type through an equivalence of the base -/ -def sigmaCongrLeft' {α₁ α₂} {β : α₁ → Sort _} (f : α₁ ≃ α₂) : (Σa : α₁, β a) ≃ Σa : α₂, β (f.symm a) := +def sigmaCongrLeft' {α₁ α₂} {β : α₁ → Sort _} (f : α₁ ≃ α₂) : + (Σa : α₁, β a) ≃ Σa : α₂, β (f.symm a) := (sigmaCongrLeft f.symm).symm /-- Transporting a sigma type through an equivalence of the base and a family of equivalences of matching fibers -/ -def sigmaCongr {α₁ α₂} {β₁ : α₁ → Sort _} {β₂ : α₂ → Sort _} (f : α₁ ≃ α₂) (F : ∀ a, β₁ a ≃ β₂ (f a)) : - Sigma β₁ ≃ Sigma β₂ := +def sigmaCongr {α₁ α₂} {β₁ : α₁ → Sort _} {β₂ : α₂ → Sort _} (f : α₁ ≃ α₂) + (F : ∀ a, β₁ a ≃ β₂ (f a)) : Sigma β₁ ≃ Sigma β₂ := (sigmaCongrRight F).trans (sigmaCongrLeft f) /-- `sigma` type with a constant fiber is equivalent to the product. -/ @@ -823,8 +838,8 @@ def sigmaAssoc {α : Type _} {β : α → Type _} (γ : ∀ a : α, β a → Typ end -protected theorem exists_unique_congr {p : α → Prop} {q : β → Prop} (f : α ≃ β) (h : ∀ {x}, p x ↔ q (f x)) : - (∃! x, p x) ↔ ∃! y, q y := by +protected theorem exists_unique_congr {p : α → Prop} {q : β → Prop} + (f : α ≃ β) (h : ∀ {x}, p x ↔ q (f x)) : (∃! x, p x) ↔ ∃! y, q y := by constructor · rintro ⟨a, ha₁, ha₂⟩ exact ⟨f a, h.1 ha₁, fun b hb => f.symm_apply_eq.1 (ha₂ (f.symm b) (h.2 (by simpa using hb)))⟩ @@ -833,10 +848,12 @@ protected theorem exists_unique_congr {p : α → Prop} {q : β → Prop} (f : exact ⟨f.symm b, h.2 (by simpa using hb₁), fun y hy => (eq_symm_apply f).2 (hb₂ _ (h.1 hy))⟩ -protected theorem exists_unique_congr_left' {p : α → Prop} (f : α ≃ β) : (∃! x, p x) ↔ ∃! y, p (f.symm y) := +protected theorem exists_unique_congr_left' {p : α → Prop} (f : α ≃ β) : + (∃! x, p x) ↔ ∃! y, p (f.symm y) := Equiv.exists_unique_congr f fun {_} => by simp -protected theorem exists_unique_congr_left {p : β → Prop} (f : α ≃ β) : (∃! x, p (f x)) ↔ ∃! y, p y := +protected theorem exists_unique_congr_left {p : β → Prop} (f : α ≃ β) : + (∃! x, p (f x)) ↔ ∃! y, p y := (Equiv.exists_unique_congr_left' f.symm).symm protected theorem forall_congr {p : α → Prop} {q : β → Prop} (f : α ≃ β) @@ -848,8 +865,8 @@ by constructor <;> intro h₂ x apply h.mpr apply h₂ -protected theorem forall_congr' {p : α → Prop} {q : β → Prop} (f : α ≃ β) (h : ∀ {x}, p (f.symm x) ↔ q x) : - (∀ x, p x) ↔ ∀ y, q y := +protected theorem forall_congr' {p : α → Prop} {q : β → Prop} (f : α ≃ β) + (h : ∀ {x}, p (f.symm x) ↔ q x) : (∀ x, p x) ↔ ∀ y, q y := (Equiv.forall_congr f.symm fun {_} => h.symm).symm -- We next build some higher arity versions of `equiv.forall_congr`. @@ -860,30 +877,34 @@ protected theorem forall_congr' {p : α → Prop} {q : β → Prop} (f : α ≃ -- it's rare to have axioms involving more than 3 elements at once.) universe ua1 ua2 ub1 ub2 ug1 ug2 -variable {α₁ : Sort ua1} {α₂ : Sort ua2} {β₁ : Sort ub1} {β₂ : Sort ub2} {γ₁ : Sort ug1} {γ₂ : Sort ug2} +variable {α₁ : Sort ua1} {α₂ : Sort ua2} {β₁ : Sort ub1} {β₂ : Sort ub2} + {γ₁ : Sort ug1} {γ₂ : Sort ug2} -protected theorem forall₂_congr {p : α₁ → β₁ → Prop} {q : α₂ → β₂ → Prop} (eα : α₁ ≃ α₂) (eβ : β₁ ≃ β₂) - (h : ∀ {x y}, p x y ↔ q (eα x) (eβ y)) : (∀ x y, p x y) ↔ ∀ x y, q x y := by +protected theorem forall₂_congr {p : α₁ → β₁ → Prop} {q : α₂ → β₂ → Prop} (eα : α₁ ≃ α₂) + (eβ : β₁ ≃ β₂) (h : ∀ {x y}, p x y ↔ q (eα x) (eβ y)) : (∀ x y, p x y) ↔ ∀ x y, q x y := by apply Equiv.forall_congr intros apply Equiv.forall_congr intros apply h -protected theorem forall₂_congr' {p : α₁ → β₁ → Prop} {q : α₂ → β₂ → Prop} (eα : α₁ ≃ α₂) (eβ : β₁ ≃ β₂) - (h : ∀ {x y}, p (eα.symm x) (eβ.symm y) ↔ q x y) : (∀ x y, p x y) ↔ ∀ x y, q x y := +protected theorem forall₂_congr' {p : α₁ → β₁ → Prop} {q : α₂ → β₂ → Prop} + (eα : α₁ ≃ α₂) (eβ : β₁ ≃ β₂) (h : ∀ {x y}, p (eα.symm x) (eβ.symm y) ↔ q x y) : + (∀ x y, p x y) ↔ ∀ x y, q x y := (Equiv.forall₂_congr eα.symm eβ.symm fun {_ _} => h.symm).symm -protected theorem forall₃_congr {p : α₁ → β₁ → γ₁ → Prop} {q : α₂ → β₂ → γ₂ → Prop} (eα : α₁ ≃ α₂) (eβ : β₁ ≃ β₂) - (eγ : γ₁ ≃ γ₂) (h : ∀ {x y z}, p x y z ↔ q (eα x) (eβ y) (eγ z)) : (∀ x y z, p x y z) ↔ ∀ x y z, q x y z := by +protected theorem forall₃_congr {p : α₁ → β₁ → γ₁ → Prop} {q : α₂ → β₂ → γ₂ → Prop} + (eα : α₁ ≃ α₂) (eβ : β₁ ≃ β₂) (eγ : γ₁ ≃ γ₂) (h : ∀ {x y z}, p x y z ↔ q (eα x) (eβ y) (eγ z)) : + (∀ x y z, p x y z) ↔ ∀ x y z, q x y z := by apply Equiv.forall₂_congr intros apply Equiv.forall_congr intros apply h -protected theorem forall₃_congr' {p : α₁ → β₁ → γ₁ → Prop} {q : α₂ → β₂ → γ₂ → Prop} (eα : α₁ ≃ α₂) (eβ : β₁ ≃ β₂) - (eγ : γ₁ ≃ γ₂) (h : ∀ {x y z}, p (eα.symm x) (eβ.symm y) (eγ.symm z) ↔ q x y z) : +protected theorem forall₃_congr' {p : α₁ → β₁ → γ₁ → Prop} {q : α₂ → β₂ → γ₂ → Prop} + (eα : α₁ ≃ α₂) (eβ : β₁ ≃ β₂) (eγ : γ₁ ≃ γ₂) + (h : ∀ {x y z}, p (eα.symm x) (eβ.symm y) (eγ.symm z) ↔ q x y z) : (∀ x y z, p x y z) ↔ ∀ x y z, q x y z := (Equiv.forall₃_congr eα.symm eβ.symm eγ.symm fun {_ _ _} => h.symm).symm @@ -893,7 +914,8 @@ protected theorem forall_congr_left' {p : α → Prop} (f : α ≃ β) : (∀ x, protected theorem forall_congr_left {p : β → Prop} (f : α ≃ β) : (∀ x, p (f x)) ↔ ∀ y, p y := (Equiv.forall_congr_left' f.symm).symm -protected theorem exists_congr_left {α β} (f : α ≃ β) {p : α → Prop} : (∃ a, p a) ↔ ∃ b, p (f.symm b) := +protected theorem exists_congr_left {α β} (f : α ≃ β) {p : α → Prop} : + (∃ a, p a) ↔ ∃ b, p (f.symm b) := ⟨fun ⟨a, h⟩ => ⟨f a, by simpa using h⟩, fun ⟨b, h⟩ => ⟨_, h⟩⟩ end Equiv @@ -902,8 +924,8 @@ namespace Quot /-- An equivalence `e : α ≃ β` generates an equivalence between quotient spaces, if `ra a₁ a₂ ↔ rb (e a₁) (e a₂). -/ -protected def congr {ra : α → α → Prop} {rb : β → β → Prop} (e : α ≃ β) (eq : ∀ a₁ a₂, ra a₁ a₂ ↔ rb (e a₁) (e a₂)) : - Quot ra ≃ Quot rb where +protected def congr {ra : α → α → Prop} {rb : β → β → Prop} (e : α ≃ β) + (eq : ∀ a₁ a₂, ra a₁ a₂ ↔ rb (e a₁) (e a₂)) : Quot ra ≃ Quot rb where toFun := Quot.map e fun a₁ a₂ => (Eq a₁ a₂).1 invFun := Quot.map e.symm fun b₁ b₂ h => @@ -918,19 +940,23 @@ protected def congr {ra : α → α → Prop} {rb : β → β → Prop} (e : α simp only [Equiv.apply_symm_apply] @[simp] -theorem congr_mk {ra : α → α → Prop} {rb : β → β → Prop} (e : α ≃ β) (eq : ∀ a₁ a₂ : α, ra a₁ a₂ ↔ rb (e a₁) (e a₂)) - (a : α) : Quot.congr e Eq (Quot.mk ra a) = Quot.mk rb (e a) := +theorem congr_mk {ra : α → α → Prop} {rb : β → β → Prop} (e : α ≃ β) + (eq : ∀ a₁ a₂ : α, ra a₁ a₂ ↔ rb (e a₁) (e a₂)) (a : α) : + Quot.congr e Eq (Quot.mk ra a) = Quot.mk rb (e a) := rfl /-- Quotients are congruent on equivalences under equality of their relation. An alternative is just to use rewriting with `eq`, but then computational proofs get stuck. -/ -protected def congrRight {r r' : α → α → Prop} (eq : ∀ a₁ a₂, r a₁ a₂ ↔ r' a₁ a₂) : Quot r ≃ Quot r' := +protected def congrRight {r r' : α → α → Prop} (eq : ∀ a₁ a₂, r a₁ a₂ ↔ r' a₁ a₂) : + Quot r ≃ Quot r' := Quot.congr (Equiv.refl α) Eq /-- An equivalence `e : α ≃ β` generates an equivalence between the quotient space of `α` by a relation `ra` and the quotient space of `β` by the image of this relation under `e`. -/ -protected def congrLeft {r : α → α → Prop} (e : α ≃ β) : Quot r ≃ Quot fun b b' => r (e.symm b) (e.symm b') := - @Quot.congr α β r (fun b b' => r (e.symm b) (e.symm b')) e fun a₁ a₂ => by simp only [e.symm_apply_apply] +protected def congrLeft {r : α → α → Prop} (e : α ≃ β) : + Quot r ≃ Quot fun b b' => r (e.symm b) (e.symm b') := + @Quot.congr α β r (fun b b' => r (e.symm b) (e.symm b')) e + fun a₁ a₂ => by simp only [e.symm_apply_apply] end Quot @@ -939,18 +965,20 @@ namespace Quotient /-- An equivalence `e : α ≃ β` generates an equivalence between quotient spaces, if `ra a₁ a₂ ↔ rb (e a₁) (e a₂). -/ protected def congr {ra : Setoid α} {rb : Setoid β} (e : α ≃ β) - (eq : ∀ a₁ a₂, @Setoid.R α ra a₁ a₂ ↔ @Setoid.R β rb (e a₁) (e a₂)) : Quotient ra ≃ Quotient rb := + (eq : ∀ a₁ a₂, @Setoid.R α ra a₁ a₂ ↔ @Setoid.R β rb (e a₁) (e a₂)) : + Quotient ra ≃ Quotient rb := Quot.congr e Eq @[simp] -theorem congr_mk {ra : Setoid α} {rb : Setoid β} (e : α ≃ β) (eq : ∀ a₁ a₂ : α, Setoid.R a₁ a₂ ↔ Setoid.R (e a₁) (e a₂)) - (a : α) : Quotient.congr e Eq (Quotient.mk a) = Quotient.mk (e a) := +theorem congr_mk {ra : Setoid α} {rb : Setoid β} (e : α ≃ β) + (eq : ∀ a₁ a₂ : α, Setoid.R a₁ a₂ ↔ Setoid.R (e a₁) (e a₂)) (a : α) : + Quotient.congr e Eq (Quotient.mk a) = Quotient.mk (e a) := rfl /-- Quotients are congruent on equivalences under equality of their relation. An alternative is just to use rewriting with `eq`, but then computational proofs get stuck. -/ -protected def congrRight {r r' : Setoid α} (eq : ∀ a₁ a₂, @Setoid.R α r a₁ a₂ ↔ @Setoid.R α r' a₁ a₂) : - Quotient r ≃ Quotient r' := +protected def congrRight {r r' : Setoid α} + (eq : ∀ a₁ a₂, @Setoid.R α r a₁ a₂ ↔ @Setoid.R α r' a₁ a₂) : Quotient r ≃ Quotient r' := Quot.congrRight Eq end Quotient From 95ad70183ff293e68c016111447f68229703446c Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?David=20W=C3=A4rn?= Date: Sun, 13 Nov 2022 15:24:37 +0100 Subject: [PATCH 015/127] finish last part + capitalisation --- Mathlib/Logic/Equiv/Defs.lean | 291 +++++++++++++++++++++------------- 1 file changed, 184 insertions(+), 107 deletions(-) diff --git a/Mathlib/Logic/Equiv/Defs.lean b/Mathlib/Logic/Equiv/Defs.lean index 3270863fc7fef..9e031499e4c0d 100644 --- a/Mathlib/Logic/Equiv/Defs.lean +++ b/Mathlib/Logic/Equiv/Defs.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Mario Carneiro -/ import Mathlib.Data.FunLike.Equiv +import Mathlib.Data.Quot import Mathlib.Init.Data.Bool.Lemmas import Mathlib.Logic.Unique import Mathlib.Tactic.Conv @@ -77,7 +78,7 @@ instance {F} [EquivLike F α β] : CoeTC F (α ≃ β) := left_inv := EquivLike.left_inv f, right_inv := EquivLike.right_inv f }⟩ -/-- `perm α` is the type of bijections from `α` to itself. -/ +/-- `Perm α` is the type of bijections from `α` to itself. -/ @[reducible] def Equiv.Perm (α : Sort _) := Equiv α α @@ -197,7 +198,7 @@ instance permUnique [Subsingleton α] : Unique (Perm α) := theorem Perm.subsingleton_eq_refl [Subsingleton α] (e : Perm α) : e = Equiv.refl α := Subsingleton.elim _ _ -/-- Transfer `decidable_eq` across an equivalence. -/ +/-- Transfer `decidableEq` across an equivalence. -/ protected def decidableEq (e : α ≃ β) [DecidableEq β] : DecidableEq α := e.injective.decidableEq @@ -370,46 +371,61 @@ theorem comp_bijective (f : α → β) (e : β ≃ γ) : Bijective (e ∘ f) ↔ /-- If `α` is equivalent to `β` and `γ` is equivalent to `δ`, then the type of equivalences `α ≃ γ` is equivalent to the type of equivalences `β ≃ δ`. -/ -def equivCongr {δ} (ab : α ≃ β) (cd : γ ≃ δ) : α ≃ γ ≃ (β ≃ δ) := - ⟨fun ac => (ab.symm.trans ac).trans cd, fun bd => ab.trans <| bd.trans <| cd.symm, fun ac => by - ext x - simp, fun ac => by - ext x - simp⟩ +def equivCongr {δ} (ab : α ≃ β) (cd : γ ≃ δ) : (α ≃ γ) ≃ (β ≃ δ) := + ⟨ fun ac => (ab.symm.trans ac).trans cd, + fun bd => ab.trans <| bd.trans <| cd.symm, + fun ac => by + ext x + simp only [trans_apply, comp_apply, symm_apply_apply], + fun ac => by + ext x + simp only [trans_apply, comp_apply, apply_symm_apply] ⟩ @[simp] -theorem equiv_congr_refl {α β} : (Equiv.refl α).equivCongr (Equiv.refl β) = Equiv.refl (α ≃ β) := by +theorem equivCongr_refl {α β} : (Equiv.refl α).equivCongr (Equiv.refl β) = Equiv.refl (α ≃ β) := by ext rfl +#align equiv.equiv_congr_refl Equiv.equivCongr_refl + @[simp] -theorem equiv_congr_symm {δ} (ab : α ≃ β) (cd : γ ≃ δ) : +theorem equivCongr_symm {δ} (ab : α ≃ β) (cd : γ ≃ δ) : (ab.equivCongr cd).symm = ab.symm.equivCongr cd.symm := by ext rfl +#align equiv.equiv_congr_symm Equiv.equivCongr_symm + @[simp] -theorem equiv_congr_trans {δ ε ζ} (ab : α ≃ β) (de : δ ≃ ε) (bc : β ≃ γ) (ef : ε ≃ ζ) : +theorem equivCongr_trans {δ ε ζ} (ab : α ≃ β) (de : δ ≃ ε) (bc : β ≃ γ) (ef : ε ≃ ζ) : (ab.equivCongr de).trans (bc.equivCongr ef) = (ab.trans bc).equivCongr (de.trans ef) := by ext rfl +#align equiv.equiv_congr_trans Equiv.equivCongr_trans + @[simp] -theorem equiv_congr_refl_left {α β γ} (bg : β ≃ γ) (e : α ≃ β) : +theorem equivCongr_refl_left {α β γ} (bg : β ≃ γ) (e : α ≃ β) : (Equiv.refl α).equivCongr bg e = e.trans bg := rfl +#align equiv.equiv_congr_refl_left Equiv.equivCongr_refl_left + @[simp] -theorem equiv_congr_refl_right {α β} (ab e : α ≃ β) : +theorem equivCongr_refl_right {α β} (ab e : α ≃ β) : ab.equivCongr (Equiv.refl β) e = ab.symm.trans e := rfl +#align equiv.equiv_congr_refl_right Equiv.equivCongr_refl_right + @[simp] -theorem equiv_congr_apply_apply {δ} (ab : α ≃ β) (cd : γ ≃ δ) (e : α ≃ γ) (x) : +theorem equivCongr_apply_apply {δ} (ab : α ≃ β) (cd : γ ≃ δ) (e : α ≃ γ) (x) : ab.equivCongr cd e x = cd (e (ab.symm x)) := rfl -section PermCongr +#align equiv.equiv_congr_apply_apply Equiv.equivCongr_apply_apply + +section permCongr variable {α' β' : Type _} (e : α' ≃ β') @@ -417,39 +433,51 @@ variable {α' β' : Type _} (e : α' ≃ β') def permCongr : Perm α' ≃ Perm β' := equivCongr e e -theorem perm_congr_def (p : Equiv.Perm α') : e.permCongr p = (e.symm.trans p).trans e := +theorem permCongr_def (p : Equiv.Perm α') : e.permCongr p = (e.symm.trans p).trans e := rfl +#align equiv.perm_congr_def Equiv.permCongr_def + @[simp] -theorem perm_congr_refl : e.permCongr (Equiv.refl _) = Equiv.refl _ := by simp [perm_congr_def] +theorem permCongr_refl : e.permCongr (Equiv.refl _) = Equiv.refl _ := by simp [permCongr_def] + +#align equiv.perm_congr_refl Equiv.permCongr_refl @[simp] -theorem perm_congr_symm : e.permCongr.symm = e.symm.permCongr := +theorem permCongr_symm : e.permCongr.symm = e.symm.permCongr := rfl +#align equiv.perm_congr_symm Equiv.permCongr_symm + @[simp] -theorem perm_congr_apply (p : Equiv.Perm α') (x) : e.permCongr p x = e (p (e.symm x)) := +theorem permCongr_apply (p : Equiv.Perm α') (x) : e.permCongr p x = e (p (e.symm x)) := rfl -theorem perm_congr_symm_apply (p : Equiv.Perm β') (x) : e.permCongr.symm p x = e.symm (p (e x)) := +#align equiv.perm_congr_apply Equiv.permCongr_apply + +theorem permCongr_symm_apply (p : Equiv.Perm β') (x) : e.permCongr.symm p x = e.symm (p (e x)) := rfl -theorem perm_congr_trans (p p' : Equiv.Perm α') : +#align equiv.perm_congr_symm_apply Equiv.permCongr_symm_apply + +theorem permCongr_trans (p p' : Equiv.Perm α') : (e.permCongr p).trans (e.permCongr p') = e.permCongr (p.trans p') := by ext - simp + simp only [trans_apply, comp_apply, permCongr_apply, symm_apply_apply] + +#align equiv.perm_congr_trans Equiv.permCongr_trans -end PermCongr +end permCongr /-- Two empty types are equivalent. -/ def equivOfIsEmpty (α β : Sort _) [IsEmpty α] [IsEmpty β] : α ≃ β := ⟨isEmptyElim, isEmptyElim, isEmptyElim, isEmptyElim⟩ -/-- If `α` is an empty type, then it is equivalent to the `empty` type. -/ +/-- If `α` is an empty type, then it is equivalent to the `Empty` type. -/ def equivEmpty (α : Sort u) [IsEmpty α] : α ≃ Empty := equivOfIsEmpty α _ -/-- If `α` is an empty type, then it is equivalent to the `pempty` type in any universe. -/ +/-- If `α` is an empty type, then it is equivalent to the `PEmpty` type in any universe. -/ def equivPempty (α : Sort v) [IsEmpty α] : α ≃ PEmpty.{u} := equivOfIsEmpty α _ @@ -457,7 +485,7 @@ def equivPempty (α : Sort v) [IsEmpty α] : α ≃ PEmpty.{u} := def equivEmptyEquiv (α : Sort u) : α ≃ Empty ≃ IsEmpty α := ⟨fun e => Function.is_empty e, @equivEmpty α, fun e => ext fun x => (e x).elim, fun _ => rfl⟩ -/-- The `Sort` of proofs of a false proposition is equivalent to `pempty`. -/ +/-- The `Sort` of proofs of a false proposition is equivalent to `PEmpty`. -/ def propEquivPempty {p : Prop} (h : ¬p) : p ≃ PEmpty := @equivPempty p <| IsEmpty.prop_iff.2 h @@ -468,22 +496,22 @@ def equivOfUnique (α β : Sort _) [Unique.{u} α] [Unique.{v} β] : α ≃ β w left_inv _ := Subsingleton.elim _ _ right_inv _ := Subsingleton.elim _ _ -/-- If `α` has a unique element, then it is equivalent to any `punit`. -/ +/-- If `α` has a unique element, then it is equivalent to any `PUnit`. -/ def equivPunit (α : Sort u) [Unique α] : α ≃ PUnit.{v} := equivOfUnique α _ -/-- The `Sort` of proofs of a true proposition is equivalent to `punit`. -/ +/-- The `Sort` of proofs of a true proposition is equivalent to `PUnit`. -/ def propEquivPunit {p : Prop} (h : p) : p ≃ PUnit.{0} := -- porting note: the following causes an error "unknown free variable [anonymous]" -- @equivPunit p <| uniqueProp h ⟨default, λ _ => h, λ _ => Subsingleton.elim _ _, λ _ => Subsingleton.elim _ _⟩ -/-- `ulift α` is equivalent to `α`. -/ +/-- `ULift α` is equivalent to `α`. -/ @[simps (config := { fullyApplied := false }) apply symmApply] protected def ulift {α : Type v} : ULift.{u} α ≃ α := ⟨ULift.down, ULift.up, ULift.up_down, fun _ => rfl⟩ -/-- `plift α` is equivalent to `α`. -/ +/-- `PLift α` is equivalent to `α`. -/ @[simps (config := { fullyApplied := false }) apply symmApply] protected def plift : PLift α ≃ α := ⟨PLift.down, PLift.up, PLift.up_down, PLift.down_up⟩ @@ -502,36 +530,44 @@ is equivalent to the type of maps `α₂ → β₂`. -/ def arrowCongr {α₁ β₁ α₂ β₂ : Sort _} (e₁ : α₁ ≃ α₂) (e₂ : β₁ ≃ β₂) : (α₁ → β₁) ≃ (α₂ → β₂) where toFun f := e₂ ∘ f ∘ e₁.symm invFun f := e₂.symm ∘ f ∘ e₁ - left_inv f := funext fun x => by simp - right_inv f := funext fun x => by simp + left_inv f := funext fun x => by simp only [comp_apply, symm_apply_apply] + right_inv f := funext fun x => by simp only [comp_apply, apply_symm_apply] #align equiv.arrow_congr_apply Equiv.arrowCongr_apply -theorem arrow_congr_comp {α₁ β₁ γ₁ α₂ β₂ γ₂ : Sort _} (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) (ec : γ₁ ≃ γ₂) +theorem arrowCongr_comp {α₁ β₁ γ₁ α₂ β₂ γ₂ : Sort _} (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) (ec : γ₁ ≃ γ₂) (f : α₁ → β₁) (g : β₁ → γ₁) : arrowCongr ea ec (g ∘ f) = arrowCongr eb ec g ∘ arrowCongr ea eb f := by ext simp only [comp, arrowCongr_apply, eb.symm_apply_apply] +#align equiv.arrow_congr_comp Equiv.arrowCongr_comp + @[simp] -theorem arrow_congr_refl {α β : Sort _} : +theorem arrowCongr_refl {α β : Sort _} : arrowCongr (Equiv.refl α) (Equiv.refl β) = Equiv.refl (α → β) := rfl +#align equiv.arrow_congr_refl Equiv.arrowCongr_refl + @[simp] -theorem arrow_congr_trans {α₁ β₁ α₂ β₂ α₃ β₃ : Sort _} +theorem arrowCongr_trans {α₁ β₁ α₂ β₂ α₃ β₃ : Sort _} (e₁ : α₁ ≃ α₂) (e₁' : β₁ ≃ β₂) (e₂ : α₂ ≃ α₃) (e₂' : β₂ ≃ β₃) : arrowCongr (e₁.trans e₂) (e₁'.trans e₂') = (arrowCongr e₁ e₁').trans (arrowCongr e₂ e₂') := rfl +#align equiv.arrow_congr_trans Equiv.arrowCongr_trans + @[simp] -theorem arrow_congr_symm {α₁ β₁ α₂ β₂ : Sort _} (e₁ : α₁ ≃ α₂) (e₂ : β₁ ≃ β₂) : +theorem arrowCongr_symm {α₁ β₁ α₂ β₂ : Sort _} (e₁ : α₁ ≃ α₂) (e₂ : β₁ ≃ β₂) : (arrowCongr e₁ e₂).symm = arrowCongr e₁.symm e₂.symm := rfl -/-- A version of `equiv.arrow_congr` in `Type`, rather than `Sort`. +#align equiv.arrow_congr_symm Equiv.arrowCongr_symm -The `equiv_rw` tactic is not able to use the default `Sort` level `equiv.arrow_congr`, +/-- A version of `Equiv.arrowCongr` in `Type`, rather than `Sort`. + +The `equiv_rw` tactic is not able to use the default `Sort` level `Equiv.arrowCongr`, because Lean's universe rules will not unify `?l_1` with `imax (1 ?m_1)`. -/ -- porting note: removing `congr` attribute @@ -540,21 +576,27 @@ def arrowCongr' {α₁ β₁ α₂ β₂ : Type _} (hα : α₁ ≃ α₂) (hβ Equiv.arrowCongr hα hβ @[simp] -theorem arrow_congr'_refl {α β : Type _} : +theorem arrowCongr'_refl {α β : Type _} : arrowCongr' (Equiv.refl α) (Equiv.refl β) = Equiv.refl (α → β) := rfl +#align equiv.arrow_congr'_refl Equiv.arrowCongr'_refl + @[simp] -theorem arrow_congr'_trans {α₁ β₁ α₂ β₂ α₃ β₃ : Type _} +theorem arrowCongr'_trans {α₁ β₁ α₂ β₂ α₃ β₃ : Type _} (e₁ : α₁ ≃ α₂) (e₁' : β₁ ≃ β₂) (e₂ : α₂ ≃ α₃) (e₂' : β₂ ≃ β₃) : arrowCongr' (e₁.trans e₂) (e₁'.trans e₂') = (arrowCongr' e₁ e₁').trans (arrowCongr' e₂ e₂') := rfl +#align equiv.arrow_congr'_trans Equiv.arrowCongr'_trans + @[simp] -theorem arrow_congr'_symm {α₁ β₁ α₂ β₂ : Type _} (e₁ : α₁ ≃ α₂) (e₂ : β₁ ≃ β₂) : +theorem arrowCongr'_symm {α₁ β₁ α₂ β₂ : Type _} (e₁ : α₁ ≃ α₂) (e₂ : β₁ ≃ β₂) : (arrowCongr' e₁ e₂).symm = arrowCongr' e₁.symm e₂.symm := rfl +#align equiv.arrow_congr'_symm Equiv.arrowCongr'_symm + /-- Conjugate a map `f : α → α` by an equivalence `α ≃ β`. -/ @[simps apply] def conj (e : α ≃ β) : (α → α) ≃ (β → β) := @@ -576,7 +618,7 @@ theorem conj_trans (e₁ : α ≃ β) (e₂ : β ≃ γ) : (e₁.trans e₂).con -- when `(∘)` is reducible, Lean can unify `f₁ ∘ f₂` with any `g` using -- `f₁ := g` and `f₂ := λ x, x`. This causes nontermination. theorem conj_comp (e : α ≃ β) (f₁ f₂ : α → α) : e.conj (f₁ ∘ f₂) = e.conj f₁ ∘ e.conj f₂ := by - apply arrow_congr_comp + apply arrowCongr_comp theorem eq_comp_symm {α β γ} (e : α ≃ β) (f : β → γ) (g : α → γ) : f = g ∘ e.symm ↔ f ∘ e = g := (e.arrowCongr (Equiv.refl γ)).symm_apply_eq.symm @@ -590,15 +632,17 @@ theorem eq_symm_comp {α β γ} (e : α ≃ β) (f : γ → α) (g : γ → β) theorem symm_comp_eq {α β γ} (e : α ≃ β) (f : γ → α) (g : γ → β) : e.symm ∘ g = f ↔ g = e ∘ f := ((Equiv.refl γ).arrowCongr e).symm_apply_eq -/-- `punit` sorts in any two universes are equivalent. -/ -def punitEquivPunit : PUnit.{v} ≃ PUnit.{w} := +/-- `PUnit` sorts in any two universes are equivalent. -/ +def punitEquivPUnit : PUnit.{v} ≃ PUnit.{w} := ⟨fun _ => PUnit.unit, fun _ => PUnit.unit, fun u => by cases u rfl, fun u => by cases u rfl⟩ -/-- `Prop` is noncomputably equivalent to `bool`. -/ +#align equiv.punit_equiv_punit Equiv.punitEquivPUnit + +/-- `Prop` is noncomputably equivalent to `Bool`. -/ noncomputable def propEquivBool : Prop ≃ Bool where toFun p := @decide p (Classical.propDecidable _) invFun b := b @@ -607,8 +651,8 @@ noncomputable def propEquivBool : Prop ≃ Bool where section -/-- The sort of maps to `punit.{v}` is equivalent to `punit.{w}`. -/ -def arrowPunitEquivPunit (α : Sort _) : (α → PUnit.{v}) ≃ PUnit.{w} where +/-- The sort of maps to `PUnit.{v}` is equivalent to `PUnit.{w}`. -/ +def arrowPUnitEquivPUnit (α : Sort _) : (α → PUnit.{v}) ≃ PUnit.{w} where toFun _ := PUnit.unit invFun _ _ := PUnit.unit left_inv f := by @@ -619,8 +663,10 @@ def arrowPunitEquivPunit (α : Sort _) : (α → PUnit.{v}) ≃ PUnit.{w} where cases u rfl -/-- If `α` is `subsingleton` and `a : α`, then the type of dependent functions `Π (i : α), β -i` is equivalent to `β i`. -/ +#align equiv.arrow_punit_equiv_punit Equiv.arrowPUnitEquivPUnit + +/-- If `α` is `Subsingleton` and `a : α`, then the type of dependent functions `Π (i : α), β i` +is equivalent to `β a`. -/ @[simps] def piSubsingleton {α} (β : α → Sort _) [Subsingleton α] (a : α) : (∀ a', β a') ≃ β a where toFun := eval a @@ -636,16 +682,16 @@ def piSubsingleton {α} (β : α → Sort _) [Subsingleton α] (a : α) : (∀ a def funUnique (α β) [Unique.{u} α] : (α → β) ≃ β := piSubsingleton _ default -/-- The sort of maps from `punit` is equivalent to the codomain. -/ +/-- The sort of maps from `PUnit` is equivalent to the codomain. -/ def punitArrowEquiv (α : Sort _) : (PUnit.{u} → α) ≃ α := funUnique PUnit.{u} α -/-- The sort of maps from `true` is equivalent to the codomain. -/ +/-- The sort of maps from `True` is equivalent to the codomain. -/ def trueArrowEquiv (α : Sort _) : (True → α) ≃ α := funUnique _ _ -/-- The sort of maps from a type that `is_empty` is equivalent to `punit`. -/ -def arrowPunitOfIsEmpty (α β : Sort _) [IsEmpty α] : (α → β) ≃ PUnit.{u} where +/-- The sort of maps from a type that `IsEmpty` is equivalent to `PUnit`. -/ +def arrowPUnitOfIsEmpty (α β : Sort _) [IsEmpty α] : (α → β) ≃ PUnit.{u} where toFun _ := PUnit.unit invFun _ := isEmptyElim left_inv _ := funext isEmptyElim @@ -653,25 +699,33 @@ def arrowPunitOfIsEmpty (α β : Sort _) [IsEmpty α] : (α → β) ≃ PUnit.{u cases u rfl -/-- The sort of maps from `empty` is equivalent to `punit`. -/ -def emptyArrowEquivPunit (α : Sort _) : (Empty → α) ≃ PUnit.{u} := - arrowPunitOfIsEmpty _ _ +#align equiv.arrow_punit_of_is_empy Equiv.arrowPUnitOfIsEmpty + +/-- The sort of maps from `Empty` is equivalent to `PUnit`. -/ +def emptyArrowEquivPUnit (α : Sort _) : (Empty → α) ≃ PUnit.{u} := + arrowPUnitOfIsEmpty _ _ + +#align equiv.empty_arrow_equiv_punit Equiv.emptyArrowEquivPUnit + +/-- The sort of maps from `PEmpty` is equivalent to `PUnit`. -/ +def pemptyArrowEquivPUnit (α : Sort _) : (PEmpty → α) ≃ PUnit.{u} := + arrowPUnitOfIsEmpty _ _ -/-- The sort of maps from `pempty` is equivalent to `punit`. -/ -def pemptyArrowEquivPunit (α : Sort _) : (PEmpty → α) ≃ PUnit.{u} := - arrowPunitOfIsEmpty _ _ +#align equiv.pempty_arrow_equiv_punit Equiv.pemptyArrowEquivPUnit -/-- The sort of maps from `false` is equivalent to `punit`. -/ -def falseArrowEquivPunit (α : Sort _) : (False → α) ≃ PUnit.{u} := - arrowPunitOfIsEmpty _ _ +/-- The sort of maps from `False` is equivalent to `PUnit`. -/ +def falseArrowEquivPUnit (α : Sort _) : (False → α) ≃ PUnit.{u} := + arrowPUnitOfIsEmpty _ _ + +#align equiv.false_arrow_equiv_punit Equiv.falseArrowEquivPUnit end section -/-- A `psigma`-type is equivalent to the corresponding `sigma`-type. -/ +/-- A `PSigma`-type is equivalent to the corresponding `Sigma`-type. -/ @[simps apply symmApply] -def psigmaEquivSigma {α} (β : α → Type _) : (Σ'i, β i) ≃ Σi, β i where +def psigmaEquivSigma {α} (β : α → Type _) : (Σ' i, β i) ≃ Σ i, β i where toFun a := ⟨a.1, a.2⟩ invFun a := ⟨a.1, a.2⟩ left_inv := λ ⟨_, _⟩ => rfl @@ -679,74 +733,88 @@ def psigmaEquivSigma {α} (β : α → Type _) : (Σ'i, β i) ≃ Σi, β i wher /-- A `psigma`-type is equivalent to the corresponding `sigma`-type. -/ @[simps apply symmApply] -def psigmaEquivSigmaPlift {α} (β : α → Sort _) : (Σ'i, β i) ≃ Σi : PLift α, PLift (β i.down) where +def psigmaEquivSigmaPLift {α} (β : α → Sort _) : (Σ' i, β i) ≃ Σ i : PLift α, PLift (β i.down) where toFun a := ⟨PLift.up a.1, PLift.up a.2⟩ invFun a := ⟨a.1.down, a.2.down⟩ left_inv := λ ⟨_, _⟩ => rfl right_inv := λ ⟨_, _⟩ => rfl +#align equiv.psigma_equiv_sigma_plift Equiv.psigmaEquivSigmaPLift + /-- A family of equivalences `Π a, β₁ a ≃ β₂ a` generates an equivalence between `Σ' a, β₁ a` and `Σ' a, β₂ a`. -/ @[simps apply] -def psigmaCongrRight {α} {β₁ β₂ : α → Sort _} (F : ∀ a, β₁ a ≃ β₂ a) : (Σ'a, β₁ a) ≃ Σ'a, β₂ a := +def psigmaCongrRight {α} {β₁ β₂ : α → Sort _} (F : ∀ a, β₁ a ≃ β₂ a) : (Σ' a, β₁ a) ≃ Σ' a, β₂ a := ⟨fun a => ⟨a.1, F a.1 a.2⟩, fun a => ⟨a.1, (F a.1).symm a.2⟩, fun ⟨a, b⟩ => congr_arg (PSigma.mk a) <| symm_apply_apply (F a) b, fun ⟨a, b⟩ => congr_arg (PSigma.mk a) <| apply_symm_apply (F a) b⟩ @[simp] -theorem psigma_congr_right_trans {α} {β₁ β₂ β₃ : α → Sort _} +theorem psigmaCongrRight_trans {α} {β₁ β₂ β₃ : α → Sort _} (F : ∀ a, β₁ a ≃ β₂ a) (G : ∀ a, β₂ a ≃ β₃ a) : (psigmaCongrRight F).trans (psigmaCongrRight G) = psigmaCongrRight fun a => (F a).trans (G a) := by ext1 x cases x rfl +#align equiv.psigma_congr_right_trans Equiv.psigmaCongrRight_trans + @[simp] -theorem psigma_congr_right_symm {α} {β₁ β₂ : α → Sort _} (F : ∀ a, β₁ a ≃ β₂ a) : +theorem psigmaCongrRight_symm {α} {β₁ β₂ : α → Sort _} (F : ∀ a, β₁ a ≃ β₂ a) : (psigmaCongrRight F).symm = psigmaCongrRight fun a => (F a).symm := by ext1 x cases x rfl +#align equiv.psigma_congr_right_symm Equiv.psigmaCongrRight_symm + @[simp] -theorem psigma_congr_right_refl {α} {β : α → Sort _} : - (psigmaCongrRight fun a => Equiv.refl (β a)) = Equiv.refl (Σ'a, β a) := by +theorem psigmaCongrRight_refl {α} {β : α → Sort _} : + (psigmaCongrRight fun a => Equiv.refl (β a)) = Equiv.refl (Σ' a, β a) := by ext1 x cases x rfl +#align equiv.psigma_congr_right_refl Equiv.psigmaCongrRight_refl + /-- A family of equivalences `Π a, β₁ a ≃ β₂ a` generates an equivalence between `Σ a, β₁ a` and `Σ a, β₂ a`. -/ @[simps apply] -def sigmaCongrRight {α} {β₁ β₂ : α → Type _} (F : ∀ a, β₁ a ≃ β₂ a) : (Σa, β₁ a) ≃ Σa, β₂ a := +def sigmaCongrRight {α} {β₁ β₂ : α → Type _} (F : ∀ a, β₁ a ≃ β₂ a) : (Σ a, β₁ a) ≃ Σ a, β₂ a := ⟨fun a => ⟨a.1, F a.1 a.2⟩, fun a => ⟨a.1, (F a.1).symm a.2⟩, fun ⟨a, b⟩ => congr_arg (Sigma.mk a) <| symm_apply_apply (F a) b, fun ⟨a, b⟩ => congr_arg (Sigma.mk a) <| apply_symm_apply (F a) b⟩ @[simp] -theorem sigma_congr_right_trans {α} {β₁ β₂ β₃ : α → Type _} +theorem sigmaCongrRight_trans {α} {β₁ β₂ β₃ : α → Type _} (F : ∀ a, β₁ a ≃ β₂ a) (G : ∀ a, β₂ a ≃ β₃ a) : (sigmaCongrRight F).trans (sigmaCongrRight G) = sigmaCongrRight fun a => (F a).trans (G a) := by ext1 x cases x rfl +#align equiv.sigmaCongrRight Equiv.sigmaCongrRight_trans + @[simp] -theorem sigma_congr_right_symm {α} {β₁ β₂ : α → Type _} (F : ∀ a, β₁ a ≃ β₂ a) : +theorem sigmaCongrRight_symm {α} {β₁ β₂ : α → Type _} (F : ∀ a, β₁ a ≃ β₂ a) : (sigmaCongrRight F).symm = sigmaCongrRight fun a => (F a).symm := by ext1 x cases x rfl +#align equiv.sigma_congr_right_symm Equiv.sigmaCongrRight_symm + @[simp] -theorem sigma_congr_right_refl {α} {β : α → Type _} : - (sigmaCongrRight fun a => Equiv.refl (β a)) = Equiv.refl (Σa, β a) := by +theorem sigmaCongrRight_refl {α} {β : α → Type _} : + (sigmaCongrRight fun a => Equiv.refl (β a)) = Equiv.refl (Σ a, β a) := by ext1 x cases x rfl +#align equiv.sigma_congr_right_refl Equiv.sigmaCongrRight_refl + /-- A `psigma` with `Prop` fibers is equivalent to the subtype. -/ -def psigmaEquivSubtype {α : Type v} (P : α → Prop) : (Σ'i, P i) ≃ Subtype P where +def psigmaEquivSubtype {α : Type v} (P : α → Prop) : (Σ' i, P i) ≃ Subtype P where toFun x := ⟨x.1, x.2⟩ invFun x := ⟨x.1, x.2⟩ left_inv x := by @@ -757,44 +825,55 @@ def psigmaEquivSubtype {α : Type v} (P : α → Prop) : (Σ'i, P i) ≃ Subtype rfl /-- A `sigma` with `plift` fibers is equivalent to the subtype. -/ -def sigmaPliftEquivSubtype {α : Type v} (P : α → Prop) : (Σi, PLift (P i)) ≃ Subtype P := +def sigmaPLiftEquivSubtype {α : Type v} (P : α → Prop) : (Σ i, PLift (P i)) ≃ Subtype P := ((psigmaEquivSigma _).symm.trans (psigmaCongrRight fun _ => Equiv.plift)).trans (psigmaEquivSubtype P) +#align equiv.sigma_plift_equiv_subtype Equiv.sigmaPLiftEquivSubtype + /-- A `sigma` with `λ i, ulift (plift (P i))` fibers is equivalent to `{ x // P x }`. Variant of `sigma_plift_equiv_subtype`. -/ -def sigmaUliftPliftEquivSubtype {α : Type v} (P : α → Prop) : - (Σi, ULift (PLift (P i))) ≃ Subtype P := - (sigmaCongrRight fun _ => Equiv.ulift).trans (sigmaPliftEquivSubtype P) +def sigmaULiftPLiftEquivSubtype {α : Type v} (P : α → Prop) : + (Σ i, ULift (PLift (P i))) ≃ Subtype P := + (sigmaCongrRight fun _ => Equiv.ulift).trans (sigmaPLiftEquivSubtype P) + +#align equiv.sigma_ulift_plift_equiv_subtype Equiv.sigmaULiftPLiftEquivSubtype namespace Perm /-- A family of permutations `Π a, perm (β a)` generates a permuation `perm (Σ a, β₁ a)`. -/ @[reducible] -def sigmaCongrRight {α} {β : α → Sort _} (F : ∀ a, Perm (β a)) : Perm (Σa, β a) := +def sigmaCongrRight {α} {β : α → Sort _} (F : ∀ a, Perm (β a)) : Perm (Σ a, β a) := Equiv.sigmaCongrRight F @[simp] -theorem sigma_congr_right_trans {α} {β : α → Sort _} (F : ∀ a, Perm (β a)) (G : ∀ a, Perm (β a)) : +theorem sigmaCongrRight_trans {α} {β : α → Sort _} (F : ∀ a, Perm (β a)) (G : ∀ a, Perm (β a)) : (sigmaCongrRight F).trans (sigmaCongrRight G) = sigmaCongrRight fun a => (F a).trans (G a) := - Equiv.sigma_congr_right_trans F G + Equiv.sigmaCongrRight_trans F G + +#align equiv.perm.sigma_congr_right_trans Equiv.Perm.sigmaCongrRight_trans @[simp] -theorem sigma_congr_right_symm {α} {β : α → Sort _} (F : ∀ a, Perm (β a)) : +theorem sigmaCongrRight_symm {α} {β : α → Sort _} (F : ∀ a, Perm (β a)) : (sigmaCongrRight F).symm = sigmaCongrRight fun a => (F a).symm := - Equiv.sigma_congr_right_symm F + Equiv.sigmaCongrRight_symm F + +#align equiv.perm.sigma_congr_right_symm Equiv.Perm.sigmaCongrRight_symm @[simp] -theorem sigma_congr_right_refl {α} {β : α → Sort _} : - (sigmaCongrRight fun a => Equiv.refl (β a)) = Equiv.refl (Σa, β a) := - Equiv.sigma_congr_right_refl +theorem sigmaCongrRight_refl {α} {β : α → Sort _} : + (sigmaCongrRight fun a => Equiv.refl (β a)) = Equiv.refl (Σ a, β a) := + Equiv.sigmaCongrRight_refl + +#align equiv.perm.sigma_congr_right_refl Equiv.Perm.sigmaCongrRight_refl end Perm /-- An equivalence `f : α₁ ≃ α₂` generates an equivalence between `Σ a, β (f a)` and `Σ a, β a`. -/ @[simps apply] -def sigmaCongrLeft {α₁ α₂} {β : α₂ → Sort _} (e : α₁ ≃ α₂) : (Σa : α₁, β (e a)) ≃ Σa : α₂, β a where +def sigmaCongrLeft {α₁ α₂} {β : α₂ → Sort _} (e : α₁ ≃ α₂) : + (Σ a : α₁, β (e a)) ≃ Σ a : α₂, β a where toFun a := ⟨e a.1, a.2⟩ invFun a := ⟨e.symm a.1, (e.right_inv a.1).symm ▸ a.2⟩ -- porting note: this was a pretty gnarly match already, and it got worse after porting @@ -809,7 +888,7 @@ def sigmaCongrLeft {α₁ α₂} {β : α₂ → Sort _} (e : α₁ ≃ α₂) : /-- Transporting a sigma type through an equivalence of the base -/ def sigmaCongrLeft' {α₁ α₂} {β : α₁ → Sort _} (f : α₁ ≃ α₂) : - (Σa : α₁, β a) ≃ Σa : α₂, β (f.symm a) := + (Σ a : α₁, β a) ≃ Σ a : α₂, β (f.symm a) := (sigmaCongrLeft f.symm).symm /-- Transporting a sigma type through an equivalence of the base and a family of equivalences @@ -820,7 +899,7 @@ def sigmaCongr {α₁ α₂} {β₁ : α₁ → Sort _} {β₂ : α₂ → Sort /-- `sigma` type with a constant fiber is equivalent to the product. -/ @[simps apply symmApply] -def sigmaEquivProd (α β : Type _) : (Σ_ : α, β) ≃ α × β := +def sigmaEquivProd (α β : Type _) : (Σ _ : α, β) ≃ α × β := ⟨fun a => ⟨a.1, a.2⟩, fun a => ⟨a.1, a.2⟩, fun ⟨_, _⟩ => rfl, fun ⟨_, _⟩ => rfl⟩ /-- If each fiber of a `sigma` type is equivalent to a fixed type, then the sigma type @@ -830,7 +909,7 @@ def sigmaEquivProdOfEquiv {α β} {β₁ : α → Sort _} (F : ∀ a, β₁ a /-- Dependent product of types is associative up to an equivalence. -/ def sigmaAssoc {α : Type _} {β : α → Type _} (γ : ∀ a : α, β a → Type _) : - (Σab : Σa : α, β a, γ ab.1 ab.2) ≃ Σa : α, Σb : β a, γ a b where + (Σ ab : Σ a : α, β a, γ ab.1 ab.2) ≃ Σ a : α, Σ b : β a, γ a b where toFun x := ⟨x.1.1, ⟨x.1.2, x.2⟩⟩ invFun x := ⟨⟨x.1, x.2.1⟩, x.2.2⟩ left_inv := fun ⟨⟨_, _⟩, _⟩ => rfl @@ -926,37 +1005,35 @@ namespace Quot if `ra a₁ a₂ ↔ rb (e a₁) (e a₂). -/ protected def congr {ra : α → α → Prop} {rb : β → β → Prop} (e : α ≃ β) (eq : ∀ a₁ a₂, ra a₁ a₂ ↔ rb (e a₁) (e a₂)) : Quot ra ≃ Quot rb where - toFun := Quot.map e fun a₁ a₂ => (Eq a₁ a₂).1 + toFun := Quot.map e fun a₁ a₂ => (eq a₁ a₂).1 invFun := Quot.map e.symm fun b₁ b₂ h => - (Eq (e.symm b₁) (e.symm b₂)).2 ((e.apply_symm_apply b₁).symm ▸ (e.apply_symm_apply b₂).symm ▸ h) + (eq (e.symm b₁) (e.symm b₂)).2 ((e.apply_symm_apply b₁).symm ▸ (e.apply_symm_apply b₂).symm ▸ h) left_inv := by rintro ⟨a⟩ - dsimp only [Quot.map] - simp only [Equiv.symm_apply_apply] + simp only [Quot.map, Equiv.symm_apply_apply] right_inv := by rintro ⟨a⟩ - dsimp only [Quot.map] - simp only [Equiv.apply_symm_apply] + simp only [Quot.map, Equiv.apply_symm_apply] @[simp] theorem congr_mk {ra : α → α → Prop} {rb : β → β → Prop} (e : α ≃ β) (eq : ∀ a₁ a₂ : α, ra a₁ a₂ ↔ rb (e a₁) (e a₂)) (a : α) : - Quot.congr e Eq (Quot.mk ra a) = Quot.mk rb (e a) := + Quot.congr e eq (Quot.mk ra a) = Quot.mk rb (e a) := rfl /-- Quotients are congruent on equivalences under equality of their relation. An alternative is just to use rewriting with `eq`, but then computational proofs get stuck. -/ protected def congrRight {r r' : α → α → Prop} (eq : ∀ a₁ a₂, r a₁ a₂ ↔ r' a₁ a₂) : Quot r ≃ Quot r' := - Quot.congr (Equiv.refl α) Eq + Quot.congr (Equiv.refl α) eq /-- An equivalence `e : α ≃ β` generates an equivalence between the quotient space of `α` by a relation `ra` and the quotient space of `β` by the image of this relation under `e`. -/ protected def congrLeft {r : α → α → Prop} (e : α ≃ β) : Quot r ≃ Quot fun b b' => r (e.symm b) (e.symm b') := @Quot.congr α β r (fun b b' => r (e.symm b) (e.symm b')) e - fun a₁ a₂ => by simp only [e.symm_apply_apply] + fun a₁ a₂ => by simp only [e.symm_apply_apply] <;> rfl end Quot @@ -965,20 +1042,20 @@ namespace Quotient /-- An equivalence `e : α ≃ β` generates an equivalence between quotient spaces, if `ra a₁ a₂ ↔ rb (e a₁) (e a₂). -/ protected def congr {ra : Setoid α} {rb : Setoid β} (e : α ≃ β) - (eq : ∀ a₁ a₂, @Setoid.R α ra a₁ a₂ ↔ @Setoid.R β rb (e a₁) (e a₂)) : + (eq : ∀ a₁ a₂, @Setoid.r α ra a₁ a₂ ↔ @Setoid.r β rb (e a₁) (e a₂)) : Quotient ra ≃ Quotient rb := - Quot.congr e Eq + Quot.congr e eq @[simp] theorem congr_mk {ra : Setoid α} {rb : Setoid β} (e : α ≃ β) - (eq : ∀ a₁ a₂ : α, Setoid.R a₁ a₂ ↔ Setoid.R (e a₁) (e a₂)) (a : α) : - Quotient.congr e Eq (Quotient.mk a) = Quotient.mk (e a) := + (eq : ∀ a₁ a₂ : α, Setoid.r a₁ a₂ ↔ Setoid.r (e a₁) (e a₂)) (a : α) : + Quotient.congr e eq (Quotient.mk ra a) = Quotient.mk rb (e a) := rfl /-- Quotients are congruent on equivalences under equality of their relation. An alternative is just to use rewriting with `eq`, but then computational proofs get stuck. -/ protected def congrRight {r r' : Setoid α} - (eq : ∀ a₁ a₂, @Setoid.R α r a₁ a₂ ↔ @Setoid.R α r' a₁ a₂) : Quotient r ≃ Quotient r' := - Quot.congrRight Eq + (eq : ∀ a₁ a₂, @Setoid.r α r a₁ a₂ ↔ @Setoid.r α r' a₁ a₂) : Quotient r ≃ Quotient r' := + Quot.congrRight eq end Quotient From 5ecbbc1629c7e3fac64022a435e4cb50c98a4c0e Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Mon, 14 Nov 2022 10:43:28 +1100 Subject: [PATCH 016/127] long lines, fix error --- Mathlib/Logic/Equiv/Defs.lean | 6 ++++-- 1 file changed, 4 insertions(+), 2 deletions(-) diff --git a/Mathlib/Logic/Equiv/Defs.lean b/Mathlib/Logic/Equiv/Defs.lean index 9e031499e4c0d..93f7dcc65e1f7 100644 --- a/Mathlib/Logic/Equiv/Defs.lean +++ b/Mathlib/Logic/Equiv/Defs.lean @@ -752,7 +752,8 @@ def psigmaCongrRight {α} {β₁ β₂ : α → Sort _} (F : ∀ a, β₁ a ≃ @[simp] theorem psigmaCongrRight_trans {α} {β₁ β₂ β₃ : α → Sort _} (F : ∀ a, β₁ a ≃ β₂ a) (G : ∀ a, β₂ a ≃ β₃ a) : - (psigmaCongrRight F).trans (psigmaCongrRight G) = psigmaCongrRight fun a => (F a).trans (G a) := by + (psigmaCongrRight F).trans (psigmaCongrRight G) = + psigmaCongrRight fun a => (F a).trans (G a) := by ext1 x cases x rfl @@ -1008,7 +1009,8 @@ protected def congr {ra : α → α → Prop} {rb : β → β → Prop} (e : α toFun := Quot.map e fun a₁ a₂ => (eq a₁ a₂).1 invFun := Quot.map e.symm fun b₁ b₂ h => - (eq (e.symm b₁) (e.symm b₂)).2 ((e.apply_symm_apply b₁).symm ▸ (e.apply_symm_apply b₂).symm ▸ h) + (eq (e.symm b₁) (e.symm b₂)).2 + ((e.apply_symm_apply b₁).symm ▸ (e.apply_symm_apply b₂).symm ▸ h) left_inv := by rintro ⟨a⟩ simp only [Quot.map, Equiv.symm_apply_apply] From bfb9fe49100e3a91fd6f12e5e89df7124546478d Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Mon, 14 Nov 2022 10:53:35 +1100 Subject: [PATCH 017/127] some additional aligns --- Mathlib/Logic/Equiv/Defs.lean | 57 ++++++++--------------------------- 1 file changed, 12 insertions(+), 45 deletions(-) diff --git a/Mathlib/Logic/Equiv/Defs.lean b/Mathlib/Logic/Equiv/Defs.lean index 93f7dcc65e1f7..8dfed10ee53f7 100644 --- a/Mathlib/Logic/Equiv/Defs.lean +++ b/Mathlib/Logic/Equiv/Defs.lean @@ -385,7 +385,6 @@ def equivCongr {δ} (ab : α ≃ β) (cd : γ ≃ δ) : (α ≃ γ) ≃ (β ≃ theorem equivCongr_refl {α β} : (Equiv.refl α).equivCongr (Equiv.refl β) = Equiv.refl (α ≃ β) := by ext rfl - #align equiv.equiv_congr_refl Equiv.equivCongr_refl @[simp] @@ -393,7 +392,6 @@ theorem equivCongr_symm {δ} (ab : α ≃ β) (cd : γ ≃ δ) : (ab.equivCongr cd).symm = ab.symm.equivCongr cd.symm := by ext rfl - #align equiv.equiv_congr_symm Equiv.equivCongr_symm @[simp] @@ -401,28 +399,24 @@ theorem equivCongr_trans {δ ε ζ} (ab : α ≃ β) (de : δ ≃ ε) (bc : β (ab.equivCongr de).trans (bc.equivCongr ef) = (ab.trans bc).equivCongr (de.trans ef) := by ext rfl - #align equiv.equiv_congr_trans Equiv.equivCongr_trans @[simp] theorem equivCongr_refl_left {α β γ} (bg : β ≃ γ) (e : α ≃ β) : (Equiv.refl α).equivCongr bg e = e.trans bg := rfl - #align equiv.equiv_congr_refl_left Equiv.equivCongr_refl_left @[simp] theorem equivCongr_refl_right {α β} (ab e : α ≃ β) : ab.equivCongr (Equiv.refl β) e = ab.symm.trans e := rfl - #align equiv.equiv_congr_refl_right Equiv.equivCongr_refl_right @[simp] theorem equivCongr_apply_apply {δ} (ab : α ≃ β) (cd : γ ≃ δ) (e : α ≃ γ) (x) : ab.equivCongr cd e x = cd (e (ab.symm x)) := rfl - #align equiv.equiv_congr_apply_apply Equiv.equivCongr_apply_apply section permCongr @@ -435,36 +429,30 @@ def permCongr : Perm α' ≃ Perm β' := theorem permCongr_def (p : Equiv.Perm α') : e.permCongr p = (e.symm.trans p).trans e := rfl - #align equiv.perm_congr_def Equiv.permCongr_def @[simp] theorem permCongr_refl : e.permCongr (Equiv.refl _) = Equiv.refl _ := by simp [permCongr_def] - #align equiv.perm_congr_refl Equiv.permCongr_refl @[simp] theorem permCongr_symm : e.permCongr.symm = e.symm.permCongr := rfl - #align equiv.perm_congr_symm Equiv.permCongr_symm @[simp] theorem permCongr_apply (p : Equiv.Perm α') (x) : e.permCongr p x = e (p (e.symm x)) := rfl - #align equiv.perm_congr_apply Equiv.permCongr_apply theorem permCongr_symm_apply (p : Equiv.Perm β') (x) : e.permCongr.symm p x = e.symm (p (e x)) := rfl - #align equiv.perm_congr_symm_apply Equiv.permCongr_symm_apply theorem permCongr_trans (p p' : Equiv.Perm α') : (e.permCongr p).trans (e.permCongr p') = e.permCongr (p.trans p') := by ext simp only [trans_apply, comp_apply, permCongr_apply, symm_apply_apply] - #align equiv.perm_congr_trans Equiv.permCongr_trans end permCongr @@ -478,16 +466,19 @@ def equivEmpty (α : Sort u) [IsEmpty α] : α ≃ Empty := equivOfIsEmpty α _ /-- If `α` is an empty type, then it is equivalent to the `PEmpty` type in any universe. -/ -def equivPempty (α : Sort v) [IsEmpty α] : α ≃ PEmpty.{u} := +def equivPEmpty (α : Sort v) [IsEmpty α] : α ≃ PEmpty.{u} := equivOfIsEmpty α _ +#align equiv.equiv_pempty Equiv.equivPEmpty /-- `α` is equivalent to an empty type iff `α` is empty. -/ def equivEmptyEquiv (α : Sort u) : α ≃ Empty ≃ IsEmpty α := - ⟨fun e => Function.is_empty e, @equivEmpty α, fun e => ext fun x => (e x).elim, fun _ => rfl⟩ + ⟨fun e => Function.isEmpty e, @equivEmpty α, fun e => ext fun x => (e x).elim, fun _ => rfl⟩ +#align equiv.equiv_empty_equiv Equiv.equivEmptyEquiv /-- The `Sort` of proofs of a false proposition is equivalent to `PEmpty`. -/ -def propEquivPempty {p : Prop} (h : ¬p) : p ≃ PEmpty := - @equivPempty p <| IsEmpty.prop_iff.2 h +def propEquivPEmpty {p : Prop} (h : ¬p) : p ≃ PEmpty := + @equivPEmpty p <| IsEmpty.prop_iff.2 h +#align equiv.prop_equiv_pempty Equiv.propEquivPEmpty /-- If both `α` and `β` have a unique element, then `α ≃ β`. -/ def equivOfUnique (α β : Sort _) [Unique.{u} α] [Unique.{v} β] : α ≃ β where @@ -497,14 +488,16 @@ def equivOfUnique (α β : Sort _) [Unique.{u} α] [Unique.{v} β] : α ≃ β w right_inv _ := Subsingleton.elim _ _ /-- If `α` has a unique element, then it is equivalent to any `PUnit`. -/ -def equivPunit (α : Sort u) [Unique α] : α ≃ PUnit.{v} := +def equivPUnit (α : Sort u) [Unique α] : α ≃ PUnit.{v} := equivOfUnique α _ +#align equiv.equiv_punit Equiv.equivPUnit /-- The `Sort` of proofs of a true proposition is equivalent to `PUnit`. -/ -def propEquivPunit {p : Prop} (h : p) : p ≃ PUnit.{0} := +def propEquivPUnit {p : Prop} (h : p) : p ≃ PUnit.{0} := -- porting note: the following causes an error "unknown free variable [anonymous]" -- @equivPunit p <| uniqueProp h ⟨default, λ _ => h, λ _ => Subsingleton.elim _ _, λ _ => Subsingleton.elim _ _⟩ +#align equiv.prop_equiv_punit Equiv.propEquivPUnit /-- `ULift α` is equivalent to `α`. -/ @[simps (config := { fullyApplied := false }) apply symmApply] @@ -532,7 +525,6 @@ def arrowCongr {α₁ β₁ α₂ β₂ : Sort _} (e₁ : α₁ ≃ α₂) (e₂ invFun f := e₂.symm ∘ f ∘ e₁ left_inv f := funext fun x => by simp only [comp_apply, symm_apply_apply] right_inv f := funext fun x => by simp only [comp_apply, apply_symm_apply] - #align equiv.arrow_congr_apply Equiv.arrowCongr_apply theorem arrowCongr_comp {α₁ β₁ γ₁ α₂ β₂ γ₂ : Sort _} (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) (ec : γ₁ ≃ γ₂) @@ -540,14 +532,12 @@ theorem arrowCongr_comp {α₁ β₁ γ₁ α₂ β₂ γ₂ : Sort _} (ea : α arrowCongr ea ec (g ∘ f) = arrowCongr eb ec g ∘ arrowCongr ea eb f := by ext simp only [comp, arrowCongr_apply, eb.symm_apply_apply] - #align equiv.arrow_congr_comp Equiv.arrowCongr_comp @[simp] theorem arrowCongr_refl {α β : Sort _} : arrowCongr (Equiv.refl α) (Equiv.refl β) = Equiv.refl (α → β) := rfl - #align equiv.arrow_congr_refl Equiv.arrowCongr_refl @[simp] @@ -555,14 +545,12 @@ theorem arrowCongr_trans {α₁ β₁ α₂ β₂ α₃ β₃ : Sort _} (e₁ : α₁ ≃ α₂) (e₁' : β₁ ≃ β₂) (e₂ : α₂ ≃ α₃) (e₂' : β₂ ≃ β₃) : arrowCongr (e₁.trans e₂) (e₁'.trans e₂') = (arrowCongr e₁ e₁').trans (arrowCongr e₂ e₂') := rfl - #align equiv.arrow_congr_trans Equiv.arrowCongr_trans @[simp] theorem arrowCongr_symm {α₁ β₁ α₂ β₂ : Sort _} (e₁ : α₁ ≃ α₂) (e₂ : β₁ ≃ β₂) : (arrowCongr e₁ e₂).symm = arrowCongr e₁.symm e₂.symm := rfl - #align equiv.arrow_congr_symm Equiv.arrowCongr_symm /-- A version of `Equiv.arrowCongr` in `Type`, rather than `Sort`. @@ -579,7 +567,6 @@ def arrowCongr' {α₁ β₁ α₂ β₂ : Type _} (hα : α₁ ≃ α₂) (hβ theorem arrowCongr'_refl {α β : Type _} : arrowCongr' (Equiv.refl α) (Equiv.refl β) = Equiv.refl (α → β) := rfl - #align equiv.arrow_congr'_refl Equiv.arrowCongr'_refl @[simp] @@ -587,14 +574,12 @@ theorem arrowCongr'_trans {α₁ β₁ α₂ β₂ α₃ β₃ : Type _} (e₁ : α₁ ≃ α₂) (e₁' : β₁ ≃ β₂) (e₂ : α₂ ≃ α₃) (e₂' : β₂ ≃ β₃) : arrowCongr' (e₁.trans e₂) (e₁'.trans e₂') = (arrowCongr' e₁ e₁').trans (arrowCongr' e₂ e₂') := rfl - #align equiv.arrow_congr'_trans Equiv.arrowCongr'_trans @[simp] theorem arrowCongr'_symm {α₁ β₁ α₂ β₂ : Type _} (e₁ : α₁ ≃ α₂) (e₂ : β₁ ≃ β₂) : (arrowCongr' e₁ e₂).symm = arrowCongr' e₁.symm e₂.symm := rfl - #align equiv.arrow_congr'_symm Equiv.arrowCongr'_symm /-- Conjugate a map `f : α → α` by an equivalence `α ≃ β`. -/ @@ -639,7 +624,6 @@ def punitEquivPUnit : PUnit.{v} ≃ PUnit.{w} := rfl, fun u => by cases u rfl⟩ - #align equiv.punit_equiv_punit Equiv.punitEquivPUnit /-- `Prop` is noncomputably equivalent to `Bool`. -/ @@ -662,7 +646,6 @@ def arrowPUnitEquivPUnit (α : Sort _) : (α → PUnit.{v}) ≃ PUnit.{w} where right_inv u := by cases u rfl - #align equiv.arrow_punit_equiv_punit Equiv.arrowPUnitEquivPUnit /-- If `α` is `Subsingleton` and `a : α`, then the type of dependent functions `Π (i : α), β i` @@ -698,25 +681,21 @@ def arrowPUnitOfIsEmpty (α β : Sort _) [IsEmpty α] : (α → β) ≃ PUnit.{u right_inv u := by cases u rfl - #align equiv.arrow_punit_of_is_empy Equiv.arrowPUnitOfIsEmpty /-- The sort of maps from `Empty` is equivalent to `PUnit`. -/ def emptyArrowEquivPUnit (α : Sort _) : (Empty → α) ≃ PUnit.{u} := arrowPUnitOfIsEmpty _ _ - #align equiv.empty_arrow_equiv_punit Equiv.emptyArrowEquivPUnit /-- The sort of maps from `PEmpty` is equivalent to `PUnit`. -/ def pemptyArrowEquivPUnit (α : Sort _) : (PEmpty → α) ≃ PUnit.{u} := arrowPUnitOfIsEmpty _ _ - #align equiv.pempty_arrow_equiv_punit Equiv.pemptyArrowEquivPUnit /-- The sort of maps from `False` is equivalent to `PUnit`. -/ def falseArrowEquivPUnit (α : Sort _) : (False → α) ≃ PUnit.{u} := arrowPUnitOfIsEmpty _ _ - #align equiv.false_arrow_equiv_punit Equiv.falseArrowEquivPUnit end @@ -738,7 +717,6 @@ def psigmaEquivSigmaPLift {α} (β : α → Sort _) : (Σ' i, β i) ≃ Σ i : P invFun a := ⟨a.1.down, a.2.down⟩ left_inv := λ ⟨_, _⟩ => rfl right_inv := λ ⟨_, _⟩ => rfl - #align equiv.psigma_equiv_sigma_plift Equiv.psigmaEquivSigmaPLift /-- A family of equivalences `Π a, β₁ a ≃ β₂ a` generates an equivalence between `Σ' a, β₁ a` and @@ -757,7 +735,6 @@ theorem psigmaCongrRight_trans {α} {β₁ β₂ β₃ : α → Sort _} ext1 x cases x rfl - #align equiv.psigma_congr_right_trans Equiv.psigmaCongrRight_trans @[simp] @@ -766,7 +743,6 @@ theorem psigmaCongrRight_symm {α} {β₁ β₂ : α → Sort _} (F : ∀ a, β ext1 x cases x rfl - #align equiv.psigma_congr_right_symm Equiv.psigmaCongrRight_symm @[simp] @@ -775,7 +751,6 @@ theorem psigmaCongrRight_refl {α} {β : α → Sort _} : ext1 x cases x rfl - #align equiv.psigma_congr_right_refl Equiv.psigmaCongrRight_refl /-- A family of equivalences `Π a, β₁ a ≃ β₂ a` generates an equivalence between `Σ a, β₁ a` and @@ -793,7 +768,6 @@ theorem sigmaCongrRight_trans {α} {β₁ β₂ β₃ : α → Type _} ext1 x cases x rfl - #align equiv.sigmaCongrRight Equiv.sigmaCongrRight_trans @[simp] @@ -802,7 +776,6 @@ theorem sigmaCongrRight_symm {α} {β₁ β₂ : α → Type _} (F : ∀ a, β ext1 x cases x rfl - #align equiv.sigma_congr_right_symm Equiv.sigmaCongrRight_symm @[simp] @@ -811,7 +784,6 @@ theorem sigmaCongrRight_refl {α} {β : α → Type _} : ext1 x cases x rfl - #align equiv.sigma_congr_right_refl Equiv.sigmaCongrRight_refl /-- A `psigma` with `Prop` fibers is equivalent to the subtype. -/ @@ -829,7 +801,6 @@ def psigmaEquivSubtype {α : Type v} (P : α → Prop) : (Σ' i, P i) ≃ Subtyp def sigmaPLiftEquivSubtype {α : Type v} (P : α → Prop) : (Σ i, PLift (P i)) ≃ Subtype P := ((psigmaEquivSigma _).symm.trans (psigmaCongrRight fun _ => Equiv.plift)).trans (psigmaEquivSubtype P) - #align equiv.sigma_plift_equiv_subtype Equiv.sigmaPLiftEquivSubtype /-- A `sigma` with `λ i, ulift (plift (P i))` fibers is equivalent to `{ x // P x }`. @@ -838,7 +809,6 @@ Variant of `sigma_plift_equiv_subtype`. def sigmaULiftPLiftEquivSubtype {α : Type v} (P : α → Prop) : (Σ i, ULift (PLift (P i))) ≃ Subtype P := (sigmaCongrRight fun _ => Equiv.ulift).trans (sigmaPLiftEquivSubtype P) - #align equiv.sigma_ulift_plift_equiv_subtype Equiv.sigmaULiftPLiftEquivSubtype namespace Perm @@ -852,21 +822,18 @@ def sigmaCongrRight {α} {β : α → Sort _} (F : ∀ a, Perm (β a)) : Perm ( theorem sigmaCongrRight_trans {α} {β : α → Sort _} (F : ∀ a, Perm (β a)) (G : ∀ a, Perm (β a)) : (sigmaCongrRight F).trans (sigmaCongrRight G) = sigmaCongrRight fun a => (F a).trans (G a) := Equiv.sigmaCongrRight_trans F G - #align equiv.perm.sigma_congr_right_trans Equiv.Perm.sigmaCongrRight_trans @[simp] theorem sigmaCongrRight_symm {α} {β : α → Sort _} (F : ∀ a, Perm (β a)) : (sigmaCongrRight F).symm = sigmaCongrRight fun a => (F a).symm := Equiv.sigmaCongrRight_symm F - #align equiv.perm.sigma_congr_right_symm Equiv.Perm.sigmaCongrRight_symm @[simp] theorem sigmaCongrRight_refl {α} {β : α → Sort _} : (sigmaCongrRight fun a => Equiv.refl (β a)) = Equiv.refl (Σ a, β a) := Equiv.sigmaCongrRight_refl - #align equiv.perm.sigma_congr_right_refl Equiv.Perm.sigmaCongrRight_refl end Perm @@ -1035,7 +1002,7 @@ by a relation `ra` and the quotient space of `β` by the image of this relation protected def congrLeft {r : α → α → Prop} (e : α ≃ β) : Quot r ≃ Quot fun b b' => r (e.symm b) (e.symm b') := @Quot.congr α β r (fun b b' => r (e.symm b) (e.symm b')) e - fun a₁ a₂ => by simp only [e.symm_apply_apply] <;> rfl + fun a₁ a₂ => by simp only [e.symm_apply_apply]; rfl end Quot From 889dce55fc8565bfd6cd010f7571dfbd11330577 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Mon, 14 Nov 2022 11:18:22 +1100 Subject: [PATCH 018/127] ? --- Mathlib/Init/Data/Bool/Lemmas.lean | 5 ----- 1 file changed, 5 deletions(-) diff --git a/Mathlib/Init/Data/Bool/Lemmas.lean b/Mathlib/Init/Data/Bool/Lemmas.lean index 4c58bd6377010..0f12f621dde42 100644 --- a/Mathlib/Init/Data/Bool/Lemmas.lean +++ b/Mathlib/Init/Data/Bool/Lemmas.lean @@ -123,11 +123,6 @@ theorem coe_sort_false : (↥false : Prop) = False := by simp theorem coe_sort_true : (↥true : Prop) = True := by simp #align coe_sort_tt Bool.coe_sort_true -@[simp] theorem decide_eq_true (b : Bool) {h} : @decide (b = true) h = b := -by cases b - . simp - . simp - theorem decide_iff (p : Prop) [d : Decidable p] : decide p = true ↔ p := by simp #align to_bool_iff Bool.decide_iff From d5fa3ff55df2af3c6db1a2cac01979f13c973b64 Mon Sep 17 00:00:00 2001 From: Yakov Pechersky Date: Sun, 13 Nov 2022 21:43:58 -0500 Subject: [PATCH 019/127] fix right_inv proof --- Mathlib/Logic/Equiv/Defs.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Logic/Equiv/Defs.lean b/Mathlib/Logic/Equiv/Defs.lean index 8dfed10ee53f7..71e87f3a50c6f 100644 --- a/Mathlib/Logic/Equiv/Defs.lean +++ b/Mathlib/Logic/Equiv/Defs.lean @@ -631,7 +631,7 @@ noncomputable def propEquivBool : Prop ≃ Bool where toFun p := @decide p (Classical.propDecidable _) invFun b := b left_inv p := by simp [@Bool.decide_iff p (Classical.propDecidable _)] - right_inv b := by simp + right_inv b := by cases b; simp section From be11406943267cea455d9776f44099d229bad58c Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Mon, 14 Nov 2022 13:52:30 +1100 Subject: [PATCH 020/127] fix again --- Mathlib/Logic/Equiv/Defs.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Logic/Equiv/Defs.lean b/Mathlib/Logic/Equiv/Defs.lean index 71e87f3a50c6f..c3d07c1009863 100644 --- a/Mathlib/Logic/Equiv/Defs.lean +++ b/Mathlib/Logic/Equiv/Defs.lean @@ -631,7 +631,7 @@ noncomputable def propEquivBool : Prop ≃ Bool where toFun p := @decide p (Classical.propDecidable _) invFun b := b left_inv p := by simp [@Bool.decide_iff p (Classical.propDecidable _)] - right_inv b := by cases b; simp + right_inv b := by cases b <;> simp section From d6d3b834517bb0ba128a1d4cefec4f86f94a6e2f Mon Sep 17 00:00:00 2001 From: Yakov Pechersky Date: Sun, 13 Nov 2022 22:03:19 -0500 Subject: [PATCH 021/127] Use Logic.Equiv.Defs instead of Data.Equiv.Basic --- Mathlib.lean | 1 - Mathlib/Data/Equiv/Basic.lean | 67 --------------------------------- Mathlib/Data/Equiv/Functor.lean | 2 +- test/Simps.lean | 2 +- 4 files changed, 2 insertions(+), 70 deletions(-) delete mode 100644 Mathlib/Data/Equiv/Basic.lean diff --git a/Mathlib.lean b/Mathlib.lean index 647476e3cdfd1..54e9f4646ac53 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -26,7 +26,6 @@ import Mathlib.Data.Bool.Basic import Mathlib.Data.Bracket import Mathlib.Data.ByteArray import Mathlib.Data.Char -import Mathlib.Data.Equiv.Basic import Mathlib.Data.Equiv.Functor import Mathlib.Data.Fin.Basic import Mathlib.Data.Fin.Fin2 diff --git a/Mathlib/Data/Equiv/Basic.lean b/Mathlib/Data/Equiv/Basic.lean deleted file mode 100644 index b46cd757c1873..0000000000000 --- a/Mathlib/Data/Equiv/Basic.lean +++ /dev/null @@ -1,67 +0,0 @@ -/- -Copyright (c) 2015 Microsoft Corporation. All rights reserved. -Released under Apache 2.0 license as described in the file LICENSE. -Authors: Leonardo de Moura, Mario Carneiro --/ -import Mathlib.Logic.Function.Basic -/-! -# Equivalence between types - -* `equiv α β` a.k.a. `α ≃ β`: a bijective map `α → β` bundled with its inverse map; we use this (and - not equality!) to express that various `Type`s or `Sort`s are equivalent. - -## Tags - -equivalence, congruence, bijective map --/ - -open Function - -variable {α : Sort u} {β : Sort v} {γ : Sort w} - -/-- `α ≃ β` is the type of functions from `α → β` with a two-sided inverse. -/ -structure Equiv (α : Sort u) (β : Sort v) where - toFun : α → β - invFun : β → α - left_inv : LeftInverse invFun toFun - right_inv : RightInverse invFun toFun - -infix:25 " ≃ " => Equiv - -namespace Equiv - -/-- `perm α` is the type of bijections from `α` to itself. -/ -@[reducible] def perm (α : Sort u) := Equiv α α - -instance : CoeFun (α ≃ β) (λ _ => α → β):= -⟨toFun⟩ - --- Does not need to be simp, since the coercion is the projection, --- which simp has built-in support for. -theorem coe_fn_mk (f : α → β) (g l r) : (Equiv.mk f g l r : α → β) = f := -rfl - -def refl (α) : α ≃ α := ⟨id, id, λ _ => rfl, λ _ => rfl⟩ - -instance : Inhabited (α ≃ α) := ⟨Equiv.refl α⟩ - -def symm (e : α ≃ β) : β ≃ α := ⟨e.invFun, e.toFun, e.right_inv, e.left_inv⟩ - - -def trans (e₁ : α ≃ β) (e₂ : β ≃ γ) : α ≃ γ := -⟨e₂ ∘ (e₁ : α → β), e₁.symm ∘ (e₂.symm : γ → β), - e₂.left_inv.comp e₁.left_inv, e₂.right_inv.comp e₁.right_inv⟩ - -@[simp] theorem inv_fun_as_coe (e : α ≃ β) : e.invFun = e.symm := rfl - -@[simp] theorem apply_symm_apply (e : α ≃ β) (x : β) : e (e.symm x) = x := -e.right_inv x - -@[simp] theorem symm_apply_apply (e : α ≃ β) (x : α) : e.symm (e x) = x := -e.left_inv x - -@[simp] theorem symm_comp_self (e : α ≃ β) : e.symm ∘ (e : α → β) = id := funext e.symm_apply_apply - -@[simp] theorem self_comp_symm (e : α ≃ β) : e ∘ (e.symm : β → α) = id := funext e.apply_symm_apply - -end Equiv diff --git a/Mathlib/Data/Equiv/Functor.lean b/Mathlib/Data/Equiv/Functor.lean index 8d055ca84ff36..60cb5fabb94f6 100644 --- a/Mathlib/Data/Equiv/Functor.lean +++ b/Mathlib/Data/Equiv/Functor.lean @@ -3,7 +3,7 @@ Copyright (c) 2019 Johan Commelin. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Johan Commelin, Simon Hudon, Scott Morrison -/ -import Mathlib.Data.Equiv.Basic +import Mathlib.Logic.Equiv.Defs /-! diff --git a/test/Simps.lean b/test/Simps.lean index 5216ca8a1c6ee..51a940ebac1ed 100644 --- a/test/Simps.lean +++ b/test/Simps.lean @@ -2,7 +2,7 @@ import Mathlib.Algebra.Group.Defs import Mathlib.Tactic.Simps.Basic import Mathlib.Tactic.RunCmd import Mathlib.Lean.Exception -import Mathlib.Data.Equiv.Basic +import Mathlib.Logic.Equiv.Defs import Mathlib.Data.Prod.Basic -- set_option trace.simps.debug true From 8fa0ee6e979f408b204e5e2e0e5f9e5314721655 Mon Sep 17 00:00:00 2001 From: Yakov Pechersky Date: Sun, 13 Nov 2022 23:43:57 -0500 Subject: [PATCH 022/127] fix and speed up Functor `map_equiv` --- Mathlib/Data/Equiv/Functor.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Data/Equiv/Functor.lean b/Mathlib/Data/Equiv/Functor.lean index 60cb5fabb94f6..9c7f1e1ebc7c2 100644 --- a/Mathlib/Data/Equiv/Functor.lean +++ b/Mathlib/Data/Equiv/Functor.lean @@ -31,8 +31,8 @@ theorem map_map (m : α → β) (g : β → γ) (x : f α) : def map_equiv (h : α ≃ β) : f α ≃ f β where toFun := map h invFun := map h.symm - left_inv x := by simp [map_map] - right_inv x := by simp [map_map] + left_inv x := by rw [map_map, h.symm_comp_self, id_map] + right_inv x := by rw [map_map, h.self_comp_symm, id_map] @[simp] lemma map_equiv_apply (h : α ≃ β) (x : f α) : From 5f8ae4c224f9a6196aa207e75c06c02022ebd70f Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Mon, 14 Nov 2022 20:47:31 +1100 Subject: [PATCH 023/127] revert change that hides the simp failure --- Mathlib/Data/Equiv/Functor.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Data/Equiv/Functor.lean b/Mathlib/Data/Equiv/Functor.lean index 9c7f1e1ebc7c2..60cb5fabb94f6 100644 --- a/Mathlib/Data/Equiv/Functor.lean +++ b/Mathlib/Data/Equiv/Functor.lean @@ -31,8 +31,8 @@ theorem map_map (m : α → β) (g : β → γ) (x : f α) : def map_equiv (h : α ≃ β) : f α ≃ f β where toFun := map h invFun := map h.symm - left_inv x := by rw [map_map, h.symm_comp_self, id_map] - right_inv x := by rw [map_map, h.self_comp_symm, id_map] + left_inv x := by simp [map_map] + right_inv x := by simp [map_map] @[simp] lemma map_equiv_apply (h : α ≃ β) (x : f α) : From 97c6582e9464bb095505b8f9b592f5333aa6fc2e Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Mon, 14 Nov 2022 20:58:18 +1100 Subject: [PATCH 024/127] comment out potentially bad simps --- Mathlib/Data/Quot.lean | 46 +++++++++++++++++++++++++++++++----------- 1 file changed, 34 insertions(+), 12 deletions(-) diff --git a/Mathlib/Data/Quot.lean b/Mathlib/Data/Quot.lean index cae5ae929b713..3b1b69132b181 100644 --- a/Mathlib/Data/Quot.lean +++ b/Mathlib/Data/Quot.lean @@ -90,7 +90,9 @@ theorem lift_mk (f : α → γ) (h : ∀ a₁ a₂, r a₁ a₂ → f a₁ = f a rfl #align quot.lift_beta Quot.lift_mk -@[simp] +-- Porting note: this may be a bad simp lemma now that Lean indexes simp lemmas mod reducible. +-- We need a linter to catch these. +-- @[simp] theorem lift_on_mk (a : α) (f : α → γ) (h : ∀ a₁ a₂, r a₁ a₂ → f a₁ = f a₂) : Quot.liftOn (Quot.mk r a) f h = f a := rfl @@ -103,7 +105,9 @@ protected def lift₂ (f : α → β → γ) (hr : ∀ a b₁ b₂, s b₁ b₂ Quot.lift (fun a => Quot.lift (f a) (hr a)) (fun a₁ a₂ ha => funext fun q => Quot.induction_on q fun b => hs a₁ a₂ b ha) q₁ q₂ -@[simp] +-- Porting note: this may be a bad simp lemma now that Lean indexes simp lemmas mod reducible. +-- We need a linter to catch these. +-- @[simp] theorem lift₂_mk (f : α → β → γ) (hr : ∀ a b₁ b₂, s b₁ b₂ → f a b₁ = f a b₂) (hs : ∀ a₁ a₂ b, r a₁ a₂ → f a₁ b = f a₂ b) (a : α) (b : β) : Quot.lift₂ f hr hs (Quot.mk r a) (Quot.mk s b) = f a b := @@ -116,7 +120,9 @@ protected def liftOn₂ (p : Quot r) (q : Quot s) (f : α → β → γ) (hr : ∀ a b₁ b₂, s b₁ b₂ → f a b₁ = f a b₂) (hs : ∀ a₁ a₂ b, r a₁ a₂ → f a₁ b = f a₂ b) : γ := Quot.lift₂ f hr hs p q -@[simp] +-- Porting note: this may be a bad simp lemma now that Lean indexes simp lemmas mod reducible. +-- We need a linter to catch these. +-- @[simp] theorem lift_on₂_mk (a : α) (b : β) (f : α → β → γ) (hr : ∀ a b₁ b₂, s b₁ b₂ → f a b₁ = f a b₂) (hs : ∀ a₁ a₂ b, r a₁ a₂ → f a₁ b = f a₂ b) : Quot.liftOn₂ (Quot.mk r a) (Quot.mk s b) f hr hs = f a b := @@ -267,12 +273,16 @@ theorem forall_quotient_iff {α : Type _} [r : Setoid α] {p : Quotient r → Pr (∀ a : Quotient r, p a) ↔ ∀ a : α, p ⟦a⟧ := ⟨fun h _ => h _, fun h a => a.induction_on h⟩ -@[simp] +-- Porting note: this may be a bad simp lemma now that Lean indexes simp lemmas mod reducible. +-- We need a linter to catch these. +-- @[simp] theorem Quotient.lift_mk [s : Setoid α] (f : α → β) (h : ∀ a b : α, a ≈ b → f a = f b) (x : α) : Quotient.lift f h (Quotient.mk s x) = f x := - rfl +rfl -@[simp] +-- Porting note: this may be a bad simp lemma now that Lean indexes simp lemmas mod reducible. +-- We need a linter to catch these. +-- @[simp] theorem Quotient.lift_comp_mk [Setoid α] (f : α → β) (h : ∀ a b : α, a ≈ b → f a = f b) : Quotient.lift f h ∘ Quotient.mk _ = f := rfl @@ -285,12 +295,16 @@ theorem Quotient.lift₂_mk {α : Sort _} {β : Sort _} {γ : Sort _} [Setoid α Quotient.lift₂ f h (Quotient.mk _ a) (Quotient.mk _ b) = f a b := rfl -@[simp] +-- Porting note: this may be a bad simp lemma now that Lean indexes simp lemmas mod reducible. +-- We need a linter to catch these. +-- @[simp] theorem Quotient.lift_on_mk [s : Setoid α] (f : α → β) (h : ∀ a b : α, a ≈ b → f a = f b) (x : α) : Quotient.liftOn (Quotient.mk s x) f h = f x := rfl -@[simp] +-- Porting note: this may be a bad simp lemma now that Lean indexes simp lemmas mod reducible. +-- We need a linter to catch these. +-- @[simp] theorem Quotient.lift_on₂_mk {α : Sort _} {β : Sort _} [Setoid α] (f : α → α → β) (h : ∀ a₁ a₂ b₁ b₂ : α, a₁ ≈ b₁ → a₂ ≈ b₂ → f a₁ a₂ = f b₁ b₂) (x y : α) : Quotient.liftOn₂ (Quotient.mk _ x) (Quotient.mk _ y) f h = f x y := @@ -535,7 +549,9 @@ protected def liftOn' (q : Quotient s₁) (f : α → φ) (h : ∀ a b, @Setoid. φ := Quotient.liftOn q f h -@[simp] +-- Porting note: this may be a bad simp lemma now that Lean indexes simp lemmas mod reducible. +-- We need a linter to catch these. +-- @[simp] protected theorem liftOn'_mk'' (f : α → φ) (h) (x : α) : Quotient.liftOn' (@Quotient.mk'' _ s₁ x) f h = f x := rfl @@ -548,7 +564,9 @@ protected def liftOn₂' (q₁ : Quotient s₁) (q₂ : Quotient s₂) (f : α (h : ∀ a₁ a₂ b₁ b₂, @Setoid.r α s₁ a₁ b₁ → @Setoid.r β s₂ a₂ b₂ → f a₁ a₂ = f b₁ b₂) : γ := Quotient.liftOn₂ q₁ q₂ f h -@[simp] +-- Porting note: this may be a bad simp lemma now that Lean indexes simp lemmas mod reducible. +-- We need a linter to catch these. +-- @[simp] protected theorem liftOn₂'_mk'' (f : α → β → γ) (h) (a : α) (b : β) : Quotient.liftOn₂' (@Quotient.mk'' _ s₁ a) (@Quotient.mk'' _ s₂ b) f h = f a b := rfl @@ -689,11 +707,15 @@ variable [s : Setoid α] protected theorem mk''_eq_mk (x : α) : Quotient.mk'' x = Quotient.mk s x := rfl -@[simp] +-- Porting note: this may be a bad simp lemma now that Lean indexes simp lemmas mod reducible. +-- We need a linter to catch these. +-- @[simp] protected theorem lift_on'_mk (x : α) (f : α → β) (h) : (Quotient.mk s x).liftOn' f h = f x := rfl -@[simp] +-- Porting note: this may be a bad simp lemma now that Lean indexes simp lemmas mod reducible. +-- We need a linter to catch these. +-- @[simp] protected theorem lift_on₂'_mk [t : Setoid β] (f : α → β → γ) (h) (a : α) (b : β) : Quotient.liftOn₂' (Quotient.mk s a) (Quotient.mk t b) f h = f a b := Quotient.liftOn₂'_mk'' _ _ _ _ From 3e6f4d1f065532fd1f779f757bdef50ab1e9bc7e Mon Sep 17 00:00:00 2001 From: ruben vorster Date: Thu, 17 Nov 2022 19:20:16 +0000 Subject: [PATCH 025/127] initial commit --- Mathlib/Logic/Equiv/Basic.lean | 1811 ++++++++++++++++++++++++++++++++ 1 file changed, 1811 insertions(+) create mode 100644 Mathlib/Logic/Equiv/Basic.lean diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean new file mode 100644 index 0000000000000..bc1ca50a6e754 --- /dev/null +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -0,0 +1,1811 @@ +/- +Copyright (c) 2015 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Leonardo de Moura, Mario Carneiro +-/ +import Mathlib.Logic.Equiv.Defs +import Mathlib.Data.Prod.Basic +import Mathlib.Data.Sigma.Basic +import Mathlib.Data.Subtype +import Mathlib.Data.Sum.Basic +import Mathlib.Logic.Function.Conjugate + +/-! +# Equivalence between types + +In this file we continue the work on equivalences begun in `logic/equiv/defs.lean`, defining + +* canonical isomorphisms between various types: e.g., + + - `equiv.sum_equiv_sigma_bool` is the canonical equivalence between the sum of two types `α ⊕ β` + and the sigma-type `Σ b : bool, cond b α β`; + + - `equiv.prod_sum_distrib : α × (β ⊕ γ) ≃ (α × β) ⊕ (α × γ)` shows that type product and type sum + satisfy the distributive law up to a canonical equivalence; + +* operations on equivalences: e.g., + + - `equiv.prod_congr ea eb : α₁ × β₁ ≃ α₂ × β₂`: combine two equivalences `ea : α₁ ≃ α₂` and + `eb : β₁ ≃ β₂` using `prod.map`. + + More definitions of this kind can be found in other files. E.g., `data/equiv/transfer_instance` + does it for many algebraic type classes like `group`, `module`, etc. + +## Tags + +equivalence, congruence, bijective map +-/ + + +open Function + +universe u v w z + +variable {α : Sort u} {β : Sort v} {γ : Sort w} + +namespace Equiv + +/-- `pprod α β` is equivalent to `α × β` -/ +@[simps apply symmApply] +def pprodEquivProd {α β : Type _} : PProd α β ≃ α × β where + toFun x := (x.1, x.2) + invFun x := ⟨x.1, x.2⟩ + left_inv := fun _ => rfl + right_inv := fun _ => rfl +#align equiv.pprod_equiv_prod Equiv.pprodEquivProd + +/-- Product of two equivalences, in terms of `pprod`. If `α ≃ β` and `γ ≃ δ`, then +`pprod α γ ≃ pprod β δ`. -/ +--@[congr, simps apply] +--porting note: NEEDS FIXING +def pprodCongr {δ : Sort z} (e₁ : α ≃ β) (e₂ : γ ≃ δ) : PProd α γ ≃ PProd β δ where + toFun x := ⟨e₁ x.1, e₂ x.2⟩ + invFun x := ⟨e₁.symm x.1, e₂.symm x.2⟩ + left_inv := fun ⟨x, y⟩ => by simp + right_inv := fun ⟨x, y⟩ => by simp +#align equiv.pprod_congr Equiv.pprodCongr + +/-- Combine two equivalences using `pprod` in the domain and `prod` in the codomain. -/ +@[simps apply symmApply] +def pprodProd {α₁ β₁ : Sort _} {α₂ β₂ : Type _} (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) : PProd α₁ β₁ ≃ α₂ × β₂ := + (ea.pprodCongr eb).trans pprodEquivProd +#align equiv.pprod_prod Equiv.pprodProd + +/-- Combine two equivalences using `pprod` in the codomain and `prod` in the domain. -/ +@[simps apply symmApply] +def prodPprod {α₁ β₁ : Type _} {α₂ β₂ : Sort _} (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) : α₁ × β₁ ≃ PProd α₂ β₂ := + (ea.symm.pprodProd eb.symm).symm +#align equiv.prod_pprod Equiv.prodPprod + +/-- `pprod α β` is equivalent to `plift α × plift β` -/ +@[simps apply symmApply] +def pprodEquivProdPlift {α β : Sort _} : PProd α β ≃ PLift α × PLift β := + Equiv.plift.symm.pprodProd Equiv.plift.symm +#align equiv.pprod_equiv_prod_plift Equiv.pprodEquivProdPlift + +/-- Product of two equivalences. If `α₁ ≃ α₂` and `β₁ ≃ β₂`, then `α₁ × β₁ ≃ α₂ × β₂`. This is +`prod.map` as an equivalence. -/ +--@[congr, simps apply] +--porting note: NEEDS FIXING +def prodCongr {α₁ β₁ α₂ β₂ : Type _} (e₁ : α₁ ≃ α₂) (e₂ : β₁ ≃ β₂) : α₁ × β₁ ≃ α₂ × β₂ := + ⟨Prod.map e₁ e₂, Prod.map e₁.symm e₂.symm, fun ⟨a, b⟩ => by simp, fun ⟨a, b⟩ => by simp⟩ +#align equiv.prod_congr Equiv.prodCongr + +@[simp] +theorem prod_congr_symm {α₁ β₁ α₂ β₂ : Type _} (e₁ : α₁ ≃ α₂) (e₂ : β₁ ≃ β₂) : + (prodCongr e₁ e₂).symm = prodCongr e₁.symm e₂.symm := + rfl +#align equiv.prod_congr_symm Equiv.prod_congr_symm + +/-- Type product is commutative up to an equivalence: `α × β ≃ β × α`. This is `prod.swap` as an +equivalence.-/ +def prodComm (α β : Type _) : α × β ≃ β × α := + ⟨Prod.swap, Prod.swap, Prod.swap_swap, Prod.swap_swap⟩ +#align equiv.prod_comm Equiv.prodComm + +@[simp] +theorem coe_prod_comm (α β : Type _) : ⇑(prodComm α β) = Prod.swap := + rfl +#align equiv.coe_prod_comm Equiv.coe_prod_comm + +@[simp] +theorem prod_comm_apply {α β : Type _} (x : α × β) : prodComm α β x = x.swap := + rfl +#align equiv.prod_comm_apply Equiv.prod_comm_apply + +@[simp] +theorem prod_comm_symm (α β) : (prodComm α β).symm = prodComm β α := + rfl +#align equiv.prod_comm_symm Equiv.prod_comm_symm + +/-- Type product is associative up to an equivalence. -/ +@[simps] +def prodAssoc (α β γ : Sort _) : (α × β) × γ ≃ α × β × γ := + ⟨fun p => (p.1.1, p.1.2, p.2), fun p => ((p.1, p.2.1), p.2.2), fun ⟨⟨a, b⟩, c⟩ => rfl, fun ⟨a, ⟨b, c⟩⟩ => rfl⟩ +#align equiv.prod_assoc Equiv.prodAssoc + +/-- Functions on `α × β` are equivalent to functions `α → β → γ`. -/ +@[simps (config := { fullyApplied := false })] +def curry (α β γ : Type _) : (α × β → γ) ≃ (α → β → γ) where + toFun := Function.curry + invFun := uncurry + left_inv := uncurry_curry + right_inv := curry_uncurry +#align equiv.curry Equiv.curry + +section + +/-- `punit` is a right identity for type product up to an equivalence. -/ +@[simps] +def prodPunit (α : Type _) : α × PUnit.{u + 1} ≃ α := + ⟨fun p => p.1, fun a => (a, PUnit.unit), fun ⟨_, PUnit.unit⟩ => rfl, fun a => rfl⟩ +#align equiv.prod_punit Equiv.prodPunit + +/-- `punit` is a left identity for type product up to an equivalence. -/ +@[simps] +def punitProd (α : Type _) : PUnit.{u + 1} × α ≃ α := + calc + PUnit × α ≃ α × PUnit := prodComm _ _ + _ ≃ α := prodPunit _ + +#align equiv.punit_prod Equiv.punitProd + +/-- Any `unique` type is a right identity for type product up to equivalence. -/ +def prodUnique (α β : Type _) [Unique β] : α × β ≃ α := + ((Equiv.refl α).prodCongr <| equivPunit β).trans <| prodPunit α +#align equiv.prod_unique Equiv.prodUnique + +@[simp] +theorem coe_prod_unique {α β : Type _} [Unique β] : ⇑(prodUnique α β) = Prod.fst := + rfl +#align equiv.coe_prod_unique Equiv.coe_prod_unique + +theorem prod_unique_apply {α β : Type _} [Unique β] (x : α × β) : prodUnique α β x = x.1 := + rfl +#align equiv.prod_unique_apply Equiv.prod_unique_apply + +@[simp] +theorem prod_unique_symm_apply {α β : Type _} [Unique β] (x : α) : (prodUnique α β).symm x = (x, default) := + rfl +#align equiv.prod_unique_symm_apply Equiv.prod_unique_symm_apply + +/-- Any `unique` type is a left identity for type product up to equivalence. -/ +def uniqueProd (α β : Type _) [Unique β] : β × α ≃ α := + ((equivPunit β).prodCongr <| Equiv.refl α).trans <| punitProd α +#align equiv.unique_prod Equiv.uniqueProd + +@[simp] +theorem coe_unique_prod {α β : Type _} [Unique β] : ⇑(uniqueProd α β) = Prod.snd := + rfl +#align equiv.coe_unique_prod Equiv.coe_unique_prod + +theorem unique_prod_apply {α β : Type _} [Unique β] (x : β × α) : uniqueProd α β x = x.2 := + rfl +#align equiv.unique_prod_apply Equiv.unique_prod_apply + +@[simp] +theorem unique_prod_symm_apply {α β : Type _} [Unique β] (x : α) : (uniqueProd α β).symm x = (default, x) := + rfl +#align equiv.unique_prod_symm_apply Equiv.unique_prod_symm_apply + +/-- `empty` type is a right absorbing element for type product up to an equivalence. -/ +def prodEmpty (α : Type _) : α × Empty ≃ Empty := + equivEmpty _ +#align equiv.prod_empty Equiv.prodEmpty + +/-- `empty` type is a left absorbing element for type product up to an equivalence. -/ +def emptyProd (α : Type _) : Empty × α ≃ Empty := + equivEmpty _ +#align equiv.empty_prod Equiv.emptyProd + +/-- `pempty` type is a right absorbing element for type product up to an equivalence. -/ +def prodPempty (α : Type _) : α × PEmpty ≃ PEmpty := + equivPempty _ +#align equiv.prod_pempty Equiv.prodPempty + +/-- `pempty` type is a left absorbing element for type product up to an equivalence. -/ +def pemptyProd (α : Type _) : PEmpty × α ≃ PEmpty := + equivPempty _ +#align equiv.pempty_prod Equiv.pemptyProd + +end + +section + +open Sum + +/-- `psum` is equivalent to `sum`. -/ +def psumEquivSum (α β : Type _) : PSum α β ≃ Sum α β where + toFun s := PSum.casesOn s inl inr + invFun := Sum.elim PSum.inl PSum.inr + left_inv s := by cases s <;> rfl + right_inv s := by cases s <;> rfl +#align equiv.psum_equiv_sum Equiv.psumEquivSum + +/-- If `α ≃ α'` and `β ≃ β'`, then `α ⊕ β ≃ α' ⊕ β'`. This is `sum.map` as an equivalence. -/ +@[simps apply] +def sumCongr {α₁ β₁ α₂ β₂ : Type _} (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) : Sum α₁ β₁ ≃ Sum α₂ β₂ := + ⟨Sum.map ea eb, Sum.map ea.symm eb.symm, fun x => by simp, fun x => by simp⟩ +#align equiv.sum_congr Equiv.sumCongr + +/-- If `α ≃ α'` and `β ≃ β'`, then `psum α β ≃ psum α' β'`. -/ +def psumCongr {δ : Sort z} (e₁ : α ≃ β) (e₂ : γ ≃ δ) : PSum α γ ≃ PSum β δ where + toFun x := PSum.casesOn x (PSum.inl ∘ e₁) (PSum.inr ∘ e₂) + invFun x := PSum.casesOn x (PSum.inl ∘ e₁.symm) (PSum.inr ∘ e₂.symm) + left_inv := by rintro (x | x) <;> simp + right_inv := by rintro (x | x) <;> simp +#align equiv.psum_congr Equiv.psumCongr + +/-- Combine two `equiv`s using `psum` in the domain and `sum` in the codomain. -/ +def psumSum {α₁ β₁ : Sort _} {α₂ β₂ : Type _} (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) : PSum α₁ β₁ ≃ Sum α₂ β₂ := + (ea.psumCongr eb).trans (psumEquivSum _ _) +#align equiv.psum_sum Equiv.psumSum + +/-- Combine two `equiv`s using `sum` in the domain and `psum` in the codomain. -/ +def sumPsum {α₁ β₁ : Type _} {α₂ β₂ : Sort _} (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) : Sum α₁ β₁ ≃ PSum α₂ β₂ := + (ea.symm.psumSum eb.symm).symm +#align equiv.sum_psum Equiv.sumPsum + +@[simp] +theorem sum_congr_trans {α₁ α₂ β₁ β₂ γ₁ γ₂ : Sort _} (e : α₁ ≃ β₁) (f : α₂ ≃ β₂) (g : β₁ ≃ γ₁) (h : β₂ ≃ γ₂) : + (Equiv.sumCongr e f).trans (Equiv.sumCongr g h) = Equiv.sumCongr (e.trans g) (f.trans h) := by + ext i + cases i <;> rfl +#align equiv.sum_congr_trans Equiv.sum_congr_trans + +@[simp] +theorem sum_congr_symm {α β γ δ : Sort _} (e : α ≃ β) (f : γ ≃ δ) : + (Equiv.sumCongr e f).symm = Equiv.sumCongr e.symm f.symm := + rfl +#align equiv.sum_congr_symm Equiv.sum_congr_symm + +@[simp] +theorem sum_congr_refl {α β : Sort _} : Equiv.sumCongr (Equiv.refl α) (Equiv.refl β) = Equiv.refl (Sum α β) := by + ext i + cases i <;> rfl +#align equiv.sum_congr_refl Equiv.sum_congr_refl + +namespace Perm + +/-- Combine a permutation of `α` and of `β` into a permutation of `α ⊕ β`. -/ +@[reducible] +def sumCongr {α β : Type _} (ea : Equiv.Perm α) (eb : Equiv.Perm β) : Equiv.Perm (Sum α β) := + Equiv.sumCongr ea eb +#align equiv.perm.sum_congr Equiv.Perm.sumCongr + +@[simp] +theorem sum_congr_apply {α β : Type _} (ea : Equiv.Perm α) (eb : Equiv.Perm β) (x : Sum α β) : + sumCongr ea eb x = Sum.map (⇑ea) (⇑eb) x := + Equiv.sum_congr_apply ea eb x +#align equiv.perm.sum_congr_apply Equiv.Perm.sum_congr_apply + +@[simp] +theorem sum_congr_trans {α β : Sort _} (e : Equiv.Perm α) (f : Equiv.Perm β) (g : Equiv.Perm α) (h : Equiv.Perm β) : + (sumCongr e f).trans (sumCongr g h) = sumCongr (e.trans g) (f.trans h) := + Equiv.sum_congr_trans e f g h +#align equiv.perm.sum_congr_trans Equiv.Perm.sum_congr_trans + +@[simp] +theorem sum_congr_symm {α β : Sort _} (e : Equiv.Perm α) (f : Equiv.Perm β) : + (sumCongr e f).symm = sumCongr e.symm f.symm := + Equiv.sum_congr_symm e f +#align equiv.perm.sum_congr_symm Equiv.Perm.sum_congr_symm + +@[simp] +theorem sum_congr_refl {α β : Sort _} : sumCongr (Equiv.refl α) (Equiv.refl β) = Equiv.refl (Sum α β) := + Equiv.sum_congr_refl +#align equiv.perm.sum_congr_refl Equiv.Perm.sum_congr_refl + +end Perm + +/-- `bool` is equivalent the sum of two `punit`s. -/ +def boolEquivPunitSumPunit : Bool ≃ Sum PUnit.{u + 1} PUnit.{v + 1} := + ⟨fun b => cond b (inr PUnit.unit) (inl PUnit.unit), Sum.elim (fun _ => false) fun _ => true, fun b => by + cases b <;> rfl, fun s => by rcases s with (⟨⟨⟩⟩ | ⟨⟨⟩⟩) <;> rfl⟩ +#align equiv.bool_equiv_punit_sum_punit Equiv.boolEquivPunitSumPunit + +/-- Sum of types is commutative up to an equivalence. This is `sum.swap` as an equivalence. -/ +@[simps (config := { fullyApplied := false }) apply] +def sumComm (α β : Type _) : Sum α β ≃ Sum β α := + ⟨Sum.swap, Sum.swap, Sum.swap_swap, Sum.swap_swap⟩ +#align equiv.sum_comm Equiv.sumComm + +@[simp] +theorem sum_comm_symm (α β) : (sumComm α β).symm = sumComm β α := + rfl +#align equiv.sum_comm_symm Equiv.sum_comm_symm + +/-- Sum of types is associative up to an equivalence. -/ +def sumAssoc (α β γ : Type _) : Sum (Sum α β) γ ≃ Sum α (Sum β γ) := + ⟨Sum.elim (Sum.elim Sum.inl (Sum.inr ∘ Sum.inl)) (Sum.inr ∘ Sum.inr), + Sum.elim (Sum.inl ∘ Sum.inl) <| Sum.elim (Sum.inl ∘ Sum.inr) Sum.inr, by rintro (⟨_ | _⟩ | _) <;> rfl, by + rintro (_ | ⟨_ | _⟩) <;> rfl⟩ +#align equiv.sum_assoc Equiv.sumAssoc + +@[simp] +theorem sum_assoc_apply_inl_inl {α β γ} (a) : sumAssoc α β γ (inl (inl a)) = inl a := + rfl +#align equiv.sum_assoc_apply_inl_inl Equiv.sum_assoc_apply_inl_inl + +@[simp] +theorem sum_assoc_apply_inl_inr {α β γ} (b) : sumAssoc α β γ (inl (inr b)) = inr (inl b) := + rfl +#align equiv.sum_assoc_apply_inl_inr Equiv.sum_assoc_apply_inl_inr + +@[simp] +theorem sum_assoc_apply_inr {α β γ} (c) : sumAssoc α β γ (inr c) = inr (inr c) := + rfl +#align equiv.sum_assoc_apply_inr Equiv.sum_assoc_apply_inr + +@[simp] +theorem sum_assoc_symm_apply_inl {α β γ} (a) : (sumAssoc α β γ).symm (inl a) = inl (inl a) := + rfl +#align equiv.sum_assoc_symm_apply_inl Equiv.sum_assoc_symm_apply_inl + +@[simp] +theorem sum_assoc_symm_apply_inr_inl {α β γ} (b) : (sumAssoc α β γ).symm (inr (inl b)) = inl (inr b) := + rfl +#align equiv.sum_assoc_symm_apply_inr_inl Equiv.sum_assoc_symm_apply_inr_inl + +@[simp] +theorem sum_assoc_symm_apply_inr_inr {α β γ} (c) : (sumAssoc α β γ).symm (inr (inr c)) = inr c := + rfl +#align equiv.sum_assoc_symm_apply_inr_inr Equiv.sum_assoc_symm_apply_inr_inr + +/-- Sum with `empty` is equivalent to the original type. -/ +@[simps symmApply] +def sumEmpty (α β : Type _) [IsEmpty β] : Sum α β ≃ α := + ⟨Sum.elim id isEmptyElim, inl, fun s => by + rcases s with (_ | x) + rfl + exact isEmptyElim x, fun a => rfl⟩ +#align equiv.sum_empty Equiv.sumEmpty + +@[simp] +theorem sum_empty_apply_inl {α β : Type _} [IsEmpty β] (a : α) : sumEmpty α β (Sum.inl a) = a := + rfl +#align equiv.sum_empty_apply_inl Equiv.sum_empty_apply_inl + +/-- The sum of `empty` with any `Sort*` is equivalent to the right summand. -/ +@[simps symmApply] +def emptySum (α β : Type _) [IsEmpty α] : Sum α β ≃ β := + (sumComm _ _).trans <| sumEmpty _ _ +#align equiv.empty_sum Equiv.emptySum + +@[simp] +theorem empty_sum_apply_inr {α β : Type _} [IsEmpty α] (b : β) : emptySum α β (Sum.inr b) = b := + rfl +#align equiv.empty_sum_apply_inr Equiv.empty_sum_apply_inr + +/-- `option α` is equivalent to `α ⊕ punit` -/ +def optionEquivSumPunit (α : Type _) : Option α ≃ Sum α PUnit.{u + 1} := + ⟨fun o => o.elim (inr PUnit.unit) inl, fun s => s.elim some fun _ => none, fun o => by cases o <;> rfl, fun s => by + rcases s with (_ | ⟨⟨⟩⟩) <;> rfl⟩ +#align equiv.option_equiv_sum_punit Equiv.optionEquivSumPunit + +@[simp] +theorem option_equiv_sum_punit_none {α} : optionEquivSumPunit α none = Sum.inr PUnit.unit := + rfl +#align equiv.option_equiv_sum_punit_none Equiv.option_equiv_sum_punit_none + +@[simp] +theorem option_equiv_sum_punit_some {α} (a) : optionEquivSumPunit α (some a) = Sum.inl a := + rfl +#align equiv.option_equiv_sum_punit_some Equiv.option_equiv_sum_punit_some + +@[simp] +theorem option_equiv_sum_punit_coe {α} (a : α) : optionEquivSumPunit α a = Sum.inl a := + rfl +#align equiv.option_equiv_sum_punit_coe Equiv.option_equiv_sum_punit_coe + +@[simp] +theorem option_equiv_sum_punit_symm_inl {α} (a) : (optionEquivSumPunit α).symm (Sum.inl a) = a := + rfl +#align equiv.option_equiv_sum_punit_symm_inl Equiv.option_equiv_sum_punit_symm_inl + +@[simp] +theorem option_equiv_sum_punit_symm_inr {α} (a) : (optionEquivSumPunit α).symm (Sum.inr a) = none := + rfl +#align equiv.option_equiv_sum_punit_symm_inr Equiv.option_equiv_sum_punit_symm_inr + +/-- The set of `x : option α` such that `is_some x` is equivalent to `α`. -/ +@[simps] +def optionIsSomeEquiv (α : Type _) : { x : Option α // x.isSome } ≃ α where + toFun o := Option.get o.2 + invFun x := ⟨some x, by decide⟩ + left_inv o := Subtype.eq <| Option.some_get _ + right_inv x := Option.get_some _ _ +#align equiv.option_is_some_equiv Equiv.optionIsSomeEquiv + +/-- The product over `option α` of `β a` is the binary product of the +product over `α` of `β (some α)` and `β none` -/ +@[simps] +def piOptionEquivProd {α : Type _} {β : Option α → Type _} : (∀ a : Option α, β a) ≃ β none × ∀ a : α, β (some a) where + toFun f := (f none, fun a => f (some a)) + invFun x a := Option.casesOn a x.fst x.snd + left_inv f := funext fun a => by cases a <;> rfl + right_inv x := by simp +#align equiv.pi_option_equiv_prod Equiv.piOptionEquivProd + +/-- `α ⊕ β` is equivalent to a `sigma`-type over `bool`. Note that this definition assumes `α` and +`β` to be types from the same universe, so it cannot by used directly to transfer theorems about +sigma types to theorems about sum types. In many cases one can use `ulift` to work around this +difficulty. -/ +def sumEquivSigmaBool (α β : Type u) : Sum α β ≃ Σb : Bool, cond b α β := + ⟨fun s => s.elim (fun x => ⟨true, x⟩) fun x => ⟨false, x⟩, fun s => + match s with + | ⟨tt, a⟩ => inl a + | ⟨ff, b⟩ => inr b, + fun s => by cases s <;> rfl, fun s => by rcases s with ⟨_ | _, _⟩ <;> rfl⟩ +#align equiv.sum_equiv_sigma_bool Equiv.sumEquivSigmaBool + +-- See also `equiv.sigma_preimage_equiv`. +/-- `sigma_fiber_equiv f` for `f : α → β` is the natural equivalence between +the type of all fibres of `f` and the total space `α`. -/ +@[simps] +def sigmaFiberEquiv {α β : Type _} (f : α → β) : (Σy : β, { x // f x = y }) ≃ α := + ⟨fun x => ↑x.2, fun x => ⟨f x, x, rfl⟩, fun ⟨y, x, rfl⟩ => rfl, fun x => rfl⟩ +#align equiv.sigma_fiber_equiv Equiv.sigmaFiberEquiv + +end + +section SumCompl + +/-- For any predicate `p` on `α`, +the sum of the two subtypes `{a // p a}` and its complement `{a // ¬ p a}` +is naturally equivalent to `α`. + +See `subtype_or_equiv` for sum types over subtypes `{x // p x}` and `{x // q x}` +that are not necessarily `is_compl p q`. -/ +def sumCompl {α : Type _} (p : α → Prop) [DecidablePred p] : Sum { a // p a } { a // ¬p a } ≃ α where + toFun := Sum.elim coe coe + invFun a := if h : p a then Sum.inl ⟨a, h⟩ else Sum.inr ⟨a, h⟩ + left_inv := by rintro (⟨x, hx⟩ | ⟨x, hx⟩) <;> dsimp <;> [rw [dif_pos], rw [dif_neg]] + right_inv a := by + dsimp + split_ifs <;> rfl +#align equiv.sum_compl Equiv.sumCompl + +@[simp] +theorem sum_compl_apply_inl {α : Type _} (p : α → Prop) [DecidablePred p] (x : { a // p a }) : + sumCompl p (Sum.inl x) = x := + rfl +#align equiv.sum_compl_apply_inl Equiv.sum_compl_apply_inl + +@[simp] +theorem sum_compl_apply_inr {α : Type _} (p : α → Prop) [DecidablePred p] (x : { a // ¬p a }) : + sumCompl p (Sum.inr x) = x := + rfl +#align equiv.sum_compl_apply_inr Equiv.sum_compl_apply_inr + +@[simp] +theorem sum_compl_apply_symm_of_pos {α : Type _} (p : α → Prop) [DecidablePred p] (a : α) (h : p a) : + (sumCompl p).symm a = Sum.inl ⟨a, h⟩ := + dif_pos h +#align equiv.sum_compl_apply_symm_of_pos Equiv.sum_compl_apply_symm_of_pos + +@[simp] +theorem sum_compl_apply_symm_of_neg {α : Type _} (p : α → Prop) [DecidablePred p] (a : α) (h : ¬p a) : + (sumCompl p).symm a = Sum.inr ⟨a, h⟩ := + dif_neg h +#align equiv.sum_compl_apply_symm_of_neg Equiv.sum_compl_apply_symm_of_neg + +/-- Combines an `equiv` between two subtypes with an `equiv` between their complements to form a + permutation. -/ +def subtypeCongr {α : Type _} {p q : α → Prop} [DecidablePred p] [DecidablePred q] (e : { x // p x } ≃ { x // q x }) + (f : { x // ¬p x } ≃ { x // ¬q x }) : Perm α := + (sumCompl p).symm.trans ((sumCongr e f).trans (sumCompl q)) +#align equiv.subtype_congr Equiv.subtypeCongr + +open Equiv + +variable {ε : Type _} {p : ε → Prop} [DecidablePred p] + +variable (ep ep' : Perm { a // p a }) (en en' : Perm { a // ¬p a }) + +/-- Combining permutations on `ε` that permute only inside or outside the subtype +split induced by `p : ε → Prop` constructs a permutation on `ε`. -/ +def Perm.subtypeCongr : Equiv.Perm ε := + permCongr (sumCompl p) (sumCongr ep en) +#align equiv.perm.subtype_congr Equiv.Perm.subtypeCongr + +theorem Perm.subtypeCongr.apply (a : ε) : ep.subtypeCongr en a = if h : p a then ep ⟨a, h⟩ else en ⟨a, h⟩ := by + by_cases h:p a <;> simp [perm.subtype_congr, h] +#align equiv.perm.subtype_congr.apply Equiv.Perm.subtypeCongr.apply + +@[simp] +theorem Perm.subtypeCongr.left_apply {a : ε} (h : p a) : ep.subtypeCongr en a = ep ⟨a, h⟩ := by + simp [perm.subtype_congr.apply, h] +#align equiv.perm.subtype_congr.left_apply Equiv.Perm.subtypeCongr.left_apply + +@[simp] +theorem Perm.subtypeCongr.left_apply_subtype (a : { a // p a }) : ep.subtypeCongr en a = ep a := by + convert perm.subtype_congr.left_apply _ _ a.property + simp +#align equiv.perm.subtype_congr.left_apply_subtype Equiv.Perm.subtypeCongr.left_apply_subtype + +@[simp] +theorem Perm.subtypeCongr.right_apply {a : ε} (h : ¬p a) : ep.subtypeCongr en a = en ⟨a, h⟩ := by + simp [perm.subtype_congr.apply, h] +#align equiv.perm.subtype_congr.right_apply Equiv.Perm.subtypeCongr.right_apply + +@[simp] +theorem Perm.subtypeCongr.right_apply_subtype (a : { a // ¬p a }) : ep.subtypeCongr en a = en a := by + convert perm.subtype_congr.right_apply _ _ a.property + simp +#align equiv.perm.subtype_congr.right_apply_subtype Equiv.Perm.subtypeCongr.right_apply_subtype + +@[simp] +theorem Perm.subtypeCongr.refl : + Perm.subtypeCongr (Equiv.refl { a // p a }) (Equiv.refl { a // ¬p a }) = Equiv.refl ε := by + ext x + by_cases h:p x <;> simp [h] +#align equiv.perm.subtype_congr.refl Equiv.Perm.subtypeCongr.refl + +@[simp] +theorem Perm.subtypeCongr.symm : (ep.subtypeCongr en).symm = Perm.subtypeCongr ep.symm en.symm := by + ext x + by_cases h:p x + · have : p (ep.symm ⟨x, h⟩) := Subtype.property _ + simp [perm.subtype_congr.apply, h, symm_apply_eq, this] + + · have : ¬p (en.symm ⟨x, h⟩) := Subtype.property (en.symm _) + simp [perm.subtype_congr.apply, h, symm_apply_eq, this] + +#align equiv.perm.subtype_congr.symm Equiv.Perm.subtypeCongr.symm + +@[simp] +theorem Perm.subtypeCongr.trans : + (ep.subtypeCongr en).trans (ep'.subtypeCongr en') = Perm.subtypeCongr (ep.trans ep') (en.trans en') := by + ext x + by_cases h:p x + · have : p (ep ⟨x, h⟩) := Subtype.property _ + simp [perm.subtype_congr.apply, h, this] + + · have : ¬p (en ⟨x, h⟩) := Subtype.property (en _) + simp [perm.subtype_congr.apply, h, symm_apply_eq, this] + +#align equiv.perm.subtype_congr.trans Equiv.Perm.subtypeCongr.trans + +end SumCompl + +section SubtypePreimage + +variable (p : α → Prop) [DecidablePred p] (x₀ : { a // p a } → β) + +/-- For a fixed function `x₀ : {a // p a} → β` defined on a subtype of `α`, +the subtype of functions `x : α → β` that agree with `x₀` on the subtype `{a // p a}` +is naturally equivalent to the type of functions `{a // ¬ p a} → β`. -/ +@[simps] +def subtypePreimage : { x : α → β // x ∘ coe = x₀ } ≃ ({ a // ¬p a } → β) where + toFun (x : { x : α → β // x ∘ coe = x₀ }) a := (x : α → β) a + invFun x := ⟨fun a => if h : p a then x₀ ⟨a, h⟩ else x ⟨a, h⟩, funext fun ⟨a, h⟩ => dif_pos h⟩ + left_inv := fun ⟨x, hx⟩ => + Subtype.val_injective <| + funext fun a => by + dsimp + split_ifs <;> [rw [← hx], skip] <;> rfl + right_inv x := + funext fun ⟨a, h⟩ => + show dite (p a) _ _ = _ by + dsimp + rw [dif_neg h] +#align equiv.subtype_preimage Equiv.subtypePreimage + +theorem subtype_preimage_symm_apply_coe_pos (x : { a // ¬p a } → β) (a : α) (h : p a) : + ((subtypePreimage p x₀).symm x : α → β) a = x₀ ⟨a, h⟩ := + dif_pos h +#align equiv.subtype_preimage_symm_apply_coe_pos Equiv.subtype_preimage_symm_apply_coe_pos + +theorem subtype_preimage_symm_apply_coe_neg (x : { a // ¬p a } → β) (a : α) (h : ¬p a) : + ((subtypePreimage p x₀).symm x : α → β) a = x ⟨a, h⟩ := + dif_neg h +#align equiv.subtype_preimage_symm_apply_coe_neg Equiv.subtype_preimage_symm_apply_coe_neg + +end SubtypePreimage + +section + +/-- A family of equivalences `Π a, β₁ a ≃ β₂ a` generates an equivalence between `Π a, β₁ a` and +`Π a, β₂ a`. -/ +def piCongrRight {α} {β₁ β₂ : α → Sort _} (F : ∀ a, β₁ a ≃ β₂ a) : (∀ a, β₁ a) ≃ ∀ a, β₂ a := + ⟨fun H a => F a (H a), fun H a => (F a).symm (H a), fun H => funext <| by simp, fun H => funext <| by simp⟩ +#align equiv.Pi_congr_right Equiv.piCongrRight + +/-- Given `φ : α → β → Sort*`, we have an equivalence between `Π a b, φ a b` and `Π b a, φ a b`. +This is `function.swap` as an `equiv`. -/ +@[simps apply] +def piComm {α β} (φ : α → β → Sort _) : (∀ a b, φ a b) ≃ ∀ b a, φ a b := + ⟨swap, swap, fun x => rfl, fun y => rfl⟩ +#align equiv.Pi_comm Equiv.piComm + +@[simp] +theorem Pi_comm_symm {α β} {φ : α → β → Sort _} : (piComm φ).symm = (Pi_comm <| swap φ) := + rfl +#align equiv.Pi_comm_symm Equiv.Pi_comm_symm + +/-- Dependent `curry` equivalence: the type of dependent functions on `Σ i, β i` is equivalent +to the type of dependent functions of two arguments (i.e., functions to the space of functions). + +This is `sigma.curry` and `sigma.uncurry` together as an equiv. -/ +def piCurry {α} {β : α → Sort _} (γ : ∀ a, β a → Sort _) : (∀ x : Σi, β i, γ x.1 x.2) ≃ ∀ a b, γ a b where + toFun := Sigma.curry + invFun := Sigma.uncurry + left_inv := Sigma.uncurry_curry + right_inv := Sigma.curry_uncurry +#align equiv.Pi_curry Equiv.piCurry + +end + +section ProdCongr + +variable {α₁ β₁ β₂ : Type _} (e : α₁ → β₁ ≃ β₂) + +/-- A family of equivalences `Π (a : α₁), β₁ ≃ β₂` generates an equivalence +between `β₁ × α₁` and `β₂ × α₁`. -/ +def prodCongrLeft : β₁ × α₁ ≃ β₂ × α₁ where + toFun ab := ⟨e ab.2 ab.1, ab.2⟩ + invFun ab := ⟨(e ab.2).symm ab.1, ab.2⟩ + left_inv := by + rintro ⟨a, b⟩ + simp + right_inv := by + rintro ⟨a, b⟩ + simp +#align equiv.prod_congr_left Equiv.prodCongrLeft + +@[simp] +theorem prod_congr_left_apply (b : β₁) (a : α₁) : prodCongrLeft e (b, a) = (e a b, a) := + rfl +#align equiv.prod_congr_left_apply Equiv.prod_congr_left_apply + +theorem prod_congr_refl_right (e : β₁ ≃ β₂) : prodCongr e (Equiv.refl α₁) = prodCongrLeft fun _ => e := by + ext ⟨a, b⟩ : 1 + simp +#align equiv.prod_congr_refl_right Equiv.prod_congr_refl_right + +/-- A family of equivalences `Π (a : α₁), β₁ ≃ β₂` generates an equivalence +between `α₁ × β₁` and `α₁ × β₂`. -/ +def prodCongrRight : α₁ × β₁ ≃ α₁ × β₂ where + toFun ab := ⟨ab.1, e ab.1 ab.2⟩ + invFun ab := ⟨ab.1, (e ab.1).symm ab.2⟩ + left_inv := by + rintro ⟨a, b⟩ + simp + right_inv := by + rintro ⟨a, b⟩ + simp +#align equiv.prod_congr_right Equiv.prodCongrRight + +@[simp] +theorem prod_congr_right_apply (a : α₁) (b : β₁) : prodCongrRight e (a, b) = (a, e a b) := + rfl +#align equiv.prod_congr_right_apply Equiv.prod_congr_right_apply + +theorem prod_congr_refl_left (e : β₁ ≃ β₂) : prodCongr (Equiv.refl α₁) e = prodCongrRight fun _ => e := by + ext ⟨a, b⟩ : 1 + simp +#align equiv.prod_congr_refl_left Equiv.prod_congr_refl_left + +@[simp] +theorem prod_congr_left_trans_prod_comm : + (prodCongrLeft e).trans (prodComm _ _) = (prodComm _ _).trans (prodCongrRight e) := by + ext ⟨a, b⟩ : 1 + simp +#align equiv.prod_congr_left_trans_prod_comm Equiv.prod_congr_left_trans_prod_comm + +@[simp] +theorem prod_congr_right_trans_prod_comm : + (prodCongrRight e).trans (prodComm _ _) = (prodComm _ _).trans (prodCongrLeft e) := by + ext ⟨a, b⟩ : 1 + simp +#align equiv.prod_congr_right_trans_prod_comm Equiv.prod_congr_right_trans_prod_comm + +theorem sigma_congr_right_sigma_equiv_prod : + (sigmaCongrRight e).trans (sigmaEquivProd α₁ β₂) = (sigmaEquivProd α₁ β₁).trans (prodCongrRight e) := by + ext ⟨a, b⟩ : 1 + simp +#align equiv.sigma_congr_right_sigma_equiv_prod Equiv.sigma_congr_right_sigma_equiv_prod + +theorem sigma_equiv_prod_sigma_congr_right : + (sigmaEquivProd α₁ β₁).symm.trans (sigmaCongrRight e) = (prodCongrRight e).trans (sigmaEquivProd α₁ β₂).symm := by + ext ⟨a, b⟩ : 1 + simp +#align equiv.sigma_equiv_prod_sigma_congr_right Equiv.sigma_equiv_prod_sigma_congr_right + +-- See also `equiv.of_preimage_equiv`. +/-- A family of equivalences between fibers gives an equivalence between domains. -/ +@[simps] +def ofFiberEquiv {α β γ : Type _} {f : α → γ} {g : β → γ} (e : ∀ c, { a // f a = c } ≃ { b // g b = c }) : α ≃ β := + (sigmaFiberEquiv f).symm.trans <| (Equiv.sigmaCongrRight e).trans (sigmaFiberEquiv g) +#align equiv.of_fiber_equiv Equiv.ofFiberEquiv + +theorem of_fiber_equiv_map {α β γ} {f : α → γ} {g : β → γ} (e : ∀ c, { a // f a = c } ≃ { b // g b = c }) (a : α) : + g (ofFiberEquiv e a) = f a := + (_ : { b // g b = _ }).Prop +#align equiv.of_fiber_equiv_map Equiv.of_fiber_equiv_map + +/-- A variation on `equiv.prod_congr` where the equivalence in the second component can depend + on the first component. A typical example is a shear mapping, explaining the name of this + declaration. -/ +@[simps (config := { fullyApplied := false })] +def prodShear {α₁ β₁ α₂ β₂ : Type _} (e₁ : α₁ ≃ α₂) (e₂ : α₁ → β₁ ≃ β₂) : α₁ × β₁ ≃ α₂ × β₂ where + toFun := fun x : α₁ × β₁ => (e₁ x.1, e₂ x.1 x.2) + invFun := fun y : α₂ × β₂ => (e₁.symm y.1, (e₂ <| e₁.symm y.1).symm y.2) + left_inv := by + rintro ⟨x₁, y₁⟩ + simp only [symm_apply_apply] + right_inv := by + rintro ⟨x₁, y₁⟩ + simp only [apply_symm_apply] +#align equiv.prod_shear Equiv.prodShear + +end ProdCongr + +namespace Perm + +variable {α₁ β₁ β₂ : Type _} [DecidableEq α₁] (a : α₁) (e : Perm β₁) + +/-- `prod_extend_right a e` extends `e : perm β` to `perm (α × β)` by sending `(a, b)` to +`(a, e b)` and keeping the other `(a', b)` fixed. -/ +def prodExtendRight : Perm (α₁ × β₁) where + toFun ab := if ab.fst = a then (a, e ab.snd) else ab + invFun ab := if ab.fst = a then (a, e.symm ab.snd) else ab + left_inv := by + rintro ⟨k', x⟩ + dsimp only + split_ifs with h <;> simp [h] + right_inv := by + rintro ⟨k', x⟩ + dsimp only + split_ifs with h <;> simp [h] +#align equiv.perm.prod_extend_right Equiv.Perm.prodExtendRight + +@[simp] +theorem prod_extend_right_apply_eq (b : β₁) : prodExtendRight a e (a, b) = (a, e b) := + if_pos rfl +#align equiv.perm.prod_extend_right_apply_eq Equiv.Perm.prod_extend_right_apply_eq + +theorem prod_extend_right_apply_ne {a a' : α₁} (h : a' ≠ a) (b : β₁) : prodExtendRight a e (a', b) = (a', b) := + if_neg h +#align equiv.perm.prod_extend_right_apply_ne Equiv.Perm.prod_extend_right_apply_ne + +theorem eq_of_prod_extend_right_ne {e : Perm β₁} {a a' : α₁} {b : β₁} (h : prodExtendRight a e (a', b) ≠ (a', b)) : + a' = a := by + contrapose! h + exact prod_extend_right_apply_ne _ h _ +#align equiv.perm.eq_of_prod_extend_right_ne Equiv.Perm.eq_of_prod_extend_right_ne + +@[simp] +theorem fst_prod_extend_right (ab : α₁ × β₁) : (prodExtendRight a e ab).fst = ab.fst := by + rw [prod_extend_right, coe_fn_mk] + split_ifs with h + · rw [h] + + · rfl + +#align equiv.perm.fst_prod_extend_right Equiv.Perm.fst_prod_extend_right + +end Perm + +section + +/-- The type of functions to a product `α × β` is equivalent to the type of pairs of functions +`γ → α` and `γ → β`. -/ +def arrowProdEquivProdArrow (α β γ : Type _) : (γ → α × β) ≃ (γ → α) × (γ → β) := + ⟨fun f => (fun c => (f c).1, fun c => (f c).2), fun p c => (p.1 c, p.2 c), fun f => funext fun c => Prod.mk.eta, + fun p => by + cases p + rfl⟩ +#align equiv.arrow_prod_equiv_prod_arrow Equiv.arrowProdEquivProdArrow + +open Sum + +/-- The type of functions on a sum type `α ⊕ β` is equivalent to the type of pairs of functions +on `α` and on `β`. -/ +def sumArrowEquivProdArrow (α β γ : Type _) : (Sum α β → γ) ≃ (α → γ) × (β → γ) := + ⟨fun f => (f ∘ inl, f ∘ inr), fun p => Sum.elim p.1 p.2, fun f => by ext ⟨⟩ <;> rfl, fun p => by + cases p + rfl⟩ +#align equiv.sum_arrow_equiv_prod_arrow Equiv.sumArrowEquivProdArrow + +@[simp] +theorem sum_arrow_equiv_prod_arrow_apply_fst {α β γ} (f : Sum α β → γ) (a : α) : + (sumArrowEquivProdArrow α β γ f).1 a = f (inl a) := + rfl +#align equiv.sum_arrow_equiv_prod_arrow_apply_fst Equiv.sum_arrow_equiv_prod_arrow_apply_fst + +@[simp] +theorem sum_arrow_equiv_prod_arrow_apply_snd {α β γ} (f : Sum α β → γ) (b : β) : + (sumArrowEquivProdArrow α β γ f).2 b = f (inr b) := + rfl +#align equiv.sum_arrow_equiv_prod_arrow_apply_snd Equiv.sum_arrow_equiv_prod_arrow_apply_snd + +@[simp] +theorem sum_arrow_equiv_prod_arrow_symm_apply_inl {α β γ} (f : α → γ) (g : β → γ) (a : α) : + ((sumArrowEquivProdArrow α β γ).symm (f, g)) (inl a) = f a := + rfl +#align equiv.sum_arrow_equiv_prod_arrow_symm_apply_inl Equiv.sum_arrow_equiv_prod_arrow_symm_apply_inl + +@[simp] +theorem sum_arrow_equiv_prod_arrow_symm_apply_inr {α β γ} (f : α → γ) (g : β → γ) (b : β) : + ((sumArrowEquivProdArrow α β γ).symm (f, g)) (inr b) = g b := + rfl +#align equiv.sum_arrow_equiv_prod_arrow_symm_apply_inr Equiv.sum_arrow_equiv_prod_arrow_symm_apply_inr + +/-- Type product is right distributive with respect to type sum up to an equivalence. -/ +def sumProdDistrib (α β γ : Sort _) : Sum α β × γ ≃ Sum (α × γ) (β × γ) := + ⟨fun p => p.1.map (fun x => (x, p.2)) fun x => (x, p.2), fun s => s.elim (Prod.map inl id) (Prod.map inr id), by + rintro ⟨_ | _, _⟩ <;> rfl, by rintro (⟨_, _⟩ | ⟨_, _⟩) <;> rfl⟩ +#align equiv.sum_prod_distrib Equiv.sumProdDistrib + +@[simp] +theorem sum_prod_distrib_apply_left {α β γ} (a : α) (c : γ) : sumProdDistrib α β γ (Sum.inl a, c) = Sum.inl (a, c) := + rfl +#align equiv.sum_prod_distrib_apply_left Equiv.sum_prod_distrib_apply_left + +@[simp] +theorem sum_prod_distrib_apply_right {α β γ} (b : β) (c : γ) : sumProdDistrib α β γ (Sum.inr b, c) = Sum.inr (b, c) := + rfl +#align equiv.sum_prod_distrib_apply_right Equiv.sum_prod_distrib_apply_right + +@[simp] +theorem sum_prod_distrib_symm_apply_left {α β γ} (a : α × γ) : (sumProdDistrib α β γ).symm (inl a) = (inl a.1, a.2) := + rfl +#align equiv.sum_prod_distrib_symm_apply_left Equiv.sum_prod_distrib_symm_apply_left + +@[simp] +theorem sum_prod_distrib_symm_apply_right {α β γ} (b : β × γ) : (sumProdDistrib α β γ).symm (inr b) = (inr b.1, b.2) := + rfl +#align equiv.sum_prod_distrib_symm_apply_right Equiv.sum_prod_distrib_symm_apply_right + +/-- Type product is left distributive with respect to type sum up to an equivalence. -/ +def prodSumDistrib (α β γ : Sort _) : α × Sum β γ ≃ Sum (α × β) (α × γ) := + calc + α × Sum β γ ≃ Sum β γ × α := prodComm _ _ + _ ≃ Sum (β × α) (γ × α) := sumProdDistrib _ _ _ + _ ≃ Sum (α × β) (α × γ) := sumCongr (prodComm _ _) (prodComm _ _) + +#align equiv.prod_sum_distrib Equiv.prodSumDistrib + +@[simp] +theorem prod_sum_distrib_apply_left {α β γ} (a : α) (b : β) : prodSumDistrib α β γ (a, Sum.inl b) = Sum.inl (a, b) := + rfl +#align equiv.prod_sum_distrib_apply_left Equiv.prod_sum_distrib_apply_left + +@[simp] +theorem prod_sum_distrib_apply_right {α β γ} (a : α) (c : γ) : prodSumDistrib α β γ (a, Sum.inr c) = Sum.inr (a, c) := + rfl +#align equiv.prod_sum_distrib_apply_right Equiv.prod_sum_distrib_apply_right + +@[simp] +theorem prod_sum_distrib_symm_apply_left {α β γ} (a : α × β) : (prodSumDistrib α β γ).symm (inl a) = (a.1, inl a.2) := + rfl +#align equiv.prod_sum_distrib_symm_apply_left Equiv.prod_sum_distrib_symm_apply_left + +@[simp] +theorem prod_sum_distrib_symm_apply_right {α β γ} (a : α × γ) : (prodSumDistrib α β γ).symm (inr a) = (a.1, inr a.2) := + rfl +#align equiv.prod_sum_distrib_symm_apply_right Equiv.prod_sum_distrib_symm_apply_right + +/-- An indexed sum of disjoint sums of types is equivalent to the sum of the indexed sums. -/ +@[simps] +def sigmaSumDistrib {ι : Type _} (α β : ι → Type _) : (Σi, Sum (α i) (β i)) ≃ Sum (Σi, α i) (Σi, β i) := + ⟨fun p => p.2.map (Sigma.mk p.1) (Sigma.mk p.1), + Sum.elim (Sigma.map id fun _ => Sum.inl) (Sigma.map id fun _ => Sum.inr), fun p => by + rcases p with ⟨i, a | b⟩ <;> rfl, fun p => by rcases p with (⟨i, a⟩ | ⟨i, b⟩) <;> rfl⟩ +#align equiv.sigma_sum_distrib Equiv.sigmaSumDistrib + +/-- The product of an indexed sum of types (formally, a `sigma`-type `Σ i, α i`) by a type `β` is +equivalent to the sum of products `Σ i, (α i × β)`. -/ +def sigmaProdDistrib {ι : Type _} (α : ι → Type _) (β : Type _) : (Σi, α i) × β ≃ Σi, α i × β := + ⟨fun p => ⟨p.1.1, (p.1.2, p.2)⟩, fun p => (⟨p.1, p.2.1⟩, p.2.2), fun p => by + rcases p with ⟨⟨_, _⟩, _⟩ + rfl, fun p => by + rcases p with ⟨_, ⟨_, _⟩⟩ + rfl⟩ +#align equiv.sigma_prod_distrib Equiv.sigmaProdDistrib + +/-- An equivalence that separates out the 0th fiber of `(Σ (n : ℕ), f n)`. -/ +def sigmaNatSucc (f : ℕ → Type u) : (Σn, f n) ≃ Sum (f 0) (Σn, f (n + 1)) := + ⟨fun x => + @Sigma.casesOn ℕ f (fun _ => Sum (f 0) (Σn, f (n + 1))) x fun n => + @Nat.casesOn (fun i => f i → Sum (f 0) (Σn : ℕ, f (n + 1))) n (fun x : f 0 => Sum.inl x) + fun (n : ℕ) (x : f n.succ) => Sum.inr ⟨n, x⟩, + Sum.elim (Sigma.mk 0) (Sigma.map Nat.succ fun _ => id), by rintro ⟨n | n, x⟩ <;> rfl, by + rintro (x | ⟨n, x⟩) <;> rfl⟩ +#align equiv.sigma_nat_succ Equiv.sigmaNatSucc + +/-- The product `bool × α` is equivalent to `α ⊕ α`. -/ +@[simps] +def boolProdEquivSum (α : Type u) : Bool × α ≃ Sum α α where + toFun p := cond p.1 (inr p.2) (inl p.2) + invFun := Sum.elim (Prod.mk false) (Prod.mk true) + left_inv := by rintro ⟨_ | _, _⟩ <;> rfl + right_inv := by rintro (_ | _) <;> rfl +#align equiv.bool_prod_equiv_sum Equiv.boolProdEquivSum + +/-- The function type `bool → α` is equivalent to `α × α`. -/ +@[simps] +def boolArrowEquivProd (α : Type u) : (Bool → α) ≃ α × α where + toFun f := (f true, f false) + invFun p b := cond b p.1 p.2 + left_inv f := funext <| Bool.forall_bool.2 ⟨rfl, rfl⟩ + right_inv := fun ⟨x, y⟩ => rfl +#align equiv.bool_arrow_equiv_prod Equiv.boolArrowEquivProd + +end + +section + +open Sum Nat + +/-- The set of natural numbers is equivalent to `ℕ ⊕ punit`. -/ +def natEquivNatSumPunit : ℕ ≃ Sum ℕ PUnit.{u + 1} where + toFun n := Nat.casesOn n (inr PUnit.unit) inl + invFun := Sum.elim Nat.succ fun _ => 0 + left_inv n := by cases n <;> rfl + right_inv := by rintro (_ | _ | _) <;> rfl +#align equiv.nat_equiv_nat_sum_punit Equiv.natEquivNatSumPunit + +/-- `ℕ ⊕ punit` is equivalent to `ℕ`. -/ +def natSumPunitEquivNat : Sum ℕ PUnit.{u + 1} ≃ ℕ := + natEquivNatSumPunit.symm +#align equiv.nat_sum_punit_equiv_nat Equiv.natSumPunitEquivNat + +/-- The type of integer numbers is equivalent to `ℕ ⊕ ℕ`. -/ +def intEquivNatSumNat : ℤ ≃ Sum ℕ ℕ where + toFun z := Int.casesOn z inl inr + invFun := Sum.elim coe Int.negSucc + left_inv := by rintro (m | n) <;> rfl + right_inv := by rintro (m | n) <;> rfl +#align equiv.int_equiv_nat_sum_nat Equiv.intEquivNatSumNat + +end + +/-- An equivalence between `α` and `β` generates an equivalence between `list α` and `list β`. -/ +def listEquivOfEquiv {α β : Type _} (e : α ≃ β) : List α ≃ List β where + toFun := List.map e + invFun := List.map e.symm + left_inv l := by rw [List.map_map, e.symm_comp_self, List.map_id] + right_inv l := by rw [List.map_map, e.self_comp_symm, List.map_id] +#align equiv.list_equiv_of_equiv Equiv.listEquivOfEquiv + +/-- If `α` is equivalent to `β`, then `unique α` is equivalent to `unique β`. -/ +def uniqueCongr (e : α ≃ β) : Unique α ≃ Unique β where + toFun h := @Equiv.unique _ _ h e.symm + invFun h := @Equiv.unique _ _ h e + left_inv _ := Subsingleton.elim _ _ + right_inv _ := Subsingleton.elim _ _ +#align equiv.unique_congr Equiv.uniqueCongr + +/-- If `α` is equivalent to `β`, then `is_empty α` is equivalent to `is_empty β`. -/ +theorem is_empty_congr (e : α ≃ β) : IsEmpty α ↔ IsEmpty β := + ⟨fun h => @function.isEmpty _ _ h e.symm, fun h => @function.isEmpty _ _ h e⟩ +#align equiv.is_empty_congr Equiv.is_empty_congr + +protected theorem is_empty (e : α ≃ β) [IsEmpty β] : IsEmpty α := + e.is_empty_congr.mpr ‹_› +#align equiv.is_empty Equiv.is_empty + +section + +open Subtype + +/-- If `α` is equivalent to `β` and the predicates `p : α → Prop` and `q : β → Prop` are equivalent +at corresponding points, then `{a // p a}` is equivalent to `{b // q b}`. +For the statement where `α = β`, that is, `e : perm α`, see `perm.subtype_perm`. -/ +def subtypeEquiv {p : α → Prop} {q : β → Prop} (e : α ≃ β) (h : ∀ a, p a ↔ q (e a)) : + { a : α // p a } ≃ { b : β // q b } where + toFun a := ⟨e a, (h _).mp a.Prop⟩ + invFun b := ⟨e.symm b, (h _).mpr ((e.apply_symm_apply b).symm ▸ b.Prop)⟩ + left_inv a := Subtype.ext <| by simp + right_inv b := Subtype.ext <| by simp +#align equiv.subtype_equiv Equiv.subtypeEquiv + +@[simp] +theorem subtype_equiv_refl {p : α → Prop} (h : ∀ a, p a ↔ p (Equiv.refl _ a) := fun a => Iff.rfl) : + (Equiv.refl α).subtypeEquiv h = Equiv.refl { a : α // p a } := by + ext + rfl +#align equiv.subtype_equiv_refl Equiv.subtype_equiv_refl + +@[simp] +theorem subtype_equiv_symm {p : α → Prop} {q : β → Prop} (e : α ≃ β) (h : ∀ a : α, p a ↔ q (e a)) : + (e.subtypeEquiv h).symm = + e.symm.subtypeEquiv fun a => by + convert (h <| e.symm a).symm + exact (e.apply_symm_apply a).symm := + rfl +#align equiv.subtype_equiv_symm Equiv.subtype_equiv_symm + +@[simp] +theorem subtype_equiv_trans {p : α → Prop} {q : β → Prop} {r : γ → Prop} (e : α ≃ β) (f : β ≃ γ) + (h : ∀ a : α, p a ↔ q (e a)) (h' : ∀ b : β, q b ↔ r (f b)) : + (e.subtypeEquiv h).trans (f.subtypeEquiv h') = (e.trans f).subtypeEquiv fun a => (h a).trans (h' <| e a) := + rfl +#align equiv.subtype_equiv_trans Equiv.subtype_equiv_trans + +@[simp] +theorem subtype_equiv_apply {p : α → Prop} {q : β → Prop} (e : α ≃ β) (h : ∀ a : α, p a ↔ q (e a)) (x : { x // p x }) : + e.subtypeEquiv h x = ⟨e x, (h _).1 x.2⟩ := + rfl +#align equiv.subtype_equiv_apply Equiv.subtype_equiv_apply + +/-- If two predicates `p` and `q` are pointwise equivalent, then `{x // p x}` is equivalent to +`{x // q x}`. -/ +@[simps] +def subtypeEquivRight {p q : α → Prop} (e : ∀ x, p x ↔ q x) : { x // p x } ≃ { x // q x } := + subtypeEquiv (Equiv.refl _) e +#align equiv.subtype_equiv_right Equiv.subtypeEquivRight + +/-- If `α ≃ β`, then for any predicate `p : β → Prop` the subtype `{a // p (e a)}` is equivalent +to the subtype `{b // p b}`. -/ +def subtypeEquivOfSubtype {p : β → Prop} (e : α ≃ β) : { a : α // p (e a) } ≃ { b : β // p b } := + subtypeEquiv e <| by simp +#align equiv.subtype_equiv_of_subtype Equiv.subtypeEquivOfSubtype + +/-- If `α ≃ β`, then for any predicate `p : α → Prop` the subtype `{a // p a}` is equivalent +to the subtype `{b // p (e.symm b)}`. This version is used by `equiv_rw`. -/ +def subtypeEquivOfSubtype' {p : α → Prop} (e : α ≃ β) : { a : α // p a } ≃ { b : β // p (e.symm b) } := + e.symm.subtypeEquivOfSubtype.symm +#align equiv.subtype_equiv_of_subtype' Equiv.subtypeEquivOfSubtype' + +/-- If two predicates are equal, then the corresponding subtypes are equivalent. -/ +def subtypeEquivProp {α : Type _} {p q : α → Prop} (h : p = q) : Subtype p ≃ Subtype q := + subtypeEquiv (Equiv.refl α) fun a => h ▸ Iff.rfl +#align equiv.subtype_equiv_prop Equiv.subtypeEquivProp + +/-- A subtype of a subtype is equivalent to the subtype of elements satisfying both predicates. This +version allows the “inner” predicate to depend on `h : p a`. -/ +@[simps] +def subtypeSubtypeEquivSubtypeExists {α : Type u} (p : α → Prop) (q : Subtype p → Prop) : + Subtype q ≃ { a : α // ∃ h : p a, q ⟨a, h⟩ } := + ⟨fun a => + ⟨a, a.1.2, by + rcases a with ⟨⟨a, hap⟩, haq⟩ + exact haq⟩, + fun a => ⟨⟨a, a.2.fst⟩, a.2.snd⟩, fun ⟨⟨a, ha⟩, h⟩ => rfl, fun ⟨a, h₁, h₂⟩ => rfl⟩ +#align equiv.subtype_subtype_equiv_subtype_exists Equiv.subtypeSubtypeEquivSubtypeExists + +/-- A subtype of a subtype is equivalent to the subtype of elements satisfying both predicates. -/ +@[simps] +def subtypeSubtypeEquivSubtypeInter {α : Type u} (p q : α → Prop) : + { x : Subtype p // q x.1 } ≃ Subtype fun x => p x ∧ q x := + (subtypeSubtypeEquivSubtypeExists p _).trans <| subtype_equiv_right fun x => exists_prop +#align equiv.subtype_subtype_equiv_subtype_inter Equiv.subtypeSubtypeEquivSubtypeInter + +/-- If the outer subtype has more restrictive predicate than the inner one, +then we can drop the latter. -/ +@[simps] +def subtypeSubtypeEquivSubtype {α : Type u} {p q : α → Prop} (h : ∀ {x}, q x → p x) : + { x : Subtype p // q x.1 } ≃ Subtype q := + (subtypeSubtypeEquivSubtypeInter p _).trans <| subtype_equiv_right fun x => and_iff_right_of_imp h +#align equiv.subtype_subtype_equiv_subtype Equiv.subtypeSubtypeEquivSubtype + +/-- If a proposition holds for all elements, then the subtype is +equivalent to the original type. -/ +@[simps apply symmApply] +def subtypeUnivEquiv {α : Type u} {p : α → Prop} (h : ∀ x, p x) : Subtype p ≃ α := + ⟨fun x => x, fun x => ⟨x, h x⟩, fun x => Subtype.eq rfl, fun x => rfl⟩ +#align equiv.subtype_univ_equiv Equiv.subtypeUnivEquiv + +/-- A subtype of a sigma-type is a sigma-type over a subtype. -/ +def subtypeSigmaEquiv {α : Type u} (p : α → Type v) (q : α → Prop) : { y : Sigma p // q y.1 } ≃ Σx : Subtype q, p x.1 := + ⟨fun x => ⟨⟨x.1.1, x.2⟩, x.1.2⟩, fun x => ⟨⟨x.1.1, x.2⟩, x.1.2⟩, fun ⟨⟨x, h⟩, y⟩ => rfl, fun ⟨⟨x, y⟩, h⟩ => rfl⟩ +#align equiv.subtype_sigma_equiv Equiv.subtypeSigmaEquiv + +/-- A sigma type over a subtype is equivalent to the sigma set over the original type, +if the fiber is empty outside of the subset -/ +def sigmaSubtypeEquivOfSubset {α : Type u} (p : α → Type v) (q : α → Prop) (h : ∀ x, p x → q x) : + (Σx : Subtype q, p x) ≃ Σx : α, p x := + (subtypeSigmaEquiv p q).symm.trans <| subtype_univ_equiv fun x => h x.1 x.2 +#align equiv.sigma_subtype_equiv_of_subset Equiv.sigmaSubtypeEquivOfSubset + +/-- If a predicate `p : β → Prop` is true on the range of a map `f : α → β`, then +`Σ y : {y // p y}, {x // f x = y}` is equivalent to `α`. -/ +def sigmaSubtypeFiberEquiv {α : Type u} {β : Type v} (f : α → β) (p : β → Prop) (h : ∀ x, p (f x)) : + (Σy : Subtype p, { x : α // f x = y }) ≃ α := + calc + _ ≃ Σy : β, { x : α // f x = y } := sigmaSubtypeEquivOfSubset _ p fun y ⟨x, h'⟩ => h' ▸ h x + _ ≃ α := sigmaFiberEquiv f + +#align equiv.sigma_subtype_fiber_equiv Equiv.sigmaSubtypeFiberEquiv + +/-- If for each `x` we have `p x ↔ q (f x)`, then `Σ y : {y // q y}, f ⁻¹' {y}` is equivalent +to `{x // p x}`. -/ +def sigmaSubtypeFiberEquivSubtype {α : Type u} {β : Type v} (f : α → β) {p : α → Prop} {q : β → Prop} + (h : ∀ x, p x ↔ q (f x)) : (Σy : Subtype q, { x : α // f x = y }) ≃ Subtype p := + calc + (Σy : Subtype q, { x : α // f x = y }) ≃ Σy : Subtype q, { x : Subtype p // Subtype.mk (f x) ((h x).1 x.2) = y } := + by + apply sigma_congr_right + intro y + symm + refine' (subtype_subtype_equiv_subtype_exists _ _).trans (subtype_equiv_right _) + intro x + exact ⟨fun ⟨hp, h'⟩ => congr_arg Subtype.val h', fun h' => ⟨(h x).2 (h'.symm ▸ y.2), Subtype.eq h'⟩⟩ + _ ≃ Subtype p := sigmaFiberEquiv fun x : Subtype p => (⟨f x, (h x).1 x.property⟩ : Subtype q) + +#align equiv.sigma_subtype_fiber_equiv_subtype Equiv.sigmaSubtypeFiberEquivSubtype + +/-- A sigma type over an `option` is equivalent to the sigma set over the original type, +if the fiber is empty at none. -/ +def sigmaOptionEquivOfSome {α : Type u} (p : Option α → Type v) (h : p none → False) : + (Σx : Option α, p x) ≃ Σx : α, p (some x) := + haveI h' : ∀ x, p x → x.isSome := by + intro x + cases x + · intro n + exfalso + exact h n + + · intro s + exact rfl + + (sigma_subtype_equiv_of_subset _ _ h').symm.trans (sigma_congr_left' (option_is_some_equiv α)) +#align equiv.sigma_option_equiv_of_some Equiv.sigmaOptionEquivOfSome + +/-- The `pi`-type `Π i, π i` is equivalent to the type of sections `f : ι → Σ i, π i` of the +`sigma` type such that for all `i` we have `(f i).fst = i`. -/ +def piEquivSubtypeSigma (ι : Type _) (π : ι → Type _) : (∀ i, π i) ≃ { f : ι → Σi, π i // ∀ i, (f i).1 = i } := + ⟨fun f => ⟨fun i => ⟨i, f i⟩, fun i => rfl⟩, fun f i => by + rw [← f.2 i] + exact (f.1 i).2, fun f => funext fun i => rfl, fun ⟨f, hf⟩ => + Subtype.eq <| + funext fun i => Sigma.eq (hf i).symm <| eq_of_heq <| rec_heq_of_heq _ <| rec_heq_of_heq _ <| HEq.refl _⟩ +#align equiv.pi_equiv_subtype_sigma Equiv.piEquivSubtypeSigma + +/-- The set of functions `f : Π a, β a` such that for all `a` we have `p a (f a)` is equivalent +to the set of functions `Π a, {b : β a // p a b}`. -/ +def subtypePiEquivPi {α : Sort u} {β : α → Sort v} {p : ∀ a, β a → Prop} : + { f : ∀ a, β a // ∀ a, p a (f a) } ≃ ∀ a, { b : β a // p a b } := + ⟨fun f a => ⟨f.1 a, f.2 a⟩, fun f => ⟨fun a => (f a).1, fun a => (f a).2⟩, by + rintro ⟨f, h⟩ + rfl, by + rintro f + funext a + exact Subtype.ext_val rfl⟩ +#align equiv.subtype_pi_equiv_pi Equiv.subtypePiEquivPi + +/-- A subtype of a product defined by componentwise conditions +is equivalent to a product of subtypes. -/ +def subtypeProdEquivProd {α : Type u} {β : Type v} {p : α → Prop} {q : β → Prop} : + { c : α × β // p c.1 ∧ q c.2 } ≃ { a // p a } × { b // q b } := + ⟨fun x => ⟨⟨x.1.1, x.2.1⟩, ⟨x.1.2, x.2.2⟩⟩, fun x => ⟨⟨x.1.1, x.2.1⟩, ⟨x.1.2, x.2.2⟩⟩, fun ⟨⟨_, _⟩, ⟨_, _⟩⟩ => rfl, + fun ⟨⟨_, _⟩, ⟨_, _⟩⟩ => rfl⟩ +#align equiv.subtype_prod_equiv_prod Equiv.subtypeProdEquivProd + +/-- A subtype of a `prod` is equivalent to a sigma type whose fibers are subtypes. -/ +def subtypeProdEquivSigmaSubtype {α β : Type _} (p : α → β → Prop) : + { x : α × β // p x.1 x.2 } ≃ Σa, { b : β // p a b } where + toFun x := ⟨x.1.1, x.1.2, x.Prop⟩ + invFun x := ⟨⟨x.1, x.2⟩, x.2.Prop⟩ + left_inv x := by ext <;> rfl + right_inv := fun ⟨a, b, pab⟩ => rfl +#align equiv.subtype_prod_equiv_sigma_subtype Equiv.subtypeProdEquivSigmaSubtype + +/-- The type `Π (i : α), β i` can be split as a product by separating the indices in `α` +depending on whether they satisfy a predicate `p` or not. -/ +@[simps] +def piEquivPiSubtypeProd {α : Type _} (p : α → Prop) (β : α → Type _) [DecidablePred p] : + (∀ i : α, β i) ≃ (∀ i : { x // p x }, β i) × ∀ i : { x // ¬p x }, β i where + toFun f := (fun x => f x, fun x => f x) + invFun f x := if h : p x then f.1 ⟨x, h⟩ else f.2 ⟨x, h⟩ + right_inv := by + rintro ⟨f, g⟩ + ext1 <;> + · ext y + rcases y with ⟨⟩ + simp only [y_property, dif_pos, dif_neg, not_false_iff, Subtype.coe_mk] + rfl + + left_inv f := by + ext x + by_cases h:p x <;> + · simp only [h, dif_neg, dif_pos, not_false_iff] + rfl + +#align equiv.pi_equiv_pi_subtype_prod Equiv.piEquivPiSubtypeProd + +/-- A product of types can be split as the binary product of one of the types and the product + of all the remaining types. -/ +@[simps] +def piSplitAt {α : Type _} [DecidableEq α] (i : α) (β : α → Type _) : (∀ j, β j) ≃ β i × ∀ j : { j // j ≠ i }, β j where + toFun f := ⟨f i, fun j => f j⟩ + invFun f j := if h : j = i then h.symm.rec f.1 else f.2 ⟨j, h⟩ + right_inv f := by + ext + exacts[dif_pos rfl, (dif_neg x.2).trans (by cases x <;> rfl)] + left_inv f := by + ext + dsimp only + split_ifs + · subst h + + · rfl + +#align equiv.pi_split_at Equiv.piSplitAt + +/-- A product of copies of a type can be split as the binary product of one copy and the product + of all the remaining copies. -/ +@[simps] +def funSplitAt {α : Type _} [DecidableEq α] (i : α) (β : Type _) : (α → β) ≃ β × ({ j // j ≠ i } → β) := + piSplitAt i _ +#align equiv.fun_split_at Equiv.funSplitAt + +end + +section SubtypeEquivCodomain + +variable {X : Type _} {Y : Type _} [DecidableEq X] {x : X} + +/-- The type of all functions `X → Y` with prescribed values for all `x' ≠ x` +is equivalent to the codomain `Y`. -/ +def subtypeEquivCodomain (f : { x' // x' ≠ x } → Y) : { g : X → Y // g ∘ coe = f } ≃ Y := + (subtypePreimage _ f).trans <| + @funUnique { x' // ¬x' ≠ x } _ <| + show Unique { x' // ¬x' ≠ x } from + @Equiv.unique _ _ + (show Unique { x' // x' = x } from { default := ⟨x, rfl⟩, uniq := fun ⟨x', h⟩ => Subtype.val_injective h }) + (subtype_equiv_right fun a => not_not) +#align equiv.subtype_equiv_codomain Equiv.subtypeEquivCodomain + +@[simp] +theorem coe_subtype_equiv_codomain (f : { x' // x' ≠ x } → Y) : + (subtypeEquivCodomain f : { g : X → Y // g ∘ coe = f } → Y) = fun g => (g : X → Y) x := + rfl +#align equiv.coe_subtype_equiv_codomain Equiv.coe_subtype_equiv_codomain + +@[simp] +theorem subtype_equiv_codomain_apply (f : { x' // x' ≠ x } → Y) (g : { g : X → Y // g ∘ coe = f }) : + subtypeEquivCodomain f g = (g : X → Y) x := + rfl +#align equiv.subtype_equiv_codomain_apply Equiv.subtype_equiv_codomain_apply + +theorem coe_subtype_equiv_codomain_symm (f : { x' // x' ≠ x } → Y) : + ((subtypeEquivCodomain f).symm : Y → { g : X → Y // g ∘ coe = f }) = fun y => + ⟨fun x' => if h : x' ≠ x then f ⟨x', h⟩ else y, by + funext x' + dsimp + erw [dif_pos x'.2, Subtype.coe_eta]⟩ := + rfl +#align equiv.coe_subtype_equiv_codomain_symm Equiv.coe_subtype_equiv_codomain_symm + +@[simp] +theorem subtype_equiv_codomain_symm_apply (f : { x' // x' ≠ x } → Y) (y : Y) (x' : X) : + ((subtypeEquivCodomain f).symm y : X → Y) x' = if h : x' ≠ x then f ⟨x', h⟩ else y := + rfl +#align equiv.subtype_equiv_codomain_symm_apply Equiv.subtype_equiv_codomain_symm_apply + +@[simp] +theorem subtype_equiv_codomain_symm_apply_eq (f : { x' // x' ≠ x } → Y) (y : Y) : + ((subtypeEquivCodomain f).symm y : X → Y) x = y := + dif_neg (not_not.mpr rfl) +#align equiv.subtype_equiv_codomain_symm_apply_eq Equiv.subtype_equiv_codomain_symm_apply_eq + +theorem subtype_equiv_codomain_symm_apply_ne (f : { x' // x' ≠ x } → Y) (y : Y) (x' : X) (h : x' ≠ x) : + ((subtypeEquivCodomain f).symm y : X → Y) x' = f ⟨x', h⟩ := + dif_pos h +#align equiv.subtype_equiv_codomain_symm_apply_ne Equiv.subtype_equiv_codomain_symm_apply_ne + +end SubtypeEquivCodomain + +/-- If `f` is a bijective function, then its domain is equivalent to its codomain. -/ +@[simps apply] +noncomputable def ofBijective (f : α → β) (hf : Bijective f) : α ≃ β where + toFun := f + invFun := Function.surjInv hf.Surjective + left_inv := Function.left_inverse_surj_inv hf + right_inv := Function.right_inverse_surj_inv _ +#align equiv.of_bijective Equiv.ofBijective + +theorem of_bijective_apply_symm_apply (f : α → β) (hf : Bijective f) (x : β) : f ((ofBijective f hf).symm x) = x := + (ofBijective f hf).apply_symm_apply x +#align equiv.of_bijective_apply_symm_apply Equiv.of_bijective_apply_symm_apply + +@[simp] +theorem of_bijective_symm_apply_apply (f : α → β) (hf : Bijective f) (x : α) : (ofBijective f hf).symm (f x) = x := + (ofBijective f hf).symm_apply_apply x +#align equiv.of_bijective_symm_apply_apply Equiv.of_bijective_symm_apply_apply + +instance : CanLift (α → β) (α ≃ β) coeFn Bijective where prf f hf := ⟨ofBijective f hf, rfl⟩ + +section + +variable {α' β' : Type _} (e : Perm α') {p : β' → Prop} [DecidablePred p] (f : α' ≃ Subtype p) + +/-- Extend the domain of `e : equiv.perm α` to one that is over `β` via `f : α → subtype p`, +where `p : β → Prop`, permuting only the `b : β` that satisfy `p b`. +This can be used to extend the domain across a function `f : α → β`, +keeping everything outside of `set.range f` fixed. For this use-case `equiv` given by `f` can +be constructed by `equiv.of_left_inverse'` or `equiv.of_left_inverse` when there is a known +inverse, or `equiv.of_injective` in the general case.`. +-/ +def Perm.extendDomain : Perm β' := + (permCongr f e).subtypeCongr (Equiv.refl _) +#align equiv.perm.extend_domain Equiv.Perm.extendDomain + +@[simp] +theorem Perm.extend_domain_apply_image (a : α') : e.extendDomain f (f a) = f (e a) := by simp [perm.extend_domain] +#align equiv.perm.extend_domain_apply_image Equiv.Perm.extend_domain_apply_image + +theorem Perm.extend_domain_apply_subtype {b : β'} (h : p b) : e.extendDomain f b = f (e (f.symm ⟨b, h⟩)) := by + simp [perm.extend_domain, h] +#align equiv.perm.extend_domain_apply_subtype Equiv.Perm.extend_domain_apply_subtype + +theorem Perm.extend_domain_apply_not_subtype {b : β'} (h : ¬p b) : e.extendDomain f b = b := by + simp [perm.extend_domain, h] +#align equiv.perm.extend_domain_apply_not_subtype Equiv.Perm.extend_domain_apply_not_subtype + +@[simp] +theorem Perm.extend_domain_refl : Perm.extendDomain (Equiv.refl _) f = Equiv.refl _ := by simp [perm.extend_domain] +#align equiv.perm.extend_domain_refl Equiv.Perm.extend_domain_refl + +@[simp] +theorem Perm.extend_domain_symm : (e.extendDomain f).symm = Perm.extendDomain e.symm f := + rfl +#align equiv.perm.extend_domain_symm Equiv.Perm.extend_domain_symm + +theorem Perm.extend_domain_trans (e e' : Perm α') : + (e.extendDomain f).trans (e'.extendDomain f) = Perm.extendDomain (e.trans e') f := by + simp [perm.extend_domain, perm_congr_trans] +#align equiv.perm.extend_domain_trans Equiv.Perm.extend_domain_trans + +end + +/-- Subtype of the quotient is equivalent to the quotient of the subtype. Let `α` be a setoid with +equivalence relation `~`. Let `p₂` be a predicate on the quotient type `α/~`, and `p₁` be the lift +of this predicate to `α`: `p₁ a ↔ p₂ ⟦a⟧`. Let `~₂` be the restriction of `~` to `{x // p₁ x}`. +Then `{x // p₂ x}` is equivalent to the quotient of `{x // p₁ x}` by `~₂`. -/ +def subtypeQuotientEquivQuotientSubtype (p₁ : α → Prop) [s₁ : Setoid α] [s₂ : Setoid (Subtype p₁)] + (p₂ : Quotient s₁ → Prop) (hp₂ : ∀ a, p₁ a ↔ p₂ ⟦a⟧) (h : ∀ x y : Subtype p₁, @Setoid.r _ s₂ x y ↔ (x : α) ≈ y) : + { x // p₂ x } ≃ Quotient s₂ where + toFun a := + Quotient.hrecOn a.1 (fun a h => ⟦⟨a, (hp₂ _).2 h⟩⟧) + (fun a b hab => hfunext (by rw [Quotient.sound hab]) fun h₁ h₂ _ => heq_of_eq (Quotient.sound ((h _ _).2 hab))) + a.2 + invFun a := + Quotient.liftOn a (fun a => (⟨⟦a.1⟧, (hp₂ _).1 a.2⟩ : { x // p₂ x })) fun a b hab => + Subtype.ext_val (Quotient.sound ((h _ _).1 hab)) + left_inv := fun ⟨a, ha⟩ => Quotient.induction_on a (fun a ha => rfl) ha + right_inv a := Quotient.induction_on a fun ⟨a, ha⟩ => rfl +#align equiv.subtype_quotient_equiv_quotient_subtype Equiv.subtypeQuotientEquivQuotientSubtype + +@[simp] +theorem subtype_quotient_equiv_quotient_subtype_mk (p₁ : α → Prop) [s₁ : Setoid α] [s₂ : Setoid (Subtype p₁)] + (p₂ : Quotient s₁ → Prop) (hp₂ : ∀ a, p₁ a ↔ p₂ ⟦a⟧) (h : ∀ x y : Subtype p₁, @Setoid.r _ s₂ x y ↔ (x : α) ≈ y) + (x hx) : subtypeQuotientEquivQuotientSubtype p₁ p₂ hp₂ h ⟨⟦x⟧, hx⟩ = ⟦⟨x, (hp₂ _).2 hx⟩⟧ := + rfl +#align equiv.subtype_quotient_equiv_quotient_subtype_mk Equiv.subtype_quotient_equiv_quotient_subtype_mk + +@[simp] +theorem subtype_quotient_equiv_quotient_subtype_symm_mk (p₁ : α → Prop) [s₁ : Setoid α] [s₂ : Setoid (Subtype p₁)] + (p₂ : Quotient s₁ → Prop) (hp₂ : ∀ a, p₁ a ↔ p₂ ⟦a⟧) (h : ∀ x y : Subtype p₁, @Setoid.r _ s₂ x y ↔ (x : α) ≈ y) + (x) : (subtypeQuotientEquivQuotientSubtype p₁ p₂ hp₂ h).symm ⟦x⟧ = ⟨⟦x⟧, (hp₂ _).1 x.Prop⟩ := + rfl +#align equiv.subtype_quotient_equiv_quotient_subtype_symm_mk Equiv.subtype_quotient_equiv_quotient_subtype_symm_mk + +section Swap + +variable [DecidableEq α] + +/-- A helper function for `equiv.swap`. -/ +def swapCore (a b r : α) : α := + if r = a then b else if r = b then a else r +#align equiv.swap_core Equiv.swapCore + +theorem swap_core_self (r a : α) : swapCore a a r = r := by + unfold swap_core + split_ifs <;> cc +#align equiv.swap_core_self Equiv.swap_core_self + +theorem swap_core_swap_core (r a b : α) : swapCore a b (swapCore a b r) = r := by + unfold swap_core + split_ifs <;> cc +#align equiv.swap_core_swap_core Equiv.swap_core_swap_core + +theorem swap_core_comm (r a b : α) : swapCore a b r = swapCore b a r := by + unfold swap_core + split_ifs <;> cc +#align equiv.swap_core_comm Equiv.swap_core_comm + +/-- `swap a b` is the permutation that swaps `a` and `b` and + leaves other values as is. -/ +def swap (a b : α) : Perm α := + ⟨swapCore a b, swapCore a b, fun r => swap_core_swap_core r a b, fun r => swap_core_swap_core r a b⟩ +#align equiv.swap Equiv.swap + +@[simp] +theorem swap_self (a : α) : swap a a = Equiv.refl _ := + ext fun r => swap_core_self r a +#align equiv.swap_self Equiv.swap_self + +theorem swap_comm (a b : α) : swap a b = swap b a := + ext fun r => swap_core_comm r _ _ +#align equiv.swap_comm Equiv.swap_comm + +theorem swap_apply_def (a b x : α) : swap a b x = if x = a then b else if x = b then a else x := + rfl +#align equiv.swap_apply_def Equiv.swap_apply_def + +@[simp] +theorem swap_apply_left (a b : α) : swap a b a = b := + if_pos rfl +#align equiv.swap_apply_left Equiv.swap_apply_left + +@[simp] +theorem swap_apply_right (a b : α) : swap a b b = a := by by_cases h:b = a <;> simp [swap_apply_def, h] +#align equiv.swap_apply_right Equiv.swap_apply_right + +theorem swap_apply_of_ne_of_ne {a b x : α} : x ≠ a → x ≠ b → swap a b x = x := by + simp (config := { contextual := true }) [swap_apply_def] +#align equiv.swap_apply_of_ne_of_ne Equiv.swap_apply_of_ne_of_ne + +@[simp] +theorem swap_swap (a b : α) : (swap a b).trans (swap a b) = Equiv.refl _ := + ext fun x => swap_core_swap_core _ _ _ +#align equiv.swap_swap Equiv.swap_swap + +@[simp] +theorem symm_swap (a b : α) : (swap a b).symm = swap a b := + rfl +#align equiv.symm_swap Equiv.symm_swap + +@[simp] +theorem swap_eq_refl_iff {x y : α} : swap x y = Equiv.refl _ ↔ x = y := by + refine' ⟨fun h => (Equiv.refl _).Injective _, fun h => h ▸ swap_self _⟩ + rw [← h, swap_apply_left, h, refl_apply] +#align equiv.swap_eq_refl_iff Equiv.swap_eq_refl_iff + +theorem swap_comp_apply {a b x : α} (π : Perm α) : + π.trans (swap a b) x = if π x = a then b else if π x = b then a else π x := by + cases π + rfl +#align equiv.swap_comp_apply Equiv.swap_comp_apply + +theorem swap_eq_update (i j : α) : (Equiv.swap i j : α → α) = update (update id j i) i j := + funext fun x => by rw [update_apply _ i j, update_apply _ j i, Equiv.swap_apply_def, id.def] +#align equiv.swap_eq_update Equiv.swap_eq_update + +theorem comp_swap_eq_update (i j : α) (f : α → β) : f ∘ Equiv.swap i j = update (update f j (f i)) i (f j) := by + rw [swap_eq_update, comp_update, comp_update, comp.right_id] +#align equiv.comp_swap_eq_update Equiv.comp_swap_eq_update + +@[simp] +theorem symm_trans_swap_trans [DecidableEq β] (a b : α) (e : α ≃ β) : + (e.symm.trans (swap a b)).trans e = swap (e a) (e b) := + Equiv.ext fun x => by + have : ∀ a, e.symm x = a ↔ x = e a := fun a => by + rw [@eq_comm _ (e.symm x)] + constructor <;> intros <;> simp_all + simp [swap_apply_def, this] + split_ifs <;> simp +#align equiv.symm_trans_swap_trans Equiv.symm_trans_swap_trans + +@[simp] +theorem trans_swap_trans_symm [DecidableEq β] (a b : β) (e : α ≃ β) : + (e.trans (swap a b)).trans e.symm = swap (e.symm a) (e.symm b) := + symm_trans_swap_trans a b e.symm +#align equiv.trans_swap_trans_symm Equiv.trans_swap_trans_symm + +@[simp] +theorem swap_apply_self (i j a : α) : swap i j (swap i j a) = a := by + rw [← Equiv.trans_apply, Equiv.swap_swap, Equiv.refl_apply] +#align equiv.swap_apply_self Equiv.swap_apply_self + +/-- A function is invariant to a swap if it is equal at both elements -/ +theorem apply_swap_eq_self {v : α → β} {i j : α} (hv : v i = v j) (k : α) : v (swap i j k) = v k := by + by_cases hi:k = i + · rw [hi, swap_apply_left, hv] + + by_cases hj:k = j + · rw [hj, swap_apply_right, hv] + + rw [swap_apply_of_ne_of_ne hi hj] +#align equiv.apply_swap_eq_self Equiv.apply_swap_eq_self + +theorem swap_apply_eq_iff {x y z w : α} : swap x y z = w ↔ z = swap x y w := by + rw [apply_eq_iff_eq_symm_apply, symm_swap] +#align equiv.swap_apply_eq_iff Equiv.swap_apply_eq_iff + +theorem swap_apply_ne_self_iff {a b x : α} : swap a b x ≠ x ↔ a ≠ b ∧ (x = a ∨ x = b) := by + by_cases hab:a = b + · simp [hab] + + by_cases hax:x = a + · simp [hax, eq_comm] + + by_cases hbx:x = b + · simp [hbx] + + simp [hab, hax, hbx, swap_apply_of_ne_of_ne] +#align equiv.swap_apply_ne_self_iff Equiv.swap_apply_ne_self_iff + +namespace Perm + +@[simp] +theorem sum_congr_swap_refl {α β : Sort _} [DecidableEq α] [DecidableEq β] (i j : α) : + Equiv.Perm.sumCongr (Equiv.swap i j) (Equiv.refl β) = Equiv.swap (Sum.inl i) (Sum.inl j) := by + ext x + cases x + · simp [Sum.map, swap_apply_def] + split_ifs <;> rfl + + · simp [Sum.map, swap_apply_of_ne_of_ne] + +#align equiv.perm.sum_congr_swap_refl Equiv.Perm.sum_congr_swap_refl + +@[simp] +theorem sum_congr_refl_swap {α β : Sort _} [DecidableEq α] [DecidableEq β] (i j : β) : + Equiv.Perm.sumCongr (Equiv.refl α) (Equiv.swap i j) = Equiv.swap (Sum.inr i) (Sum.inr j) := by + ext x + cases x + · simp [Sum.map, swap_apply_of_ne_of_ne] + + · simp [Sum.map, swap_apply_def] + split_ifs <;> rfl + +#align equiv.perm.sum_congr_refl_swap Equiv.Perm.sum_congr_refl_swap + +end Perm + +/-- Augment an equivalence with a prescribed mapping `f a = b` -/ +def setValue (f : α ≃ β) (a : α) (b : β) : α ≃ β := + (swap a (f.symm b)).trans f +#align equiv.set_value Equiv.setValue + +@[simp] +theorem set_value_eq (f : α ≃ β) (a : α) (b : β) : setValue f a b a = b := by + dsimp [set_value] + simp [swap_apply_left] +#align equiv.set_value_eq Equiv.set_value_eq + +end Swap + +end Equiv + +namespace Function.Involutive + +/-- Convert an involutive function `f` to a permutation with `to_fun = inv_fun = f`. -/ +def toPerm (f : α → α) (h : Involutive f) : Equiv.Perm α := + ⟨f, f, h.LeftInverse, h.RightInverse⟩ +#align function.involutive.to_perm Function.Involutive.toPerm + +@[simp] +theorem coe_to_perm {f : α → α} (h : Involutive f) : (h.toPerm f : α → α) = f := + rfl +#align function.involutive.coe_to_perm Function.Involutive.coe_to_perm + +@[simp] +theorem to_perm_symm {f : α → α} (h : Involutive f) : (h.toPerm f).symm = h.toPerm f := + rfl +#align function.involutive.to_perm_symm Function.Involutive.to_perm_symm + +theorem to_perm_involutive {f : α → α} (h : Involutive f) : Involutive (h.toPerm f) := + h +#align function.involutive.to_perm_involutive Function.Involutive.to_perm_involutive + +end Function.Involutive + +theorem PLift.eq_up_iff_down_eq {x : PLift α} {y : α} : x = PLift.up y ↔ x.down = y := + Equiv.plift.eq_symm_apply +#align plift.eq_up_iff_down_eq PLift.eq_up_iff_down_eq + +theorem Function.Injective.map_swap {α β : Type _} [DecidableEq α] [DecidableEq β] {f : α → β} + (hf : Function.Injective f) (x y z : α) : f (Equiv.swap x y z) = Equiv.swap (f x) (f y) (f z) := by + conv_rhs => rw [Equiv.swap_apply_def] + split_ifs with h₁ h₂ + · rw [hf h₁, Equiv.swap_apply_left] + + · rw [hf h₂, Equiv.swap_apply_right] + + · rw [Equiv.swap_apply_of_ne_of_ne (mt (congr_arg f) h₁) (mt (congr_arg f) h₂)] + +#align function.injective.map_swap Function.Injective.map_swap + +namespace Equiv + +section + +variable (P : α → Sort w) (e : α ≃ β) + +/-- Transport dependent functions through an equivalence of the base space. +-/ +@[simps] +def piCongrLeft' : (∀ a, P a) ≃ ∀ b, P (e.symm b) where + toFun f x := f (e.symm x) + invFun f x := by + rw [← e.symm_apply_apply x] + exact f (e x) + left_inv f := + funext fun x => + eq_of_heq + ((eq_rec_heq _ _).trans + (by + dsimp + rw [e.symm_apply_apply])) + right_inv f := funext fun x => eq_of_heq ((eq_rec_heq _ _).trans (by rw [e.apply_symm_apply])) +#align equiv.Pi_congr_left' Equiv.piCongrLeft' + +end + +section + +variable (P : β → Sort w) (e : α ≃ β) + +/-- Transporting dependent functions through an equivalence of the base, +expressed as a "simplification". +-/ +def piCongrLeft : (∀ a, P (e a)) ≃ ∀ b, P b := + (piCongrLeft' P e.symm).symm +#align equiv.Pi_congr_left Equiv.piCongrLeft + +end + +section + +variable {W : α → Sort w} {Z : β → Sort z} (h₁ : α ≃ β) (h₂ : ∀ a : α, W a ≃ Z (h₁ a)) + +/-- Transport dependent functions through +an equivalence of the base spaces and a family +of equivalences of the matching fibers. +-/ +def piCongr : (∀ a, W a) ≃ ∀ b, Z b := + (Equiv.piCongrRight h₂).trans (Equiv.piCongrLeft _ h₁) +#align equiv.Pi_congr Equiv.piCongr + +@[simp] +theorem coe_Pi_congr_symm : ((h₁.piCongr h₂).symm : (∀ b, Z b) → ∀ a, W a) = fun f a => (h₂ a).symm (f (h₁ a)) := + rfl +#align equiv.coe_Pi_congr_symm Equiv.coe_Pi_congr_symm + +theorem Pi_congr_symm_apply (f : ∀ b, Z b) : (h₁.piCongr h₂).symm f = fun a => (h₂ a).symm (f (h₁ a)) := + rfl +#align equiv.Pi_congr_symm_apply Equiv.Pi_congr_symm_apply + +@[simp] +theorem Pi_congr_apply_apply (f : ∀ a, W a) (a : α) : h₁.piCongr h₂ f (h₁ a) = h₂ a (f a) := by + change cast _ ((h₂ (h₁.symm (h₁ a))) (f (h₁.symm (h₁ a)))) = (h₂ a) (f a) + generalize_proofs hZa + revert hZa + rw [h₁.symm_apply_apply a] + simp +#align equiv.Pi_congr_apply_apply Equiv.Pi_congr_apply_apply + +end + +section + +variable {W : α → Sort w} {Z : β → Sort z} (h₁ : α ≃ β) (h₂ : ∀ b : β, W (h₁.symm b) ≃ Z b) + +/-- Transport dependent functions through +an equivalence of the base spaces and a family +of equivalences of the matching fibres. +-/ +def piCongr' : (∀ a, W a) ≃ ∀ b, Z b := + (piCongr h₁.symm fun b => (h₂ b).symm).symm +#align equiv.Pi_congr' Equiv.piCongr' + +@[simp] +theorem coe_Pi_congr' : (h₁.piCongr' h₂ : (∀ a, W a) → ∀ b, Z b) = fun f b => h₂ b <| f <| h₁.symm b := + rfl +#align equiv.coe_Pi_congr' Equiv.coe_Pi_congr' + +theorem Pi_congr'_apply (f : ∀ a, W a) : h₁.piCongr' h₂ f = fun b => h₂ b <| f <| h₁.symm b := + rfl +#align equiv.Pi_congr'_apply Equiv.Pi_congr'_apply + +@[simp] +theorem Pi_congr'_symm_apply_symm_apply (f : ∀ b, Z b) (b : β) : + (h₁.piCongr' h₂).symm f (h₁.symm b) = (h₂ b).symm (f b) := by + change cast _ ((h₂ (h₁ (h₁.symm b))).symm (f (h₁ (h₁.symm b)))) = (h₂ b).symm (f b) + generalize_proofs hWb + revert hWb + generalize hb : h₁ (h₁.symm b) = b' + rw [h₁.apply_symm_apply b] at hb + subst hb + simp +#align equiv.Pi_congr'_symm_apply_symm_apply Equiv.Pi_congr'_symm_apply_symm_apply + +end + +section BinaryOp + +variable {α₁ β₁ : Type _} (e : α₁ ≃ β₁) (f : α₁ → α₁ → α₁) + +theorem semiconj_conj (f : α₁ → α₁) : Semiconj e f (e.conj f) := fun x => by simp +#align equiv.semiconj_conj Equiv.semiconj_conj + +theorem semiconj₂_conj : Semiconj₂ e f (e.arrowCongr e.conj f) := fun x y => by simp +#align equiv.semiconj₂_conj Equiv.semiconj₂_conj + +instance [IsAssociative α₁ f] : IsAssociative β₁ (e.arrowCongr (e.arrowCongr e) f) := + (e.semiconj₂_conj f).is_associative_right e.Surjective + +instance [IsIdempotent α₁ f] : IsIdempotent β₁ (e.arrowCongr (e.arrowCongr e) f) := + (e.semiconj₂_conj f).is_idempotent_right e.Surjective + +instance [IsLeftCancel α₁ f] : IsLeftCancel β₁ (e.arrowCongr (e.arrowCongr e) f) := + ⟨e.Surjective.forall₃.2 fun x y z => by simpa using @IsLeftCancel.left_cancel _ f _ x y z⟩ + +instance [IsRightCancel α₁ f] : IsRightCancel β₁ (e.arrowCongr (e.arrowCongr e) f) := + ⟨e.Surjective.forall₃.2 fun x y z => by simpa using @IsRightCancel.right_cancel _ f _ x y z⟩ + +end BinaryOp + +end Equiv + +theorem Function.Injective.swap_apply [DecidableEq α] [DecidableEq β] {f : α → β} (hf : Function.Injective f) + (x y z : α) : Equiv.swap (f x) (f y) (f z) = f (Equiv.swap x y z) := by + by_cases hx:z = x + · simp [hx] + + by_cases hy:z = y + · simp [hy] + + rw [Equiv.swap_apply_of_ne_of_ne hx hy, Equiv.swap_apply_of_ne_of_ne (hf.ne hx) (hf.ne hy)] +#align function.injective.swap_apply Function.Injective.swap_apply + +theorem Function.Injective.swap_comp [DecidableEq α] [DecidableEq β] {f : α → β} (hf : Function.Injective f) (x y : α) : + Equiv.swap (f x) (f y) ∘ f = f ∘ Equiv.swap x y := + funext fun z => hf.swap_apply _ _ _ +#align function.injective.swap_comp Function.Injective.swap_comp + +/-- If `α` is a subsingleton, then it is equivalent to `α × α`. -/ +def subsingletonProdSelfEquiv {α : Type _} [Subsingleton α] : α × α ≃ α where + toFun p := p.1 + invFun a := (a, a) + left_inv p := Subsingleton.elim _ _ + right_inv p := Subsingleton.elim _ _ +#align subsingleton_prod_self_equiv subsingletonProdSelfEquiv + +/-- To give an equivalence between two subsingleton types, it is sufficient to give any two + functions between them. -/ +def equivOfSubsingletonOfSubsingleton [Subsingleton α] [Subsingleton β] (f : α → β) (g : β → α) : α ≃ β where + toFun := f + invFun := g + left_inv _ := Subsingleton.elim _ _ + right_inv _ := Subsingleton.elim _ _ +#align equiv_of_subsingleton_of_subsingleton equivOfSubsingletonOfSubsingleton + +/-- A nonempty subsingleton type is (noncomputably) equivalent to `punit`. -/ +noncomputable def Equiv.punitOfNonemptyOfSubsingleton {α : Sort _} [h : Nonempty α] [Subsingleton α] : α ≃ PUnit.{v} := + equivOfSubsingletonOfSubsingleton (fun _ => PUnit.unit) fun _ => h.some +#align equiv.punit_of_nonempty_of_subsingleton Equiv.punitOfNonemptyOfSubsingleton + +/-- `unique (unique α)` is equivalent to `unique α`. -/ +def uniqueUniqueEquiv : Unique (Unique α) ≃ Unique α := + equivOfSubsingletonOfSubsingleton (fun h => h.default) fun h => + { default := h, uniq := fun _ => Subsingleton.elim _ _ } +#align unique_unique_equiv uniqueUniqueEquiv + +namespace Function + +theorem update_comp_equiv {α β α' : Sort _} [DecidableEq α'] [DecidableEq α] (f : α → β) (g : α' ≃ α) (a : α) (v : β) : + update f a v ∘ g = update (f ∘ g) (g.symm a) v := by + rw [← update_comp_eq_of_injective _ g.injective, g.apply_symm_apply] +#align function.update_comp_equiv Function.update_comp_equiv + +theorem update_apply_equiv_apply {α β α' : Sort _} [DecidableEq α'] [DecidableEq α] (f : α → β) (g : α' ≃ α) (a : α) + (v : β) (a' : α') : update f a v (g a') = update (f ∘ g) (g.symm a) v a' := + congr_fun (update_comp_equiv f g a v) a' +#align function.update_apply_equiv_apply Function.update_apply_equiv_apply + +theorem Pi_congr_left'_update [DecidableEq α] [DecidableEq β] (P : α → Sort _) (e : α ≃ β) (f : ∀ a, P a) (b : β) + (x : P (e.symm b)) : e.piCongrLeft' P (update f (e.symm b) x) = update (e.piCongrLeft' P f) b x := by + ext b' + rcases eq_or_ne b' b with (rfl | h) + · simp + + · simp [h] + +#align function.Pi_congr_left'_update Function.Pi_congr_left'_update + +theorem Pi_congr_left'_symm_update [DecidableEq α] [DecidableEq β] (P : α → Sort _) (e : α ≃ β) (f : ∀ b, P (e.symm b)) + (b : β) (x : P (e.symm b)) : + (e.piCongrLeft' P).symm (update f b x) = update ((e.piCongrLeft' P).symm f) (e.symm b) x := by + simp [(e.Pi_congr_left' P).symm_apply_eq, Pi_congr_left'_update] +#align function.Pi_congr_left'_symm_update Function.Pi_congr_left'_symm_update + +end Function From 2a6c0a1a456c0cddb90f64c4bb35caa3e7032674 Mon Sep 17 00:00:00 2001 From: ruben vorster Date: Thu, 17 Nov 2022 19:43:24 +0000 Subject: [PATCH 026/127] add Trans instance for Equiv --- Mathlib/Logic/Equiv/Defs.lean | 3 +++ 1 file changed, 3 insertions(+) diff --git a/Mathlib/Logic/Equiv/Defs.lean b/Mathlib/Logic/Equiv/Defs.lean index c3d07c1009863..04c72500e20ae 100644 --- a/Mathlib/Logic/Equiv/Defs.lean +++ b/Mathlib/Logic/Equiv/Defs.lean @@ -157,6 +157,9 @@ initialize_simps_projections Equiv (toFun → apply, invFun → symmApply) protected def trans (e₁ : α ≃ β) (e₂ : β ≃ γ) : α ≃ γ := ⟨e₂ ∘ e₁, e₁.symm ∘ e₂.symm, e₂.left_inv.comp e₁.left_inv, e₂.right_inv.comp e₁.right_inv⟩ +instance : Trans Equiv Equiv Equiv where + trans := Equiv.trans + -- porting note: this lemma is now useless since coercions are eagerly unfolded /- @[simp] From 8299dc9768e1c2f2a8b30f50684a68c279df3176 Mon Sep 17 00:00:00 2001 From: ruben vorster Date: Thu, 17 Nov 2022 19:54:25 +0000 Subject: [PATCH 027/127] fix prodComm --- Mathlib/Logic/Equiv/Basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index bc1ca50a6e754..f5365d0b090a2 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -104,7 +104,7 @@ def prodComm (α β : Type _) : α × β ≃ β × α := #align equiv.prod_comm Equiv.prodComm @[simp] -theorem coe_prod_comm (α β : Type _) : ⇑(prodComm α β) = Prod.swap := +theorem coe_prod_comm (α β : Type _) : (⇑(prodComm α β) : α × β → β × α) = Prod.swap := rfl #align equiv.coe_prod_comm Equiv.coe_prod_comm From 04d962a6c111fff857cc8e4682de13924266ce6a Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Fri, 18 Nov 2022 10:17:24 +1100 Subject: [PATCH 028/127] linting --- Mathlib/Logic/Equiv/Defs.lean | 12 ++++++------ 1 file changed, 6 insertions(+), 6 deletions(-) diff --git a/Mathlib/Logic/Equiv/Defs.lean b/Mathlib/Logic/Equiv/Defs.lean index 04c72500e20ae..0ea372a3bf1ce 100644 --- a/Mathlib/Logic/Equiv/Defs.lean +++ b/Mathlib/Logic/Equiv/Defs.lean @@ -503,12 +503,12 @@ def propEquivPUnit {p : Prop} (h : p) : p ≃ PUnit.{0} := #align equiv.prop_equiv_punit Equiv.propEquivPUnit /-- `ULift α` is equivalent to `α`. -/ -@[simps (config := { fullyApplied := false }) apply symmApply] +@[simps (config := { fullyApplied := false }) apply] protected def ulift {α : Type v} : ULift.{u} α ≃ α := ⟨ULift.down, ULift.up, ULift.up_down, fun _ => rfl⟩ /-- `PLift α` is equivalent to `α`. -/ -@[simps (config := { fullyApplied := false }) apply symmApply] +@[simps (config := { fullyApplied := false }) apply] protected def plift : PLift α ≃ α := ⟨PLift.down, PLift.up, PLift.up_down, PLift.down_up⟩ @@ -664,7 +664,7 @@ def piSubsingleton {α} (β : α → Sort _) [Subsingleton α] (a : α) : (∀ a right_inv b := rfl /-- If `α` has a unique term, then the type of function `α → β` is equivalent to `β`. -/ -@[simps (config := { fullyApplied := false })] +@[simps (config := { fullyApplied := false }) apply] def funUnique (α β) [Unique.{u} α] : (α → β) ≃ β := piSubsingleton _ default @@ -730,7 +730,7 @@ def psigmaCongrRight {α} {β₁ β₂ : α → Sort _} (F : ∀ a, β₁ a ≃ congr_arg (PSigma.mk a) <| symm_apply_apply (F a) b, fun ⟨a, b⟩ => congr_arg (PSigma.mk a) <| apply_symm_apply (F a) b⟩ -@[simp] +-- Porting note: simp can now simplify the LHS, so I have removed `@[simp]` theorem psigmaCongrRight_trans {α} {β₁ β₂ β₃ : α → Sort _} (F : ∀ a, β₁ a ≃ β₂ a) (G : ∀ a, β₂ a ≃ β₃ a) : (psigmaCongrRight F).trans (psigmaCongrRight G) = @@ -740,7 +740,7 @@ theorem psigmaCongrRight_trans {α} {β₁ β₂ β₃ : α → Sort _} rfl #align equiv.psigma_congr_right_trans Equiv.psigmaCongrRight_trans -@[simp] +-- Porting note: simp can now simplify the LHS, so I have removed `@[simp]` theorem psigmaCongrRight_symm {α} {β₁ β₂ : α → Sort _} (F : ∀ a, β₁ a ≃ β₂ a) : (psigmaCongrRight F).symm = psigmaCongrRight fun a => (F a).symm := by ext1 x @@ -748,7 +748,7 @@ theorem psigmaCongrRight_symm {α} {β₁ β₂ : α → Sort _} (F : ∀ a, β rfl #align equiv.psigma_congr_right_symm Equiv.psigmaCongrRight_symm -@[simp] +-- Porting note: simp can now prove this, so I have removed `@[simp]` theorem psigmaCongrRight_refl {α} {β : α → Sort _} : (psigmaCongrRight fun a => Equiv.refl (β a)) = Equiv.refl (Σ' a, β a) := by ext1 x From c74802b044579955e8296e0801319c09f4bf6ec7 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Fri, 18 Nov 2022 10:29:58 +1100 Subject: [PATCH 029/127] lint --- Mathlib/Logic/Equiv/Defs.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/Mathlib/Logic/Equiv/Defs.lean b/Mathlib/Logic/Equiv/Defs.lean index 0ea372a3bf1ce..6d0794bba8a4d 100644 --- a/Mathlib/Logic/Equiv/Defs.lean +++ b/Mathlib/Logic/Equiv/Defs.lean @@ -764,7 +764,7 @@ def sigmaCongrRight {α} {β₁ β₂ : α → Type _} (F : ∀ a, β₁ a ≃ congr_arg (Sigma.mk a) <| symm_apply_apply (F a) b, fun ⟨a, b⟩ => congr_arg (Sigma.mk a) <| apply_symm_apply (F a) b⟩ -@[simp] +-- Porting note: simp can now simplify the LHS, so I have removed `@[simp]` theorem sigmaCongrRight_trans {α} {β₁ β₂ β₃ : α → Type _} (F : ∀ a, β₁ a ≃ β₂ a) (G : ∀ a, β₂ a ≃ β₃ a) : (sigmaCongrRight F).trans (sigmaCongrRight G) = sigmaCongrRight fun a => (F a).trans (G a) := by @@ -773,7 +773,7 @@ theorem sigmaCongrRight_trans {α} {β₁ β₂ β₃ : α → Type _} rfl #align equiv.sigmaCongrRight Equiv.sigmaCongrRight_trans -@[simp] +-- Porting note: simp can now simplify the LHS, so I have removed `@[simp]` theorem sigmaCongrRight_symm {α} {β₁ β₂ : α → Type _} (F : ∀ a, β₁ a ≃ β₂ a) : (sigmaCongrRight F).symm = sigmaCongrRight fun a => (F a).symm := by ext1 x @@ -781,7 +781,7 @@ theorem sigmaCongrRight_symm {α} {β₁ β₂ : α → Type _} (F : ∀ a, β rfl #align equiv.sigma_congr_right_symm Equiv.sigmaCongrRight_symm -@[simp] +-- Porting note: simp can now prove this, so I have removed `@[simp]` theorem sigmaCongrRight_refl {α} {β : α → Type _} : (sigmaCongrRight fun a => Equiv.refl (β a)) = Equiv.refl (Σ a, β a) := by ext1 x From a067a608a854c9a93cadddd3c146ec2f71e2cabd Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Fri, 18 Nov 2022 10:48:25 +1100 Subject: [PATCH 030/127] Mathlib.lean --- Mathlib.lean | 1 + 1 file changed, 1 insertion(+) diff --git a/Mathlib.lean b/Mathlib.lean index 06b0afd5deb89..3ecf53db9abc5 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -95,6 +95,7 @@ import Mathlib.Lean.Expr.Traverse import Mathlib.Lean.LocalContext import Mathlib.Lean.Meta import Mathlib.Logic.Basic +import Mathlib.Logic.Equiv.Defs import Mathlib.Logic.Equiv.LocalEquiv import Mathlib.Logic.Equiv.MfldSimpsAttr import Mathlib.Logic.Function.Basic From 4f566ca49e820b6d366342c8d4b7c0b30c961dfe Mon Sep 17 00:00:00 2001 From: Mario Carneiro Date: Thu, 17 Nov 2022 19:35:46 -0500 Subject: [PATCH 031/127] style --- Mathlib/Logic/Equiv/Defs.lean | 576 +++++++++++----------------------- 1 file changed, 179 insertions(+), 397 deletions(-) diff --git a/Mathlib/Logic/Equiv/Defs.lean b/Mathlib/Logic/Equiv/Defs.lean index 6d0794bba8a4d..63933f731ed32 100644 --- a/Mathlib/Logic/Equiv/Defs.lean +++ b/Mathlib/Logic/Equiv/Defs.lean @@ -90,13 +90,9 @@ instance : EquivLike (α ≃ β) α β where inv := invFun left_inv := left_inv right_inv := right_inv - coe_injective' e₁ e₂ h₁ h₂ := by - cases e₁ - cases e₂ - congr + coe_injective' e₁ e₂ h₁ h₂ := by cases e₁; cases e₂; congr -instance : CoeFun (α ≃ β) fun _ => α → β := - ⟨toFun⟩ +instance : CoeFun (α ≃ β) fun _ => α → β := ⟨toFun⟩ /-- The map `(r ≃ s) → (r → s)` is injective. -/ theorem coe_fn_injective : @Function.Injective (α ≃ β) (α → β) (fun e => e) := @@ -105,9 +101,7 @@ theorem coe_fn_injective : @Function.Injective (α ≃ β) (α → β) (fun e => protected theorem coe_inj {e₁ e₂ : α ≃ β} : (e₁ : α → β) = e₂ ↔ e₁ = e₂ := @FunLike.coe_fn_eq _ _ _ _ e₁ e₂ -@[ext] -theorem ext {f g : Equiv α β} (H : ∀ x, f x = g x) : f = g := - FunLike.ext f g H +@[ext] theorem ext {f g : Equiv α β} (H : ∀ x, f x = g x) : f = g := FunLike.ext f g H protected theorem congr_arg {f : Equiv α β} {x x' : α} : x = x' → f x = f x' := FunLike.congr_arg f @@ -115,12 +109,9 @@ protected theorem congr_arg {f : Equiv α β} {x x' : α} : x = x' → f x = f x protected theorem congr_fun {f g : Equiv α β} (h : f = g) (x : α) : f x = g x := FunLike.congr_fun h x -theorem ext_iff {f g : Equiv α β} : f = g ↔ ∀ x, f x = g x := - FunLike.ext_iff +theorem ext_iff {f g : Equiv α β} : f = g ↔ ∀ x, f x = g x := FunLike.ext_iff -@[ext] -theorem Perm.ext {σ τ : Equiv.Perm α} (H : ∀ x, σ x = τ x) : σ = τ := - Equiv.ext H +@[ext] theorem Perm.ext {σ τ : Equiv.Perm α} (H : ∀ x, σ x = τ x) : σ = τ := Equiv.ext H protected theorem Perm.congr_arg {f : Equiv.Perm α} {x x' : α} : x = x' → f x = f x' := Equiv.congr_arg @@ -128,26 +119,20 @@ protected theorem Perm.congr_arg {f : Equiv.Perm α} {x x' : α} : x = x' → f protected theorem Perm.congr_fun {f g : Equiv.Perm α} (h : f = g) (x : α) : f x = g x := Equiv.congr_fun h x -theorem Perm.ext_iff {σ τ : Equiv.Perm α} : σ = τ ↔ ∀ x, σ x = τ x := - Equiv.ext_iff +theorem Perm.ext_iff {σ τ : Equiv.Perm α} : σ = τ ↔ ∀ x, σ x = τ x := Equiv.ext_iff /-- Any type is equivalent to itself. -/ -@[refl] -protected def refl (α : Sort _) : α ≃ α := - ⟨id, id, fun _ => rfl, fun _ => rfl⟩ +@[refl] protected def refl (α : Sort _) : α ≃ α := ⟨id, id, fun _ => rfl, fun _ => rfl⟩ -instance inhabited' : Inhabited (α ≃ α) := - ⟨Equiv.refl α⟩ +instance inhabited' : Inhabited (α ≃ α) := ⟨Equiv.refl α⟩ /-- Inverse of an equivalence `e : α ≃ β`. -/ -- Porting note: `symm` attribute rejects this lemma because of implicit arguments. -- @[symm] -protected def symm (e : α ≃ β) : β ≃ α := - ⟨e.invFun, e.toFun, e.right_inv, e.left_inv⟩ +protected def symm (e : α ≃ β) : β ≃ α := ⟨e.invFun, e.toFun, e.right_inv, e.left_inv⟩ /-- See Note [custom simps projection] -/ -def Simps.symmApply (e : α ≃ β) : β → α := - e.symm +def Simps.symmApply (e : α ≃ β) : β → α := e.symm initialize_simps_projections Equiv (toFun → apply, invFun → symmApply) @@ -161,24 +146,15 @@ instance : Trans Equiv Equiv Equiv where trans := Equiv.trans -- porting note: this lemma is now useless since coercions are eagerly unfolded -/- -@[simp] -theorem to_fun_as_coe (e : α ≃ β) : e.toFun = e := - rfl --/ +-- @[simp] theorem to_fun_as_coe (e : α ≃ β) : e.toFun = e := rfl -@[simp] -theorem inv_fun_as_coe (e : α ≃ β) : e.invFun = e.symm := - rfl +@[simp] theorem inv_fun_as_coe (e : α ≃ β) : e.invFun = e.symm := rfl -protected theorem injective (e : α ≃ β) : Injective e := - EquivLike.injective e +protected theorem injective (e : α ≃ β) : Injective e := EquivLike.injective e -protected theorem surjective (e : α ≃ β) : Surjective e := - EquivLike.surjective e +protected theorem surjective (e : α ≃ β) : Surjective e := EquivLike.surjective e -protected theorem bijective (e : α ≃ β) : Bijective e := - EquivLike.bijective e +protected theorem bijective (e : α ≃ β) : Bijective e := EquivLike.bijective e protected theorem subsingleton (e : α ≃ β) [Subsingleton β] : Subsingleton α := e.injective.subsingleton @@ -205,35 +181,23 @@ theorem Perm.subsingleton_eq_refl [Subsingleton α] (e : Perm α) : e = Equiv.re protected def decidableEq (e : α ≃ β) [DecidableEq β] : DecidableEq α := e.injective.decidableEq -theorem nonempty_congr (e : α ≃ β) : Nonempty α ↔ Nonempty β := - Nonempty.congr e e.symm +theorem nonempty_congr (e : α ≃ β) : Nonempty α ↔ Nonempty β := Nonempty.congr e e.symm -protected theorem nonempty (e : α ≃ β) [Nonempty β] : Nonempty α := - e.nonempty_congr.mpr ‹_› +protected theorem nonempty (e : α ≃ β) [Nonempty β] : Nonempty α := e.nonempty_congr.mpr ‹_› /-- If `α ≃ β` and `β` is inhabited, then so is `α`. -/ -protected def inhabited [Inhabited β] (e : α ≃ β) : Inhabited α := - ⟨e.symm default⟩ +protected def inhabited [Inhabited β] (e : α ≃ β) : Inhabited α := ⟨e.symm default⟩ /-- If `α ≃ β` and `β` is a singleton type, then so is `α`. -/ -protected def unique [Unique β] (e : α ≃ β) : Unique α := - e.symm.surjective.unique +protected def unique [Unique β] (e : α ≃ β) : Unique α := e.symm.surjective.unique /-- Equivalence between equal types. -/ protected def cast {α β : Sort _} (h : α = β) : α ≃ β := - ⟨cast h, cast h.symm, fun x => by - cases h - rfl, fun x => by - cases h - rfl⟩ - -@[simp] -theorem coe_fn_symm_mk (f : α → β) (g l r) : ((Equiv.mk f g l r).symm : β → α) = g := - rfl + ⟨cast h, cast h.symm, fun _ => by cases h; rfl, fun _ => by cases h; rfl⟩ -@[simp] -theorem coe_refl : (Equiv.refl α : α → α) = id := - rfl +@[simp] theorem coe_fn_symm_mk (f : α → β) (g l r) : ((Equiv.mk f g l r).symm : β → α) = g := rfl + +@[simp] theorem coe_refl : (Equiv.refl α : α → α) = id := rfl /-- This cannot be a `simp` lemmas as it incorrectly matches against `e : α ≃ synonym α`, when `synonym α` is semireducible. This makes a mess of `multiplicative.of_add` etc. -/ @@ -242,74 +206,47 @@ theorem Perm.coe_subsingleton {α : Type _} [Subsingleton α] (e : Perm α) : (e -- porting note: marking this as `@[simp]` because `simp` doesn't fire on `coe_refl` -- in an expression such as `Equiv.refl a x` -@[simp] -theorem refl_apply (x : α) : Equiv.refl α x = x := - rfl +@[simp] theorem refl_apply (x : α) : Equiv.refl α x = x := rfl -@[simp] -theorem coe_trans (f : α ≃ β) (g : β ≃ γ) : (f.trans g : α → γ) = g ∘ f := - rfl +@[simp] theorem coe_trans (f : α ≃ β) (g : β ≃ γ) : (f.trans g : α → γ) = g ∘ f := rfl -- porting note: marking this as `@[simp]` because `simp` doesn't fire on `coe_trans` -- in an expression such as `Equiv.trans f g x` -@[simp] -theorem trans_apply (f : α ≃ β) (g : β ≃ γ) (a : α) : (f.trans g) a = g (f a) := - rfl +@[simp] theorem trans_apply (f : α ≃ β) (g : β ≃ γ) (a : α) : (f.trans g) a = g (f a) := rfl -@[simp] -theorem apply_symm_apply (e : α ≃ β) (x : β) : e (e.symm x) = x := - e.right_inv x +@[simp] theorem apply_symm_apply (e : α ≃ β) (x : β) : e (e.symm x) = x := e.right_inv x -@[simp] -theorem symm_apply_apply (e : α ≃ β) (x : α) : e.symm (e x) = x := - e.left_inv x +@[simp] theorem symm_apply_apply (e : α ≃ β) (x : α) : e.symm (e x) = x := e.left_inv x -@[simp] -theorem symm_comp_self (e : α ≃ β) : e.symm ∘ e = id := - funext e.symm_apply_apply +@[simp] theorem symm_comp_self (e : α ≃ β) : e.symm ∘ e = id := funext e.symm_apply_apply -@[simp] -theorem self_comp_symm (e : α ≃ β) : e ∘ e.symm = id := - funext e.apply_symm_apply +@[simp] theorem self_comp_symm (e : α ≃ β) : e ∘ e.symm = id := funext e.apply_symm_apply -@[simp] -theorem symm_trans_apply (f : α ≃ β) (g : β ≃ γ) (a : γ) : (f.trans g).symm a = f.symm (g.symm a) := - rfl +@[simp] theorem symm_trans_apply (f : α ≃ β) (g : β ≃ γ) (a : γ) : + (f.trans g).symm a = f.symm (g.symm a) := rfl -- The `simp` attribute is needed to make this a `dsimp` lemma. -- `simp` will always rewrite with `equiv.symm_symm` before this has a chance to fire. -@[simp, nolint simpNF] -theorem symm_symm_apply (f : α ≃ β) (b : α) : f.symm.symm b = f b := - rfl +@[simp, nolint simpNF] theorem symm_symm_apply (f : α ≃ β) (b : α) : f.symm.symm b = f b := rfl -theorem apply_eq_iff_eq (f : α ≃ β) {x y : α} : f x = f y ↔ x = y := - EquivLike.apply_eq_iff_eq f +theorem apply_eq_iff_eq (f : α ≃ β) {x y : α} : f x = f y ↔ x = y := EquivLike.apply_eq_iff_eq f -theorem apply_eq_iff_eq_symm_apply (f : α ≃ β) : f x = y ↔ x = f.symm y - := by conv_lhs => rw [ ← apply_symm_apply f y ] - rw [ apply_eq_iff_eq ] +theorem apply_eq_iff_eq_symm_apply (f : α ≃ β) : f x = y ↔ x = f.symm y := by + conv_lhs => rw [← apply_symm_apply f y] + rw [apply_eq_iff_eq] -@[simp] -theorem cast_apply {α β} (h : α = β) (x : α) : Equiv.cast h x = cast h x := - rfl +@[simp] theorem cast_apply {α β} (h : α = β) (x : α) : Equiv.cast h x = cast h x := rfl -@[simp] -theorem cast_symm {α β} (h : α = β) : (Equiv.cast h).symm = Equiv.cast h.symm := - rfl +@[simp] theorem cast_symm {α β} (h : α = β) : (Equiv.cast h).symm = Equiv.cast h.symm := rfl -@[simp] -theorem cast_refl {α} (h : α = α := rfl) : Equiv.cast h = Equiv.refl α := - rfl +@[simp] theorem cast_refl {α} (h : α = α := rfl) : Equiv.cast h = Equiv.refl α := rfl -@[simp] -theorem cast_trans {α β γ} (h : α = β) (h2 : β = γ) : +@[simp] theorem cast_trans {α β γ} (h : α = β) (h2 : β = γ) : (Equiv.cast h).trans (Equiv.cast h2) = Equiv.cast (h.trans h2) := - ext fun x => by substs h h2 - rfl + ext fun x => by substs h h2; rfl -theorem cast_eq_iff_heq {α β} (h : α = β) {a : α} {b : β} : Equiv.cast h a = b ↔ HEq a b := -by subst h - simp [coe_refl] +theorem cast_eq_iff_heq {α β} (h : α = β) {a : α} {b : β} : Equiv.cast h a = b ↔ HEq a b := by + subst h; simp [coe_refl] theorem symm_apply_eq {α β} (e : α ≃ β) {x y} : e.symm x = y ↔ x = e y := ⟨fun H => by simp [H.symm], fun H => by simp [H]⟩ @@ -317,42 +254,24 @@ theorem symm_apply_eq {α β} (e : α ≃ β) {x y} : e.symm x = y ↔ x = e y : theorem eq_symm_apply {α β} (e : α ≃ β) {x y} : y = e.symm x ↔ e y = x := (eq_comm.trans e.symm_apply_eq).trans eq_comm -@[simp] -theorem symm_symm (e : α ≃ β) : e.symm.symm = e := by - cases e - rfl +@[simp] theorem symm_symm (e : α ≃ β) : e.symm.symm = e := by cases e; rfl -@[simp] -theorem trans_refl (e : α ≃ β) : e.trans (Equiv.refl β) = e := by - cases e - rfl +@[simp] theorem trans_refl (e : α ≃ β) : e.trans (Equiv.refl β) = e := by cases e; rfl -@[simp] -theorem refl_symm : (Equiv.refl α).symm = Equiv.refl α := - rfl +@[simp] theorem refl_symm : (Equiv.refl α).symm = Equiv.refl α := rfl -@[simp] -theorem refl_trans (e : α ≃ β) : (Equiv.refl α).trans e = e := by - cases e - rfl +@[simp] theorem refl_trans (e : α ≃ β) : (Equiv.refl α).trans e = e := by cases e; rfl -@[simp] -theorem symm_trans_self (e : α ≃ β) : e.symm.trans e = Equiv.refl β := - ext (by simp) +@[simp] theorem symm_trans_self (e : α ≃ β) : e.symm.trans e = Equiv.refl β := ext <| by simp -@[simp] -theorem self_trans_symm (e : α ≃ β) : e.trans e.symm = Equiv.refl α := - ext (by simp) +@[simp] theorem self_trans_symm (e : α ≃ β) : e.trans e.symm = Equiv.refl α := ext <| by simp theorem trans_assoc {δ} (ab : α ≃ β) (bc : β ≃ γ) (cd : γ ≃ δ) : - (ab.trans bc).trans cd = ab.trans (bc.trans cd) := - Equiv.ext fun _ => rfl + (ab.trans bc).trans cd = ab.trans (bc.trans cd) := Equiv.ext fun _ => rfl -theorem left_inverse_symm (f : Equiv α β) : LeftInverse f.symm f := - f.left_inv +theorem left_inverse_symm (f : Equiv α β) : LeftInverse f.symm f := f.left_inv -theorem right_inverse_symm (f : Equiv α β) : Function.RightInverse f.symm f := - f.right_inv +theorem right_inverse_symm (f : Equiv α β) : Function.RightInverse f.symm f := f.right_inv theorem injective_comp (e : α ≃ β) (f : β → γ) : Injective (f ∘ e) ↔ Injective f := EquivLike.injective_comp e f @@ -374,52 +293,35 @@ theorem comp_bijective (f : α → β) (e : β ≃ γ) : Bijective (e ∘ f) ↔ /-- If `α` is equivalent to `β` and `γ` is equivalent to `δ`, then the type of equivalences `α ≃ γ` is equivalent to the type of equivalences `β ≃ δ`. -/ -def equivCongr {δ} (ab : α ≃ β) (cd : γ ≃ δ) : (α ≃ γ) ≃ (β ≃ δ) := - ⟨ fun ac => (ab.symm.trans ac).trans cd, - fun bd => ab.trans <| bd.trans <| cd.symm, - fun ac => by - ext x - simp only [trans_apply, comp_apply, symm_apply_apply], - fun ac => by - ext x - simp only [trans_apply, comp_apply, apply_symm_apply] ⟩ - -@[simp] -theorem equivCongr_refl {α β} : (Equiv.refl α).equivCongr (Equiv.refl β) = Equiv.refl (α ≃ β) := by - ext - rfl +def equivCongr (ab : α ≃ β) (cd : γ ≃ δ) : (α ≃ γ) ≃ (β ≃ δ) where + toFun ac := (ab.symm.trans ac).trans cd + invFun bd := ab.trans <| bd.trans <| cd.symm + left_inv ac := by ext x; simp only [trans_apply, comp_apply, symm_apply_apply] + right_inv ac := by ext x; simp only [trans_apply, comp_apply, apply_symm_apply] + +@[simp] theorem equivCongr_refl {α β} : + (Equiv.refl α).equivCongr (Equiv.refl β) = Equiv.refl (α ≃ β) := by ext; rfl #align equiv.equiv_congr_refl Equiv.equivCongr_refl -@[simp] -theorem equivCongr_symm {δ} (ab : α ≃ β) (cd : γ ≃ δ) : - (ab.equivCongr cd).symm = ab.symm.equivCongr cd.symm := by - ext - rfl +@[simp] theorem equivCongr_symm {δ} (ab : α ≃ β) (cd : γ ≃ δ) : + (ab.equivCongr cd).symm = ab.symm.equivCongr cd.symm := by ext; rfl #align equiv.equiv_congr_symm Equiv.equivCongr_symm -@[simp] -theorem equivCongr_trans {δ ε ζ} (ab : α ≃ β) (de : δ ≃ ε) (bc : β ≃ γ) (ef : ε ≃ ζ) : +@[simp] theorem equivCongr_trans {δ ε ζ} (ab : α ≃ β) (de : δ ≃ ε) (bc : β ≃ γ) (ef : ε ≃ ζ) : (ab.equivCongr de).trans (bc.equivCongr ef) = (ab.trans bc).equivCongr (de.trans ef) := by - ext - rfl + ext; rfl #align equiv.equiv_congr_trans Equiv.equivCongr_trans -@[simp] -theorem equivCongr_refl_left {α β γ} (bg : β ≃ γ) (e : α ≃ β) : - (Equiv.refl α).equivCongr bg e = e.trans bg := - rfl +@[simp] theorem equivCongr_refl_left {α β γ} (bg : β ≃ γ) (e : α ≃ β) : + (Equiv.refl α).equivCongr bg e = e.trans bg := rfl #align equiv.equiv_congr_refl_left Equiv.equivCongr_refl_left -@[simp] -theorem equivCongr_refl_right {α β} (ab e : α ≃ β) : - ab.equivCongr (Equiv.refl β) e = ab.symm.trans e := - rfl +@[simp] theorem equivCongr_refl_right {α β} (ab e : α ≃ β) : + ab.equivCongr (Equiv.refl β) e = ab.symm.trans e := rfl #align equiv.equiv_congr_refl_right Equiv.equivCongr_refl_right -@[simp] -theorem equivCongr_apply_apply {δ} (ab : α ≃ β) (cd : γ ≃ δ) (e : α ≃ γ) (x) : - ab.equivCongr cd e x = cd (e (ab.symm x)) := - rfl +@[simp] theorem equivCongr_apply_apply {δ} (ab : α ≃ β) (cd : γ ≃ δ) (e : α ≃ γ) (x) : + ab.equivCongr cd e x = cd (e (ab.symm x)) := rfl #align equiv.equiv_congr_apply_apply Equiv.equivCongr_apply_apply section permCongr @@ -427,35 +329,28 @@ section permCongr variable {α' β' : Type _} (e : α' ≃ β') /-- If `α` is equivalent to `β`, then `perm α` is equivalent to `perm β`. -/ -def permCongr : Perm α' ≃ Perm β' := - equivCongr e e +def permCongr : Perm α' ≃ Perm β' := equivCongr e e -theorem permCongr_def (p : Equiv.Perm α') : e.permCongr p = (e.symm.trans p).trans e := - rfl +theorem permCongr_def (p : Equiv.Perm α') : e.permCongr p = (e.symm.trans p).trans e := rfl #align equiv.perm_congr_def Equiv.permCongr_def -@[simp] -theorem permCongr_refl : e.permCongr (Equiv.refl _) = Equiv.refl _ := by simp [permCongr_def] +@[simp] theorem permCongr_refl : e.permCongr (Equiv.refl _) = Equiv.refl _ := by + simp [permCongr_def] #align equiv.perm_congr_refl Equiv.permCongr_refl -@[simp] -theorem permCongr_symm : e.permCongr.symm = e.symm.permCongr := - rfl +@[simp] theorem permCongr_symm : e.permCongr.symm = e.symm.permCongr := rfl #align equiv.perm_congr_symm Equiv.permCongr_symm -@[simp] -theorem permCongr_apply (p : Equiv.Perm α') (x) : e.permCongr p x = e (p (e.symm x)) := - rfl +@[simp] theorem permCongr_apply (p : Equiv.Perm α') (x) : e.permCongr p x = e (p (e.symm x)) := rfl #align equiv.perm_congr_apply Equiv.permCongr_apply -theorem permCongr_symm_apply (p : Equiv.Perm β') (x) : e.permCongr.symm p x = e.symm (p (e x)) := - rfl +theorem permCongr_symm_apply (p : Equiv.Perm β') (x) : + e.permCongr.symm p x = e.symm (p (e x)) := rfl #align equiv.perm_congr_symm_apply Equiv.permCongr_symm_apply theorem permCongr_trans (p p' : Equiv.Perm α') : (e.permCongr p).trans (e.permCongr p') = e.permCongr (p.trans p') := by - ext - simp only [trans_apply, comp_apply, permCongr_apply, symm_apply_apply] + ext; simp only [trans_apply, comp_apply, permCongr_apply, symm_apply_apply] #align equiv.perm_congr_trans Equiv.permCongr_trans end permCongr @@ -465,12 +360,10 @@ def equivOfIsEmpty (α β : Sort _) [IsEmpty α] [IsEmpty β] : α ≃ β := ⟨isEmptyElim, isEmptyElim, isEmptyElim, isEmptyElim⟩ /-- If `α` is an empty type, then it is equivalent to the `Empty` type. -/ -def equivEmpty (α : Sort u) [IsEmpty α] : α ≃ Empty := - equivOfIsEmpty α _ +def equivEmpty (α : Sort u) [IsEmpty α] : α ≃ Empty := equivOfIsEmpty α _ /-- If `α` is an empty type, then it is equivalent to the `PEmpty` type in any universe. -/ -def equivPEmpty (α : Sort v) [IsEmpty α] : α ≃ PEmpty.{u} := - equivOfIsEmpty α _ +def equivPEmpty (α : Sort v) [IsEmpty α] : α ≃ PEmpty.{u} := equivOfIsEmpty α _ #align equiv.equiv_pempty Equiv.equivPEmpty /-- `α` is equivalent to an empty type iff `α` is empty. -/ @@ -479,8 +372,7 @@ def equivEmptyEquiv (α : Sort u) : α ≃ Empty ≃ IsEmpty α := #align equiv.equiv_empty_equiv Equiv.equivEmptyEquiv /-- The `Sort` of proofs of a false proposition is equivalent to `PEmpty`. -/ -def propEquivPEmpty {p : Prop} (h : ¬p) : p ≃ PEmpty := - @equivPEmpty p <| IsEmpty.prop_iff.2 h +def propEquivPEmpty {p : Prop} (h : ¬p) : p ≃ PEmpty := @equivPEmpty p <| IsEmpty.prop_iff.2 h #align equiv.prop_equiv_pempty Equiv.propEquivPEmpty /-- If both `α` and `β` have a unique element, then `α ≃ β`. -/ @@ -491,15 +383,11 @@ def equivOfUnique (α β : Sort _) [Unique.{u} α] [Unique.{v} β] : α ≃ β w right_inv _ := Subsingleton.elim _ _ /-- If `α` has a unique element, then it is equivalent to any `PUnit`. -/ -def equivPUnit (α : Sort u) [Unique α] : α ≃ PUnit.{v} := - equivOfUnique α _ +def equivPUnit (α : Sort u) [Unique α] : α ≃ PUnit.{v} := equivOfUnique α _ #align equiv.equiv_punit Equiv.equivPUnit /-- The `Sort` of proofs of a true proposition is equivalent to `PUnit`. -/ -def propEquivPUnit {p : Prop} (h : p) : p ≃ PUnit.{0} := --- porting note: the following causes an error "unknown free variable [anonymous]" --- @equivPunit p <| uniqueProp h -⟨default, λ _ => h, λ _ => Subsingleton.elim _ _, λ _ => Subsingleton.elim _ _⟩ +def propEquivPUnit {p : Prop} (h : p) : p ≃ PUnit.{0} := @equivPUnit p <| uniqueProp h #align equiv.prop_equiv_punit Equiv.propEquivPUnit /-- `ULift α` is equivalent to `α`. -/ @@ -509,15 +397,10 @@ protected def ulift {α : Type v} : ULift.{u} α ≃ α := /-- `PLift α` is equivalent to `α`. -/ @[simps (config := { fullyApplied := false }) apply] -protected def plift : PLift α ≃ α := - ⟨PLift.down, PLift.up, PLift.up_down, PLift.down_up⟩ +protected def plift : PLift α ≃ α := ⟨PLift.down, PLift.up, PLift.up_down, PLift.down_up⟩ /-- equivalence of propositions is the same as iff -/ -def ofIff {P Q : Prop} (h : P ↔ Q) : P ≃ Q where - toFun := h.mp - invFun := h.mpr - left_inv _ := rfl - right_inv _ := rfl +def ofIff {P Q : Prop} (h : P ↔ Q) : P ≃ Q := ⟨h.mp, h.mpr, fun _ => rfl, fun _ => rfl⟩ /-- If `α₁` is equivalent to `α₂` and `β₁` is equivalent to `β₂`, then the type of maps `α₁ → β₁` is equivalent to the type of maps `α₂ → β₂`. -/ @@ -533,27 +416,19 @@ def arrowCongr {α₁ β₁ α₂ β₂ : Sort _} (e₁ : α₁ ≃ α₂) (e₂ theorem arrowCongr_comp {α₁ β₁ γ₁ α₂ β₂ γ₂ : Sort _} (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) (ec : γ₁ ≃ γ₂) (f : α₁ → β₁) (g : β₁ → γ₁) : arrowCongr ea ec (g ∘ f) = arrowCongr eb ec g ∘ arrowCongr ea eb f := by - ext - simp only [comp, arrowCongr_apply, eb.symm_apply_apply] + ext; simp only [comp, arrowCongr_apply, eb.symm_apply_apply] #align equiv.arrow_congr_comp Equiv.arrowCongr_comp -@[simp] -theorem arrowCongr_refl {α β : Sort _} : - arrowCongr (Equiv.refl α) (Equiv.refl β) = Equiv.refl (α → β) := - rfl +@[simp] theorem arrowCongr_refl {α β : Sort _} : + arrowCongr (Equiv.refl α) (Equiv.refl β) = Equiv.refl (α → β) := rfl #align equiv.arrow_congr_refl Equiv.arrowCongr_refl -@[simp] -theorem arrowCongr_trans {α₁ β₁ α₂ β₂ α₃ β₃ : Sort _} - (e₁ : α₁ ≃ α₂) (e₁' : β₁ ≃ β₂) (e₂ : α₂ ≃ α₃) (e₂' : β₂ ≃ β₃) : - arrowCongr (e₁.trans e₂) (e₁'.trans e₂') = (arrowCongr e₁ e₁').trans (arrowCongr e₂ e₂') := - rfl +@[simp] theorem arrowCongr_trans (e₁ : α₁ ≃ α₂) (e₁' : β₁ ≃ β₂) (e₂ : α₂ ≃ α₃) (e₂' : β₂ ≃ β₃) : + arrowCongr (e₁.trans e₂) (e₁'.trans e₂') = (arrowCongr e₁ e₁').trans (arrowCongr e₂ e₂') := rfl #align equiv.arrow_congr_trans Equiv.arrowCongr_trans -@[simp] -theorem arrowCongr_symm {α₁ β₁ α₂ β₂ : Sort _} (e₁ : α₁ ≃ α₂) (e₂ : β₁ ≃ β₂) : - (arrowCongr e₁ e₂).symm = arrowCongr e₁.symm e₂.symm := - rfl +@[simp] theorem arrowCongr_symm (e₁ : α₁ ≃ α₂) (e₂ : β₁ ≃ β₂) : + (arrowCongr e₁ e₂).symm = arrowCongr e₁.symm e₂.symm := rfl #align equiv.arrow_congr_symm Equiv.arrowCongr_symm /-- A version of `Equiv.arrowCongr` in `Type`, rather than `Sort`. @@ -566,41 +441,29 @@ because Lean's universe rules will not unify `?l_1` with `imax (1 ?m_1)`. def arrowCongr' {α₁ β₁ α₂ β₂ : Type _} (hα : α₁ ≃ α₂) (hβ : β₁ ≃ β₂) : (α₁ → β₁) ≃ (α₂ → β₂) := Equiv.arrowCongr hα hβ -@[simp] -theorem arrowCongr'_refl {α β : Type _} : - arrowCongr' (Equiv.refl α) (Equiv.refl β) = Equiv.refl (α → β) := - rfl +@[simp] theorem arrowCongr'_refl {α β : Type _} : + arrowCongr' (Equiv.refl α) (Equiv.refl β) = Equiv.refl (α → β) := rfl #align equiv.arrow_congr'_refl Equiv.arrowCongr'_refl -@[simp] -theorem arrowCongr'_trans {α₁ β₁ α₂ β₂ α₃ β₃ : Type _} +@[simp] theorem arrowCongr'_trans (e₁ : α₁ ≃ α₂) (e₁' : β₁ ≃ β₂) (e₂ : α₂ ≃ α₃) (e₂' : β₂ ≃ β₃) : arrowCongr' (e₁.trans e₂) (e₁'.trans e₂') = (arrowCongr' e₁ e₁').trans (arrowCongr' e₂ e₂') := rfl #align equiv.arrow_congr'_trans Equiv.arrowCongr'_trans -@[simp] -theorem arrowCongr'_symm {α₁ β₁ α₂ β₂ : Type _} (e₁ : α₁ ≃ α₂) (e₂ : β₁ ≃ β₂) : - (arrowCongr' e₁ e₂).symm = arrowCongr' e₁.symm e₂.symm := - rfl +@[simp] theorem arrowCongr'_symm (e₁ : α₁ ≃ α₂) (e₂ : β₁ ≃ β₂) : + (arrowCongr' e₁ e₂).symm = arrowCongr' e₁.symm e₂.symm := rfl #align equiv.arrow_congr'_symm Equiv.arrowCongr'_symm /-- Conjugate a map `f : α → α` by an equivalence `α ≃ β`. -/ -@[simps apply] -def conj (e : α ≃ β) : (α → α) ≃ (β → β) := - arrowCongr e e +@[simps apply] def conj (e : α ≃ β) : (α → α) ≃ (β → β) := arrowCongr e e -@[simp] -theorem conj_refl : conj (Equiv.refl α) = Equiv.refl (α → α) := - rfl +@[simp] theorem conj_refl : conj (Equiv.refl α) = Equiv.refl (α → α) := rfl -@[simp] -theorem conj_symm (e : α ≃ β) : e.conj.symm = e.symm.conj := - rfl +@[simp] theorem conj_symm (e : α ≃ β) : e.conj.symm = e.symm.conj := rfl -@[simp] -theorem conj_trans (e₁ : α ≃ β) (e₂ : β ≃ γ) : (e₁.trans e₂).conj = e₁.conj.trans e₂.conj := - rfl +@[simp] theorem conj_trans (e₁ : α ≃ β) (e₂ : β ≃ γ) : + (e₁.trans e₂).conj = e₁.conj.trans e₂.conj := rfl -- This should not be a simp lemma as long as `(∘)` is reducible: -- when `(∘)` is reducible, Lean can unify `f₁ ∘ f₂` with any `g` using @@ -622,11 +485,7 @@ theorem symm_comp_eq {α β γ} (e : α ≃ β) (f : γ → α) (g : γ → β) /-- `PUnit` sorts in any two universes are equivalent. -/ def punitEquivPUnit : PUnit.{v} ≃ PUnit.{w} := - ⟨fun _ => PUnit.unit, fun _ => PUnit.unit, fun u => by - cases u - rfl, fun u => by - cases u - rfl⟩ + ⟨fun _ => .unit, fun _ => .unit, fun ⟨⟩ => rfl, fun ⟨⟩ => rfl⟩ #align equiv.punit_equiv_punit Equiv.punitEquivPUnit /-- `Prop` is noncomputably equivalent to `Bool`. -/ @@ -639,66 +498,46 @@ noncomputable def propEquivBool : Prop ≃ Bool where section /-- The sort of maps to `PUnit.{v}` is equivalent to `PUnit.{w}`. -/ -def arrowPUnitEquivPUnit (α : Sort _) : (α → PUnit.{v}) ≃ PUnit.{w} where - toFun _ := PUnit.unit - invFun _ _ := PUnit.unit - left_inv f := by - funext x - cases f x - rfl - right_inv u := by - cases u - rfl +def arrowPUnitEquivPUnit (α : Sort _) : (α → PUnit.{v}) ≃ PUnit.{w} := + ⟨fun _ => .unit, fun _ _ => .unit, fun _ => rfl, fun _ => rfl⟩ #align equiv.arrow_punit_equiv_punit Equiv.arrowPUnitEquivPUnit /-- If `α` is `Subsingleton` and `a : α`, then the type of dependent functions `Π (i : α), β i` is equivalent to `β a`. -/ -@[simps] -def piSubsingleton {α} (β : α → Sort _) [Subsingleton α] (a : α) : (∀ a', β a') ≃ β a where +@[simps] def piSubsingleton (β : α → Sort _) [Subsingleton α] (a : α) : (∀ a', β a') ≃ β a where toFun := eval a invFun x b := cast (congr_arg β <| Subsingleton.elim a b) x - left_inv f := - funext fun b => by - rw [Subsingleton.elim b a] - rfl - right_inv b := rfl + left_inv _ := funext fun b => by rw [Subsingleton.elim b a]; rfl + right_inv _ := rfl /-- If `α` has a unique term, then the type of function `α → β` is equivalent to `β`. -/ @[simps (config := { fullyApplied := false }) apply] -def funUnique (α β) [Unique.{u} α] : (α → β) ≃ β := - piSubsingleton _ default +def funUnique (α β) [Unique.{u} α] : (α → β) ≃ β := piSubsingleton _ default /-- The sort of maps from `PUnit` is equivalent to the codomain. -/ -def punitArrowEquiv (α : Sort _) : (PUnit.{u} → α) ≃ α := - funUnique PUnit.{u} α +def punitArrowEquiv (α : Sort _) : (PUnit.{u} → α) ≃ α := funUnique PUnit.{u} α /-- The sort of maps from `True` is equivalent to the codomain. -/ -def trueArrowEquiv (α : Sort _) : (True → α) ≃ α := - funUnique _ _ +def trueArrowEquiv (α : Sort _) : (True → α) ≃ α := funUnique _ _ /-- The sort of maps from a type that `IsEmpty` is equivalent to `PUnit`. -/ def arrowPUnitOfIsEmpty (α β : Sort _) [IsEmpty α] : (α → β) ≃ PUnit.{u} where toFun _ := PUnit.unit invFun _ := isEmptyElim left_inv _ := funext isEmptyElim - right_inv u := by - cases u - rfl + right_inv _ := rfl #align equiv.arrow_punit_of_is_empy Equiv.arrowPUnitOfIsEmpty /-- The sort of maps from `Empty` is equivalent to `PUnit`. -/ -def emptyArrowEquivPUnit (α : Sort _) : (Empty → α) ≃ PUnit.{u} := - arrowPUnitOfIsEmpty _ _ +def emptyArrowEquivPUnit (α : Sort _) : (Empty → α) ≃ PUnit.{u} := arrowPUnitOfIsEmpty _ _ #align equiv.empty_arrow_equiv_punit Equiv.emptyArrowEquivPUnit /-- The sort of maps from `PEmpty` is equivalent to `PUnit`. -/ -def pemptyArrowEquivPUnit (α : Sort _) : (PEmpty → α) ≃ PUnit.{u} := - arrowPUnitOfIsEmpty _ _ +def pemptyArrowEquivPUnit (α : Sort _) : (PEmpty → α) ≃ PUnit.{u} := arrowPUnitOfIsEmpty _ _ #align equiv.pempty_arrow_equiv_punit Equiv.pemptyArrowEquivPUnit /-- The sort of maps from `False` is equivalent to `PUnit`. -/ -def falseArrowEquivPUnit (α : Sort _) : (False → α) ≃ PUnit.{u} := - arrowPUnitOfIsEmpty _ _ +def falseArrowEquivPUnit (α : Sort _) : (False → α) ≃ PUnit.{u} := arrowPUnitOfIsEmpty _ _ #align equiv.false_arrow_equiv_punit Equiv.falseArrowEquivPUnit end @@ -710,95 +549,76 @@ section def psigmaEquivSigma {α} (β : α → Type _) : (Σ' i, β i) ≃ Σ i, β i where toFun a := ⟨a.1, a.2⟩ invFun a := ⟨a.1, a.2⟩ - left_inv := λ ⟨_, _⟩ => rfl - right_inv := λ ⟨_, _⟩ => rfl + left_inv | ⟨_, _⟩ => rfl + right_inv | ⟨_, _⟩ => rfl /-- A `psigma`-type is equivalent to the corresponding `sigma`-type. -/ @[simps apply symmApply] def psigmaEquivSigmaPLift {α} (β : α → Sort _) : (Σ' i, β i) ≃ Σ i : PLift α, PLift (β i.down) where toFun a := ⟨PLift.up a.1, PLift.up a.2⟩ invFun a := ⟨a.1.down, a.2.down⟩ - left_inv := λ ⟨_, _⟩ => rfl - right_inv := λ ⟨_, _⟩ => rfl + left_inv | ⟨_, _⟩ => rfl + right_inv | ⟨_, _⟩ => rfl #align equiv.psigma_equiv_sigma_plift Equiv.psigmaEquivSigmaPLift /-- A family of equivalences `Π a, β₁ a ≃ β₂ a` generates an equivalence between `Σ' a, β₁ a` and `Σ' a, β₂ a`. -/ @[simps apply] -def psigmaCongrRight {α} {β₁ β₂ : α → Sort _} (F : ∀ a, β₁ a ≃ β₂ a) : (Σ' a, β₁ a) ≃ Σ' a, β₂ a := - ⟨fun a => ⟨a.1, F a.1 a.2⟩, fun a => ⟨a.1, (F a.1).symm a.2⟩, fun ⟨a, b⟩ => - congr_arg (PSigma.mk a) <| symm_apply_apply (F a) b, fun ⟨a, b⟩ => - congr_arg (PSigma.mk a) <| apply_symm_apply (F a) b⟩ +def psigmaCongrRight {β₁ β₂ : α → Sort _} (F : ∀ a, β₁ a ≃ β₂ a) : (Σ' a, β₁ a) ≃ Σ' a, β₂ a where + toFun a := ⟨a.1, F a.1 a.2⟩ + invFun a := ⟨a.1, (F a.1).symm a.2⟩ + left_inv | ⟨a, b⟩ => congr_arg (PSigma.mk a) <| symm_apply_apply (F a) b + right_inv | ⟨a, b⟩ => congr_arg (PSigma.mk a) <| apply_symm_apply (F a) b -- Porting note: simp can now simplify the LHS, so I have removed `@[simp]` theorem psigmaCongrRight_trans {α} {β₁ β₂ β₃ : α → Sort _} (F : ∀ a, β₁ a ≃ β₂ a) (G : ∀ a, β₂ a ≃ β₃ a) : (psigmaCongrRight F).trans (psigmaCongrRight G) = - psigmaCongrRight fun a => (F a).trans (G a) := by - ext1 x - cases x - rfl + psigmaCongrRight fun a => (F a).trans (G a) := rfl #align equiv.psigma_congr_right_trans Equiv.psigmaCongrRight_trans -- Porting note: simp can now simplify the LHS, so I have removed `@[simp]` theorem psigmaCongrRight_symm {α} {β₁ β₂ : α → Sort _} (F : ∀ a, β₁ a ≃ β₂ a) : - (psigmaCongrRight F).symm = psigmaCongrRight fun a => (F a).symm := by - ext1 x - cases x - rfl + (psigmaCongrRight F).symm = psigmaCongrRight fun a => (F a).symm := rfl #align equiv.psigma_congr_right_symm Equiv.psigmaCongrRight_symm -- Porting note: simp can now prove this, so I have removed `@[simp]` theorem psigmaCongrRight_refl {α} {β : α → Sort _} : - (psigmaCongrRight fun a => Equiv.refl (β a)) = Equiv.refl (Σ' a, β a) := by - ext1 x - cases x - rfl + (psigmaCongrRight fun a => Equiv.refl (β a)) = Equiv.refl (Σ' a, β a) := rfl #align equiv.psigma_congr_right_refl Equiv.psigmaCongrRight_refl /-- A family of equivalences `Π a, β₁ a ≃ β₂ a` generates an equivalence between `Σ a, β₁ a` and `Σ a, β₂ a`. -/ @[simps apply] -def sigmaCongrRight {α} {β₁ β₂ : α → Type _} (F : ∀ a, β₁ a ≃ β₂ a) : (Σ a, β₁ a) ≃ Σ a, β₂ a := - ⟨fun a => ⟨a.1, F a.1 a.2⟩, fun a => ⟨a.1, (F a.1).symm a.2⟩, fun ⟨a, b⟩ => - congr_arg (Sigma.mk a) <| symm_apply_apply (F a) b, fun ⟨a, b⟩ => - congr_arg (Sigma.mk a) <| apply_symm_apply (F a) b⟩ +def sigmaCongrRight {α} {β₁ β₂ : α → Type _} (F : ∀ a, β₁ a ≃ β₂ a) : (Σ a, β₁ a) ≃ Σ a, β₂ a where + toFun a := ⟨a.1, F a.1 a.2⟩ + invFun a := ⟨a.1, (F a.1).symm a.2⟩ + left_inv | ⟨a, b⟩ => congr_arg (Sigma.mk a) <| symm_apply_apply (F a) b + right_inv | ⟨a, b⟩ => congr_arg (Sigma.mk a) <| apply_symm_apply (F a) b -- Porting note: simp can now simplify the LHS, so I have removed `@[simp]` theorem sigmaCongrRight_trans {α} {β₁ β₂ β₃ : α → Type _} (F : ∀ a, β₁ a ≃ β₂ a) (G : ∀ a, β₂ a ≃ β₃ a) : - (sigmaCongrRight F).trans (sigmaCongrRight G) = sigmaCongrRight fun a => (F a).trans (G a) := by - ext1 x - cases x - rfl + (sigmaCongrRight F).trans (sigmaCongrRight G) = + sigmaCongrRight fun a => (F a).trans (G a) := rfl #align equiv.sigmaCongrRight Equiv.sigmaCongrRight_trans -- Porting note: simp can now simplify the LHS, so I have removed `@[simp]` theorem sigmaCongrRight_symm {α} {β₁ β₂ : α → Type _} (F : ∀ a, β₁ a ≃ β₂ a) : - (sigmaCongrRight F).symm = sigmaCongrRight fun a => (F a).symm := by - ext1 x - cases x - rfl + (sigmaCongrRight F).symm = sigmaCongrRight fun a => (F a).symm := rfl #align equiv.sigma_congr_right_symm Equiv.sigmaCongrRight_symm -- Porting note: simp can now prove this, so I have removed `@[simp]` theorem sigmaCongrRight_refl {α} {β : α → Type _} : - (sigmaCongrRight fun a => Equiv.refl (β a)) = Equiv.refl (Σ a, β a) := by - ext1 x - cases x - rfl + (sigmaCongrRight fun a => Equiv.refl (β a)) = Equiv.refl (Σ a, β a) := rfl #align equiv.sigma_congr_right_refl Equiv.sigmaCongrRight_refl /-- A `psigma` with `Prop` fibers is equivalent to the subtype. -/ def psigmaEquivSubtype {α : Type v} (P : α → Prop) : (Σ' i, P i) ≃ Subtype P where toFun x := ⟨x.1, x.2⟩ invFun x := ⟨x.1, x.2⟩ - left_inv x := by - cases x - rfl - right_inv x := by - cases x - rfl + left_inv _ := rfl + right_inv _ := rfl /-- A `sigma` with `plift` fibers is equivalent to the subtype. -/ def sigmaPLiftEquivSubtype {α : Type v} (P : α → Prop) : (Σ i, PLift (P i)) ≃ Subtype P := @@ -817,24 +637,21 @@ def sigmaULiftPLiftEquivSubtype {α : Type v} (P : α → Prop) : namespace Perm /-- A family of permutations `Π a, perm (β a)` generates a permuation `perm (Σ a, β₁ a)`. -/ -@[reducible] -def sigmaCongrRight {α} {β : α → Sort _} (F : ∀ a, Perm (β a)) : Perm (Σ a, β a) := +@[reducible] def sigmaCongrRight {α} {β : α → Sort _} (F : ∀ a, Perm (β a)) : Perm (Σ a, β a) := Equiv.sigmaCongrRight F -@[simp] -theorem sigmaCongrRight_trans {α} {β : α → Sort _} (F : ∀ a, Perm (β a)) (G : ∀ a, Perm (β a)) : +@[simp] theorem sigmaCongrRight_trans {α} {β : α → Sort _} + (F : ∀ a, Perm (β a)) (G : ∀ a, Perm (β a)) : (sigmaCongrRight F).trans (sigmaCongrRight G) = sigmaCongrRight fun a => (F a).trans (G a) := Equiv.sigmaCongrRight_trans F G #align equiv.perm.sigma_congr_right_trans Equiv.Perm.sigmaCongrRight_trans -@[simp] -theorem sigmaCongrRight_symm {α} {β : α → Sort _} (F : ∀ a, Perm (β a)) : +@[simp] theorem sigmaCongrRight_symm {α} {β : α → Sort _} (F : ∀ a, Perm (β a)) : (sigmaCongrRight F).symm = sigmaCongrRight fun a => (F a).symm := Equiv.sigmaCongrRight_symm F #align equiv.perm.sigma_congr_right_symm Equiv.Perm.sigmaCongrRight_symm -@[simp] -theorem sigmaCongrRight_refl {α} {β : α → Sort _} : +@[simp] theorem sigmaCongrRight_refl {α} {β : α → Sort _} : (sigmaCongrRight fun a => Equiv.refl (β a)) = Equiv.refl (Σ a, β a) := Equiv.sigmaCongrRight_refl #align equiv.perm.sigma_congr_right_refl Equiv.Perm.sigmaCongrRight_refl @@ -842,25 +659,23 @@ theorem sigmaCongrRight_refl {α} {β : α → Sort _} : end Perm /-- An equivalence `f : α₁ ≃ α₂` generates an equivalence between `Σ a, β (f a)` and `Σ a, β a`. -/ -@[simps apply] -def sigmaCongrLeft {α₁ α₂} {β : α₂ → Sort _} (e : α₁ ≃ α₂) : +@[simps apply] def sigmaCongrLeft {β : α₂ → Sort _} (e : α₁ ≃ α₂) : (Σ a : α₁, β (e a)) ≃ Σ a : α₂, β a where toFun a := ⟨e a.1, a.2⟩ invFun a := ⟨e.symm a.1, (e.right_inv a.1).symm ▸ a.2⟩ -- porting note: this was a pretty gnarly match already, and it got worse after porting left_inv := fun ⟨a, b⟩ => - match (motive := ∀ (a') (h : a' = a), @Sigma.mk _ (β ∘ e) _ ((congr_arg e h.symm) ▸ b) = ⟨a, b⟩) + match (motive := ∀ a' (h : a' = a), Sigma.mk _ (congr_arg e h.symm ▸ b) = ⟨a, b⟩) e.symm (e a), e.left_inv a with | _, rfl => rfl right_inv := fun ⟨a, b⟩ => - match (motive := ∀ (a') (h : a' = a), Sigma.mk a' (h.symm ▸ b) = ⟨a, b⟩) + match (motive := ∀ a' (h : a' = a), Sigma.mk a' (h.symm ▸ b) = ⟨a, b⟩) e (e.symm a), e.apply_symm_apply _ with | _, rfl => rfl /-- Transporting a sigma type through an equivalence of the base -/ def sigmaCongrLeft' {α₁ α₂} {β : α₁ → Sort _} (f : α₁ ≃ α₂) : - (Σ a : α₁, β a) ≃ Σ a : α₂, β (f.symm a) := - (sigmaCongrLeft f.symm).symm + (Σ a : α₁, β a) ≃ Σ a : α₂, β (f.symm a) := (sigmaCongrLeft f.symm).symm /-- Transporting a sigma type through an equivalence of the base and a family of equivalences of matching fibers -/ @@ -869,8 +684,7 @@ def sigmaCongr {α₁ α₂} {β₁ : α₁ → Sort _} {β₂ : α₂ → Sort (sigmaCongrRight F).trans (sigmaCongrLeft f) /-- `sigma` type with a constant fiber is equivalent to the product. -/ -@[simps apply symmApply] -def sigmaEquivProd (α β : Type _) : (Σ _ : α, β) ≃ α × β := +@[simps apply symmApply] def sigmaEquivProd (α β : Type _) : (Σ _ : α, β) ≃ α × β := ⟨fun a => ⟨a.1, a.2⟩, fun a => ⟨a.1, a.2⟩, fun ⟨_, _⟩ => rfl, fun ⟨_, _⟩ => rfl⟩ /-- If each fiber of a `sigma` type is equivalent to a fixed type, then the sigma type @@ -883,8 +697,8 @@ def sigmaAssoc {α : Type _} {β : α → Type _} (γ : ∀ a : α, β a → Typ (Σ ab : Σ a : α, β a, γ ab.1 ab.2) ≃ Σ a : α, Σ b : β a, γ a b where toFun x := ⟨x.1.1, ⟨x.1.2, x.2⟩⟩ invFun x := ⟨⟨x.1, x.2.1⟩, x.2.2⟩ - left_inv := fun ⟨⟨_, _⟩, _⟩ => rfl - right_inv := fun ⟨_, ⟨_, _⟩⟩ => rfl + left_inv _ := rfl + right_inv _ := rfl end @@ -893,31 +707,24 @@ protected theorem exists_unique_congr {p : α → Prop} {q : β → Prop} constructor · rintro ⟨a, ha₁, ha₂⟩ exact ⟨f a, h.1 ha₁, fun b hb => f.symm_apply_eq.1 (ha₂ (f.symm b) (h.2 (by simpa using hb)))⟩ - · rintro ⟨b, hb₁, hb₂⟩ exact ⟨f.symm b, h.2 (by simpa using hb₁), fun y hy => (eq_symm_apply f).2 (hb₂ _ (h.1 hy))⟩ - protected theorem exists_unique_congr_left' {p : α → Prop} (f : α ≃ β) : - (∃! x, p x) ↔ ∃! y, p (f.symm y) := - Equiv.exists_unique_congr f fun {_} => by simp + (∃! x, p x) ↔ ∃! y, p (f.symm y) := Equiv.exists_unique_congr f fun {_} => by simp protected theorem exists_unique_congr_left {p : β → Prop} (f : α ≃ β) : - (∃! x, p (f x)) ↔ ∃! y, p y := - (Equiv.exists_unique_congr_left' f.symm).symm + (∃! x, p (f x)) ↔ ∃! y, p y := (Equiv.exists_unique_congr_left' f.symm).symm protected theorem forall_congr {p : α → Prop} {q : β → Prop} (f : α ≃ β) - (h : ∀ {x}, p x ↔ q (f x)) : (∀ x, p x) ↔ (∀ y, q y) := -by constructor <;> intro h₂ x - . rw [ ← f.right_inv x ] - apply h.mp - apply h₂ - apply h.mpr - apply h₂ + (h : ∀ {x}, p x ↔ q (f x)) : (∀ x, p x) ↔ (∀ y, q y) := by + constructor <;> intro h₂ x + . rw [← f.right_inv x]; apply h.mp; apply h₂ + · apply h.mpr; apply h₂ protected theorem forall_congr' {p : α → Prop} {q : β → Prop} (f : α ≃ β) (h : ∀ {x}, p (f.symm x) ↔ q x) : (∀ x, p x) ↔ ∀ y, q y := - (Equiv.forall_congr f.symm fun {_} => h.symm).symm + (Equiv.forall_congr f.symm h.symm).symm -- We next build some higher arity versions of `equiv.forall_congr`. -- Although they appear to just be repeated applications of `equiv.forall_congr`, @@ -925,41 +732,28 @@ protected theorem forall_congr' {p : α → Prop} {q : β → Prop} (f : α ≃ -- In particular, they are necessary in `equiv_rw`. -- (Stopping at ternary functions seems reasonable: at least in 1-categorical mathematics, -- it's rare to have axioms involving more than 3 elements at once.) -universe ua1 ua2 ub1 ub2 ug1 ug2 - -variable {α₁ : Sort ua1} {α₂ : Sort ua2} {β₁ : Sort ub1} {β₂ : Sort ub2} - {γ₁ : Sort ug1} {γ₂ : Sort ug2} protected theorem forall₂_congr {p : α₁ → β₁ → Prop} {q : α₂ → β₂ → Prop} (eα : α₁ ≃ α₂) - (eβ : β₁ ≃ β₂) (h : ∀ {x y}, p x y ↔ q (eα x) (eβ y)) : (∀ x y, p x y) ↔ ∀ x y, q x y := by - apply Equiv.forall_congr - intros - apply Equiv.forall_congr - intros - apply h + (eβ : β₁ ≃ β₂) (h : ∀ {x y}, p x y ↔ q (eα x) (eβ y)) : (∀ x y, p x y) ↔ ∀ x y, q x y := + Equiv.forall_congr _ <| Equiv.forall_congr _ h protected theorem forall₂_congr' {p : α₁ → β₁ → Prop} {q : α₂ → β₂ → Prop} (eα : α₁ ≃ α₂) (eβ : β₁ ≃ β₂) (h : ∀ {x y}, p (eα.symm x) (eβ.symm y) ↔ q x y) : - (∀ x y, p x y) ↔ ∀ x y, q x y := - (Equiv.forall₂_congr eα.symm eβ.symm fun {_ _} => h.symm).symm + (∀ x y, p x y) ↔ ∀ x y, q x y := (Equiv.forall₂_congr eα.symm eβ.symm h.symm).symm protected theorem forall₃_congr {p : α₁ → β₁ → γ₁ → Prop} {q : α₂ → β₂ → γ₂ → Prop} (eα : α₁ ≃ α₂) (eβ : β₁ ≃ β₂) (eγ : γ₁ ≃ γ₂) (h : ∀ {x y z}, p x y z ↔ q (eα x) (eβ y) (eγ z)) : - (∀ x y z, p x y z) ↔ ∀ x y z, q x y z := by - apply Equiv.forall₂_congr - intros - apply Equiv.forall_congr - intros - apply h + (∀ x y z, p x y z) ↔ ∀ x y z, q x y z := + Equiv.forall₂_congr _ _ <| Equiv.forall_congr _ h protected theorem forall₃_congr' {p : α₁ → β₁ → γ₁ → Prop} {q : α₂ → β₂ → γ₂ → Prop} (eα : α₁ ≃ α₂) (eβ : β₁ ≃ β₂) (eγ : γ₁ ≃ γ₂) (h : ∀ {x y z}, p (eα.symm x) (eβ.symm y) (eγ.symm z) ↔ q x y z) : (∀ x y z, p x y z) ↔ ∀ x y z, q x y z := - (Equiv.forall₃_congr eα.symm eβ.symm eγ.symm fun {_ _ _} => h.symm).symm + (Equiv.forall₃_congr eα.symm eβ.symm eγ.symm h.symm).symm protected theorem forall_congr_left' {p : α → Prop} (f : α ≃ β) : (∀ x, p x) ↔ ∀ y, p (f.symm y) := - Equiv.forall_congr f fun {_} => by simp + Equiv.forall_congr f <| by simp protected theorem forall_congr_left {p : β → Prop} (f : α ≃ β) : (∀ x, p (f x)) ↔ ∀ y, p y := (Equiv.forall_congr_left' f.symm).symm @@ -977,35 +771,26 @@ if `ra a₁ a₂ ↔ rb (e a₁) (e a₂). -/ protected def congr {ra : α → α → Prop} {rb : β → β → Prop} (e : α ≃ β) (eq : ∀ a₁ a₂, ra a₁ a₂ ↔ rb (e a₁) (e a₂)) : Quot ra ≃ Quot rb where toFun := Quot.map e fun a₁ a₂ => (eq a₁ a₂).1 - invFun := - Quot.map e.symm fun b₁ b₂ h => - (eq (e.symm b₁) (e.symm b₂)).2 - ((e.apply_symm_apply b₁).symm ▸ (e.apply_symm_apply b₂).symm ▸ h) - left_inv := by - rintro ⟨a⟩ - simp only [Quot.map, Equiv.symm_apply_apply] - right_inv := by - rintro ⟨a⟩ - simp only [Quot.map, Equiv.apply_symm_apply] - -@[simp] -theorem congr_mk {ra : α → α → Prop} {rb : β → β → Prop} (e : α ≃ β) + invFun := Quot.map e.symm fun b₁ b₂ h => + (eq (e.symm b₁) (e.symm b₂)).2 + ((e.apply_symm_apply b₁).symm ▸ (e.apply_symm_apply b₂).symm ▸ h) + left_inv := by rintro ⟨a⟩; simp only [Quot.map, Equiv.symm_apply_apply] + right_inv := by rintro ⟨a⟩; simp only [Quot.map, Equiv.apply_symm_apply] + +@[simp] theorem congr_mk {ra : α → α → Prop} {rb : β → β → Prop} (e : α ≃ β) (eq : ∀ a₁ a₂ : α, ra a₁ a₂ ↔ rb (e a₁) (e a₂)) (a : α) : - Quot.congr e eq (Quot.mk ra a) = Quot.mk rb (e a) := - rfl + Quot.congr e eq (Quot.mk ra a) = Quot.mk rb (e a) := rfl /-- Quotients are congruent on equivalences under equality of their relation. An alternative is just to use rewriting with `eq`, but then computational proofs get stuck. -/ protected def congrRight {r r' : α → α → Prop} (eq : ∀ a₁ a₂, r a₁ a₂ ↔ r' a₁ a₂) : - Quot r ≃ Quot r' := - Quot.congr (Equiv.refl α) eq + Quot r ≃ Quot r' := Quot.congr (Equiv.refl α) eq /-- An equivalence `e : α ≃ β` generates an equivalence between the quotient space of `α` by a relation `ra` and the quotient space of `β` by the image of this relation under `e`. -/ protected def congrLeft {r : α → α → Prop} (e : α ≃ β) : Quot r ≃ Quot fun b b' => r (e.symm b) (e.symm b') := - @Quot.congr α β r (fun b b' => r (e.symm b) (e.symm b')) e - fun a₁ a₂ => by simp only [e.symm_apply_apply]; rfl + Quot.congr e fun _ _ => by simp only [e.symm_apply_apply]; rfl end Quot @@ -1015,14 +800,11 @@ namespace Quotient if `ra a₁ a₂ ↔ rb (e a₁) (e a₂). -/ protected def congr {ra : Setoid α} {rb : Setoid β} (e : α ≃ β) (eq : ∀ a₁ a₂, @Setoid.r α ra a₁ a₂ ↔ @Setoid.r β rb (e a₁) (e a₂)) : - Quotient ra ≃ Quotient rb := - Quot.congr e eq + Quotient ra ≃ Quotient rb := Quot.congr e eq -@[simp] -theorem congr_mk {ra : Setoid α} {rb : Setoid β} (e : α ≃ β) +@[simp] theorem congr_mk {ra : Setoid α} {rb : Setoid β} (e : α ≃ β) (eq : ∀ a₁ a₂ : α, Setoid.r a₁ a₂ ↔ Setoid.r (e a₁) (e a₂)) (a : α) : - Quotient.congr e eq (Quotient.mk ra a) = Quotient.mk rb (e a) := - rfl + Quotient.congr e eq (Quotient.mk ra a) = Quotient.mk rb (e a) := rfl /-- Quotients are congruent on equivalences under equality of their relation. An alternative is just to use rewriting with `eq`, but then computational proofs get stuck. -/ From ac33b9d27259ee1582b2f4491743559df335d454 Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Fri, 18 Nov 2022 19:14:12 +0000 Subject: [PATCH 032/127] add Logic.Equiv.Basic to Mathlib.lean --- Mathlib.lean | 1 + 1 file changed, 1 insertion(+) diff --git a/Mathlib.lean b/Mathlib.lean index 18e74d824fe2d..555ceb63d9d9e 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -96,6 +96,7 @@ import Mathlib.Lean.Expr.Traverse import Mathlib.Lean.LocalContext import Mathlib.Lean.Meta import Mathlib.Logic.Basic +import Mathlib.Logic.Equiv.Basic import Mathlib.Logic.Equiv.Defs import Mathlib.Logic.Equiv.LocalEquiv import Mathlib.Logic.Equiv.MfldSimpsAttr From c387d5303b84e1fbb1647f9fcb43e50b0c0038c5 Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Fri, 18 Nov 2022 19:14:23 +0000 Subject: [PATCH 033/127] epsilon --- Mathlib/Logic/Equiv/Basic.lean | 5 +++-- 1 file changed, 3 insertions(+), 2 deletions(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index f5365d0b090a2..7db8b999ac3e2 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -121,7 +121,8 @@ theorem prod_comm_symm (α β) : (prodComm α β).symm = prodComm β α := /-- Type product is associative up to an equivalence. -/ @[simps] def prodAssoc (α β γ : Sort _) : (α × β) × γ ≃ α × β × γ := - ⟨fun p => (p.1.1, p.1.2, p.2), fun p => ((p.1, p.2.1), p.2.2), fun ⟨⟨a, b⟩, c⟩ => rfl, fun ⟨a, ⟨b, c⟩⟩ => rfl⟩ + ⟨fun p => (p.1.1, p.1.2, p.2), fun p => ((p.1, p.2.1), p.2.2), fun ⟨⟨_, _⟩, _⟩ => rfl, + fun ⟨_, ⟨_, _⟩⟩ => rfl⟩ #align equiv.prod_assoc Equiv.prodAssoc /-- Functions on `α × β` are equivalent to functions `α → β → γ`. -/ @@ -138,7 +139,7 @@ section /-- `punit` is a right identity for type product up to an equivalence. -/ @[simps] def prodPunit (α : Type _) : α × PUnit.{u + 1} ≃ α := - ⟨fun p => p.1, fun a => (a, PUnit.unit), fun ⟨_, PUnit.unit⟩ => rfl, fun a => rfl⟩ + ⟨fun p => p.1, fun a => (a, PUnit.unit), fun ⟨_, PUnit.unit⟩ => rfl, fun _ => rfl⟩ #align equiv.prod_punit Equiv.prodPunit /-- `punit` is a left identity for type product up to an equivalence. -/ From f935c141666c0b7e9fc1e68e1dfe6eef08938174 Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Fri, 18 Nov 2022 19:35:54 +0000 Subject: [PATCH 034/127] epsilon more --- Mathlib/Logic/Equiv/Basic.lean | 21 +++++++++++++-------- 1 file changed, 13 insertions(+), 8 deletions(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 7db8b999ac3e2..09edded656c33 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -10,6 +10,9 @@ import Mathlib.Data.Subtype import Mathlib.Data.Sum.Basic import Mathlib.Logic.Function.Conjugate +-- **TODO** remove these later +set_option autoImplicit false + /-! # Equivalence between types @@ -86,7 +89,8 @@ def pprodEquivProdPlift {α β : Sort _} : PProd α β ≃ PLift α × PLift β /-- Product of two equivalences. If `α₁ ≃ α₂` and `β₁ ≃ β₂`, then `α₁ × β₁ ≃ α₂ × β₂`. This is `prod.map` as an equivalence. -/ --@[congr, simps apply] ---porting note: NEEDS FIXING +--**TODO** need to sort out those attributes +--@[simps apply] def prodCongr {α₁ β₁ α₂ β₂ : Type _} (e₁ : α₁ ≃ α₂) (e₂ : β₁ ≃ β₂) : α₁ × β₁ ≃ α₂ × β₂ := ⟨Prod.map e₁ e₂, Prod.map e₁.symm e₂.symm, fun ⟨a, b⟩ => by simp, fun ⟨a, b⟩ => by simp⟩ #align equiv.prod_congr Equiv.prodCongr @@ -138,26 +142,27 @@ section /-- `punit` is a right identity for type product up to an equivalence. -/ @[simps] -def prodPunit (α : Type _) : α × PUnit.{u + 1} ≃ α := +def prodPUnit (α : Type _) : α × PUnit.{u + 1} ≃ α := ⟨fun p => p.1, fun a => (a, PUnit.unit), fun ⟨_, PUnit.unit⟩ => rfl, fun _ => rfl⟩ #align equiv.prod_punit Equiv.prodPunit /-- `punit` is a left identity for type product up to an equivalence. -/ @[simps] -def punitProd (α : Type _) : PUnit.{u + 1} × α ≃ α := +def pUnitProd (α : Type _) : PUnit.{u + 1} × α ≃ α := calc PUnit × α ≃ α × PUnit := prodComm _ _ - _ ≃ α := prodPunit _ + _ ≃ α := prodPUnit _ #align equiv.punit_prod Equiv.punitProd +-- **TODO** fix weird universe error /-- Any `unique` type is a right identity for type product up to equivalence. -/ -def prodUnique (α β : Type _) [Unique β] : α × β ≃ α := - ((Equiv.refl α).prodCongr <| equivPunit β).trans <| prodPunit α +def prodUnique (α : Type u) (β : Type v) [Unique β] : α × β ≃ α := + ((Equiv.refl α).prodCongr <| equivPUnit β).trans <| prodPUnit α #align equiv.prod_unique Equiv.prodUnique @[simp] -theorem coe_prod_unique {α β : Type _} [Unique β] : ⇑(prodUnique α β) = Prod.fst := +theorem coe_prod_unique {α β : Type _} [Unique β] : (⇑(prodUnique α β) : α × β → α) = Prod.fst := rfl #align equiv.coe_prod_unique Equiv.coe_prod_unique @@ -188,7 +193,7 @@ theorem unique_prod_apply {α β : Type _} [Unique β] (x : β × α) : uniquePr theorem unique_prod_symm_apply {α β : Type _} [Unique β] (x : α) : (uniqueProd α β).symm x = (default, x) := rfl #align equiv.unique_prod_symm_apply Equiv.unique_prod_symm_apply - +#exit /-- `empty` type is a right absorbing element for type product up to an equivalence. -/ def prodEmpty (α : Type _) : α × Empty ≃ Empty := equivEmpty _ From d3af30195d14f331554415dbdddc3100b180f012 Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Fri, 18 Nov 2022 20:37:49 +0000 Subject: [PATCH 035/127] epsilon more --- Mathlib/Logic/Equiv/Basic.lean | 30 +++++++++++++++--------------- 1 file changed, 15 insertions(+), 15 deletions(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 09edded656c33..0159acbfafdbf 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -59,8 +59,8 @@ def pprodEquivProd {α β : Type _} : PProd α β ≃ α × β where /-- Product of two equivalences, in terms of `pprod`. If `α ≃ β` and `γ ≃ δ`, then `pprod α γ ≃ pprod β δ`. -/ ---@[congr, simps apply] ---porting note: NEEDS FIXING +-- porting note: in Lean 3 this had @[congr]` +@[simps apply] def pprodCongr {δ : Sort z} (e₁ : α ≃ β) (e₂ : γ ≃ δ) : PProd α γ ≃ PProd β δ where toFun x := ⟨e₁ x.1, e₂ x.2⟩ invFun x := ⟨e₁.symm x.1, e₂.symm x.2⟩ @@ -76,7 +76,7 @@ def pprodProd {α₁ β₁ : Sort _} {α₂ β₂ : Type _} (ea : α₁ ≃ α /-- Combine two equivalences using `pprod` in the codomain and `prod` in the domain. -/ @[simps apply symmApply] -def prodPprod {α₁ β₁ : Type _} {α₂ β₂ : Sort _} (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) : α₁ × β₁ ≃ PProd α₂ β₂ := +def prodPProd {α₁ β₁ : Type _} {α₂ β₂ : Sort _} (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) : α₁ × β₁ ≃ PProd α₂ β₂ := (ea.symm.pprodProd eb.symm).symm #align equiv.prod_pprod Equiv.prodPprod @@ -88,9 +88,8 @@ def pprodEquivProdPlift {α β : Sort _} : PProd α β ≃ PLift α × PLift β /-- Product of two equivalences. If `α₁ ≃ α₂` and `β₁ ≃ β₂`, then `α₁ × β₁ ≃ α₂ × β₂`. This is `prod.map` as an equivalence. -/ ---@[congr, simps apply] ---**TODO** need to sort out those attributes ---@[simps apply] +-- porting note: in Lean 3 there was also a @[congr] tag +@[simps apply] def prodCongr {α₁ β₁ α₂ β₂ : Type _} (e₁ : α₁ ≃ α₂) (e₂ : β₁ ≃ β₂) : α₁ × β₁ ≃ α₂ × β₂ := ⟨Prod.map e₁ e₂, Prod.map e₁.symm e₂.symm, fun ⟨a, b⟩ => by simp, fun ⟨a, b⟩ => by simp⟩ #align equiv.prod_congr Equiv.prodCongr @@ -155,10 +154,9 @@ def pUnitProd (α : Type _) : PUnit.{u + 1} × α ≃ α := #align equiv.punit_prod Equiv.punitProd --- **TODO** fix weird universe error /-- Any `unique` type is a right identity for type product up to equivalence. -/ def prodUnique (α : Type u) (β : Type v) [Unique β] : α × β ≃ α := - ((Equiv.refl α).prodCongr <| equivPUnit β).trans <| prodPUnit α + ((Equiv.refl α).prodCongr <| equivPUnit.{v+1,v+1} β).trans <| prodPUnit α #align equiv.prod_unique Equiv.prodUnique @[simp] @@ -177,11 +175,11 @@ theorem prod_unique_symm_apply {α β : Type _} [Unique β] (x : α) : (prodUniq /-- Any `unique` type is a left identity for type product up to equivalence. -/ def uniqueProd (α β : Type _) [Unique β] : β × α ≃ α := - ((equivPunit β).prodCongr <| Equiv.refl α).trans <| punitProd α + ((equivPUnit.{v+1,v+1} β).prodCongr <| Equiv.refl α).trans <| pUnitProd α #align equiv.unique_prod Equiv.uniqueProd @[simp] -theorem coe_unique_prod {α β : Type _} [Unique β] : ⇑(uniqueProd α β) = Prod.snd := +theorem coe_unique_prod {α β : Type _} [Unique β] : (⇑(uniqueProd α β) : β × α → α) = Prod.snd := rfl #align equiv.coe_unique_prod Equiv.coe_unique_prod @@ -190,10 +188,11 @@ theorem unique_prod_apply {α β : Type _} [Unique β] (x : β × α) : uniquePr #align equiv.unique_prod_apply Equiv.unique_prod_apply @[simp] -theorem unique_prod_symm_apply {α β : Type _} [Unique β] (x : α) : (uniqueProd α β).symm x = (default, x) := +theorem unique_prod_symm_apply {α β : Type _} [Unique β] (x : α) : + (uniqueProd α β).symm x = (default, x) := rfl #align equiv.unique_prod_symm_apply Equiv.unique_prod_symm_apply -#exit + /-- `empty` type is a right absorbing element for type product up to an equivalence. -/ def prodEmpty (α : Type _) : α × Empty ≃ Empty := equivEmpty _ @@ -206,12 +205,12 @@ def emptyProd (α : Type _) : Empty × α ≃ Empty := /-- `pempty` type is a right absorbing element for type product up to an equivalence. -/ def prodPempty (α : Type _) : α × PEmpty ≃ PEmpty := - equivPempty _ + equivPEmpty _ #align equiv.prod_pempty Equiv.prodPempty /-- `pempty` type is a left absorbing element for type product up to an equivalence. -/ def pemptyProd (α : Type _) : PEmpty × α ≃ PEmpty := - equivPempty _ + equivPEmpty _ #align equiv.pempty_prod Equiv.pemptyProd end @@ -282,7 +281,7 @@ def sumCongr {α β : Type _} (ea : Equiv.Perm α) (eb : Equiv.Perm β) : Equiv. @[simp] theorem sum_congr_apply {α β : Type _} (ea : Equiv.Perm α) (eb : Equiv.Perm β) (x : Sum α β) : sumCongr ea eb x = Sum.map (⇑ea) (⇑eb) x := - Equiv.sum_congr_apply ea eb x + Equiv.sumCongr_apply ea eb x #align equiv.perm.sum_congr_apply Equiv.Perm.sum_congr_apply @[simp] @@ -393,6 +392,7 @@ def optionEquivSumPunit (α : Type _) : Option α ≃ Sum α PUnit.{u + 1} := theorem option_equiv_sum_punit_none {α} : optionEquivSumPunit α none = Sum.inr PUnit.unit := rfl #align equiv.option_equiv_sum_punit_none Equiv.option_equiv_sum_punit_none +#exit @[simp] theorem option_equiv_sum_punit_some {α} (a) : optionEquivSumPunit α (some a) = Sum.inl a := From 8c43b711b89db194978d04eb959413daebb61f2c Mon Sep 17 00:00:00 2001 From: 77Tigers <36101478+77Tigers@users.noreply.github.com> Date: Fri, 18 Nov 2022 21:18:12 +0000 Subject: [PATCH 036/127] Update Mathlib/Logic/Equiv/Basic.lean Co-authored-by: Scott Morrison --- Mathlib/Logic/Equiv/Basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 09edded656c33..fa101a46d56bd 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -16,7 +16,7 @@ set_option autoImplicit false /-! # Equivalence between types -In this file we continue the work on equivalences begun in `logic/equiv/defs.lean`, defining +In this file we continue the work on equivalences begun in `Logic/Equiv/Defs.lean`, defining * canonical isomorphisms between various types: e.g., From f209c6c596639f1f7cca87d9d473890daef77aa9 Mon Sep 17 00:00:00 2001 From: ruben vorster Date: Fri, 18 Nov 2022 21:48:16 +0000 Subject: [PATCH 037/127] satisfy long line linter (mostly) --- Mathlib/Logic/Equiv/Basic.lean | 275 +++++++++++++++++++++------------ 1 file changed, 176 insertions(+), 99 deletions(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index fa101a46d56bd..c3533573537df 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -2,6 +2,7 @@ Copyright (c) 2015 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Mario Carneiro +Ported by: Kevin Buzzard, Ruben Vorster -/ import Mathlib.Logic.Equiv.Defs import Mathlib.Data.Prod.Basic @@ -70,13 +71,15 @@ def pprodCongr {δ : Sort z} (e₁ : α ≃ β) (e₂ : γ ≃ δ) : PProd α γ /-- Combine two equivalences using `pprod` in the domain and `prod` in the codomain. -/ @[simps apply symmApply] -def pprodProd {α₁ β₁ : Sort _} {α₂ β₂ : Type _} (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) : PProd α₁ β₁ ≃ α₂ × β₂ := +def pprodProd {α₁ β₁ : Sort _} {α₂ β₂ : Type _} (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) +: PProd α₁ β₁ ≃ α₂ × β₂ := (ea.pprodCongr eb).trans pprodEquivProd #align equiv.pprod_prod Equiv.pprodProd /-- Combine two equivalences using `pprod` in the codomain and `prod` in the domain. -/ @[simps apply symmApply] -def prodPprod {α₁ β₁ : Type _} {α₂ β₂ : Sort _} (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) : α₁ × β₁ ≃ PProd α₂ β₂ := +def prodPprod {α₁ β₁ : Type _} {α₂ β₂ : Sort _} (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) +: α₁ × β₁ ≃ PProd α₂ β₂ := (ea.symm.pprodProd eb.symm).symm #align equiv.prod_pprod Equiv.prodPprod @@ -171,7 +174,8 @@ theorem prod_unique_apply {α β : Type _} [Unique β] (x : α × β) : prodUniq #align equiv.prod_unique_apply Equiv.prod_unique_apply @[simp] -theorem prod_unique_symm_apply {α β : Type _} [Unique β] (x : α) : (prodUnique α β).symm x = (x, default) := +theorem prod_unique_symm_apply {α β : Type _} [Unique β] (x : α) +: (prodUnique α β).symm x = (x, default) := rfl #align equiv.prod_unique_symm_apply Equiv.prod_unique_symm_apply @@ -190,7 +194,8 @@ theorem unique_prod_apply {α β : Type _} [Unique β] (x : β × α) : uniquePr #align equiv.unique_prod_apply Equiv.unique_prod_apply @[simp] -theorem unique_prod_symm_apply {α β : Type _} [Unique β] (x : α) : (uniqueProd α β).symm x = (default, x) := +theorem unique_prod_symm_apply {α β : Type _} [Unique β] (x : α) +: (uniqueProd α β).symm x = (default, x) := rfl #align equiv.unique_prod_symm_apply Equiv.unique_prod_symm_apply #exit @@ -243,17 +248,20 @@ def psumCongr {δ : Sort z} (e₁ : α ≃ β) (e₂ : γ ≃ δ) : PSum α γ #align equiv.psum_congr Equiv.psumCongr /-- Combine two `equiv`s using `psum` in the domain and `sum` in the codomain. -/ -def psumSum {α₁ β₁ : Sort _} {α₂ β₂ : Type _} (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) : PSum α₁ β₁ ≃ Sum α₂ β₂ := +def psumSum {α₁ β₁ : Sort _} {α₂ β₂ : Type _} (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) : + PSum α₁ β₁ ≃ Sum α₂ β₂ := (ea.psumCongr eb).trans (psumEquivSum _ _) #align equiv.psum_sum Equiv.psumSum /-- Combine two `equiv`s using `sum` in the domain and `psum` in the codomain. -/ -def sumPsum {α₁ β₁ : Type _} {α₂ β₂ : Sort _} (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) : Sum α₁ β₁ ≃ PSum α₂ β₂ := +def sumPsum {α₁ β₁ : Type _} {α₂ β₂ : Sort _} (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) : + Sum α₁ β₁ ≃ PSum α₂ β₂ := (ea.symm.psumSum eb.symm).symm #align equiv.sum_psum Equiv.sumPsum @[simp] -theorem sum_congr_trans {α₁ α₂ β₁ β₂ γ₁ γ₂ : Sort _} (e : α₁ ≃ β₁) (f : α₂ ≃ β₂) (g : β₁ ≃ γ₁) (h : β₂ ≃ γ₂) : +theorem sum_congr_trans {α₁ α₂ β₁ β₂ γ₁ γ₂ : Sort _} +(e : α₁ ≃ β₁) (f : α₂ ≃ β₂) (g : β₁ ≃ γ₁) (h : β₂ ≃ γ₂) : (Equiv.sumCongr e f).trans (Equiv.sumCongr g h) = Equiv.sumCongr (e.trans g) (f.trans h) := by ext i cases i <;> rfl @@ -266,7 +274,8 @@ theorem sum_congr_symm {α β γ δ : Sort _} (e : α ≃ β) (f : γ ≃ δ) : #align equiv.sum_congr_symm Equiv.sum_congr_symm @[simp] -theorem sum_congr_refl {α β : Sort _} : Equiv.sumCongr (Equiv.refl α) (Equiv.refl β) = Equiv.refl (Sum α β) := by +theorem sum_congr_refl {α β : Sort _} : + Equiv.sumCongr (Equiv.refl α) (Equiv.refl β) = Equiv.refl (Sum α β) := by ext i cases i <;> rfl #align equiv.sum_congr_refl Equiv.sum_congr_refl @@ -286,7 +295,8 @@ theorem sum_congr_apply {α β : Type _} (ea : Equiv.Perm α) (eb : Equiv.Perm #align equiv.perm.sum_congr_apply Equiv.Perm.sum_congr_apply @[simp] -theorem sum_congr_trans {α β : Sort _} (e : Equiv.Perm α) (f : Equiv.Perm β) (g : Equiv.Perm α) (h : Equiv.Perm β) : +theorem sum_congr_trans {α β : Sort _} +(e : Equiv.Perm α) (f : Equiv.Perm β) (g : Equiv.Perm α) (h : Equiv.Perm β) : (sumCongr e f).trans (sumCongr g h) = sumCongr (e.trans g) (f.trans h) := Equiv.sum_congr_trans e f g h #align equiv.perm.sum_congr_trans Equiv.Perm.sum_congr_trans @@ -298,7 +308,8 @@ theorem sum_congr_symm {α β : Sort _} (e : Equiv.Perm α) (f : Equiv.Perm β) #align equiv.perm.sum_congr_symm Equiv.Perm.sum_congr_symm @[simp] -theorem sum_congr_refl {α β : Sort _} : sumCongr (Equiv.refl α) (Equiv.refl β) = Equiv.refl (Sum α β) := +theorem sum_congr_refl {α β : Sort _} : + sumCongr (Equiv.refl α) (Equiv.refl β) = Equiv.refl (Sum α β) := Equiv.sum_congr_refl #align equiv.perm.sum_congr_refl Equiv.Perm.sum_congr_refl @@ -306,8 +317,8 @@ end Perm /-- `bool` is equivalent the sum of two `punit`s. -/ def boolEquivPunitSumPunit : Bool ≃ Sum PUnit.{u + 1} PUnit.{v + 1} := - ⟨fun b => cond b (inr PUnit.unit) (inl PUnit.unit), Sum.elim (fun _ => false) fun _ => true, fun b => by - cases b <;> rfl, fun s => by rcases s with (⟨⟨⟩⟩ | ⟨⟨⟩⟩) <;> rfl⟩ + ⟨fun b => cond b (inr PUnit.unit) (inl PUnit.unit), Sum.elim (fun _ => false) fun _ => true, + fun b => by cases b <;> rfl, fun s => by rcases s with (⟨⟨⟩⟩ | ⟨⟨⟩⟩) <;> rfl⟩ #align equiv.bool_equiv_punit_sum_punit Equiv.boolEquivPunitSumPunit /-- Sum of types is commutative up to an equivalence. This is `sum.swap` as an equivalence. -/ @@ -324,7 +335,8 @@ theorem sum_comm_symm (α β) : (sumComm α β).symm = sumComm β α := /-- Sum of types is associative up to an equivalence. -/ def sumAssoc (α β γ : Type _) : Sum (Sum α β) γ ≃ Sum α (Sum β γ) := ⟨Sum.elim (Sum.elim Sum.inl (Sum.inr ∘ Sum.inl)) (Sum.inr ∘ Sum.inr), - Sum.elim (Sum.inl ∘ Sum.inl) <| Sum.elim (Sum.inl ∘ Sum.inr) Sum.inr, by rintro (⟨_ | _⟩ | _) <;> rfl, by + Sum.elim (Sum.inl ∘ Sum.inl) <| Sum.elim (Sum.inl ∘ Sum.inr) Sum.inr, + by rintro (⟨_ | _⟩ | _) <;> rfl, by rintro (_ | ⟨_ | _⟩) <;> rfl⟩ #align equiv.sum_assoc Equiv.sumAssoc @@ -349,7 +361,8 @@ theorem sum_assoc_symm_apply_inl {α β γ} (a) : (sumAssoc α β γ).symm (inl #align equiv.sum_assoc_symm_apply_inl Equiv.sum_assoc_symm_apply_inl @[simp] -theorem sum_assoc_symm_apply_inr_inl {α β γ} (b) : (sumAssoc α β γ).symm (inr (inl b)) = inl (inr b) := +theorem sum_assoc_symm_apply_inr_inl {α β γ} (b) : + (sumAssoc α β γ).symm (inr (inl b)) = inl (inr b) := rfl #align equiv.sum_assoc_symm_apply_inr_inl Equiv.sum_assoc_symm_apply_inr_inl @@ -385,8 +398,9 @@ theorem empty_sum_apply_inr {α β : Type _} [IsEmpty α] (b : β) : emptySum α /-- `option α` is equivalent to `α ⊕ punit` -/ def optionEquivSumPunit (α : Type _) : Option α ≃ Sum α PUnit.{u + 1} := - ⟨fun o => o.elim (inr PUnit.unit) inl, fun s => s.elim some fun _ => none, fun o => by cases o <;> rfl, fun s => by - rcases s with (_ | ⟨⟨⟩⟩) <;> rfl⟩ + ⟨fun o => o.elim (inr PUnit.unit) inl, fun s => s.elim some fun _ => none, + fun o => by cases o <;> rfl, fun s => by + rcases s with (_ | ⟨⟨⟩⟩) <;> rfl⟩ #align equiv.option_equiv_sum_punit Equiv.optionEquivSumPunit @[simp] @@ -426,7 +440,8 @@ def optionIsSomeEquiv (α : Type _) : { x : Option α // x.isSome } ≃ α where /-- The product over `option α` of `β a` is the binary product of the product over `α` of `β (some α)` and `β none` -/ @[simps] -def piOptionEquivProd {α : Type _} {β : Option α → Type _} : (∀ a : Option α, β a) ≃ β none × ∀ a : α, β (some a) where +def piOptionEquivProd {α : Type _} {β : Option α → Type _} : + (∀ a : Option α, β a) ≃ β none × ∀ a : α, β (some a) where toFun f := (f none, fun a => f (some a)) invFun x a := Option.casesOn a x.fst x.snd left_inv f := funext fun a => by cases a <;> rfl @@ -463,7 +478,8 @@ is naturally equivalent to `α`. See `subtype_or_equiv` for sum types over subtypes `{x // p x}` and `{x // q x}` that are not necessarily `is_compl p q`. -/ -def sumCompl {α : Type _} (p : α → Prop) [DecidablePred p] : Sum { a // p a } { a // ¬p a } ≃ α where +def sumCompl {α : Type _} (p : α → Prop) [DecidablePred p] : + Sum { a // p a } { a // ¬p a } ≃ α where toFun := Sum.elim coe coe invFun a := if h : p a then Sum.inl ⟨a, h⟩ else Sum.inr ⟨a, h⟩ left_inv := by rintro (⟨x, hx⟩ | ⟨x, hx⟩) <;> dsimp <;> [rw [dif_pos], rw [dif_neg]] @@ -485,20 +501,23 @@ theorem sum_compl_apply_inr {α : Type _} (p : α → Prop) [DecidablePred p] (x #align equiv.sum_compl_apply_inr Equiv.sum_compl_apply_inr @[simp] -theorem sum_compl_apply_symm_of_pos {α : Type _} (p : α → Prop) [DecidablePred p] (a : α) (h : p a) : +theorem sum_compl_apply_symm_of_pos {α : Type _} (p : α → Prop) [DecidablePred p] (a : α) + (h : p a) : (sumCompl p).symm a = Sum.inl ⟨a, h⟩ := dif_pos h #align equiv.sum_compl_apply_symm_of_pos Equiv.sum_compl_apply_symm_of_pos @[simp] -theorem sum_compl_apply_symm_of_neg {α : Type _} (p : α → Prop) [DecidablePred p] (a : α) (h : ¬p a) : +theorem sum_compl_apply_symm_of_neg {α : Type _} (p : α → Prop) [DecidablePred p] (a : α) + (h : ¬p a) : (sumCompl p).symm a = Sum.inr ⟨a, h⟩ := dif_neg h #align equiv.sum_compl_apply_symm_of_neg Equiv.sum_compl_apply_symm_of_neg /-- Combines an `equiv` between two subtypes with an `equiv` between their complements to form a permutation. -/ -def subtypeCongr {α : Type _} {p q : α → Prop} [DecidablePred p] [DecidablePred q] (e : { x // p x } ≃ { x // q x }) +def subtypeCongr {α : Type _} {p q : α → Prop} [DecidablePred p] [DecidablePred q] + (e : { x // p x } ≃ { x // q x }) (f : { x // ¬p x } ≃ { x // ¬q x }) : Perm α := (sumCompl p).symm.trans ((sumCongr e f).trans (sumCompl q)) #align equiv.subtype_congr Equiv.subtypeCongr @@ -515,7 +534,8 @@ def Perm.subtypeCongr : Equiv.Perm ε := permCongr (sumCompl p) (sumCongr ep en) #align equiv.perm.subtype_congr Equiv.Perm.subtypeCongr -theorem Perm.subtypeCongr.apply (a : ε) : ep.subtypeCongr en a = if h : p a then ep ⟨a, h⟩ else en ⟨a, h⟩ := by +theorem Perm.subtypeCongr.apply (a : ε) : + ep.subtypeCongr en a = if h : p a then ep ⟨a, h⟩ else en ⟨a, h⟩ := by by_cases h:p a <;> simp [perm.subtype_congr, h] #align equiv.perm.subtype_congr.apply Equiv.Perm.subtypeCongr.apply @@ -536,7 +556,8 @@ theorem Perm.subtypeCongr.right_apply {a : ε} (h : ¬p a) : ep.subtypeCongr en #align equiv.perm.subtype_congr.right_apply Equiv.Perm.subtypeCongr.right_apply @[simp] -theorem Perm.subtypeCongr.right_apply_subtype (a : { a // ¬p a }) : ep.subtypeCongr en a = en a := by +theorem Perm.subtypeCongr.right_apply_subtype (a : { a // ¬p a }) : + ep.subtypeCongr en a = en a := by convert perm.subtype_congr.right_apply _ _ a.property simp #align equiv.perm.subtype_congr.right_apply_subtype Equiv.Perm.subtypeCongr.right_apply_subtype @@ -562,7 +583,8 @@ theorem Perm.subtypeCongr.symm : (ep.subtypeCongr en).symm = Perm.subtypeCongr e @[simp] theorem Perm.subtypeCongr.trans : - (ep.subtypeCongr en).trans (ep'.subtypeCongr en') = Perm.subtypeCongr (ep.trans ep') (en.trans en') := by + (ep.subtypeCongr en).trans (ep'.subtypeCongr en') + = Perm.subtypeCongr (ep.trans ep') (en.trans en') := by ext x by_cases h:p x · have : p (ep ⟨x, h⟩) := Subtype.property _ @@ -615,7 +637,8 @@ section /-- A family of equivalences `Π a, β₁ a ≃ β₂ a` generates an equivalence between `Π a, β₁ a` and `Π a, β₂ a`. -/ def piCongrRight {α} {β₁ β₂ : α → Sort _} (F : ∀ a, β₁ a ≃ β₂ a) : (∀ a, β₁ a) ≃ ∀ a, β₂ a := - ⟨fun H a => F a (H a), fun H a => (F a).symm (H a), fun H => funext <| by simp, fun H => funext <| by simp⟩ + ⟨fun H a => F a (H a), fun H a => (F a).symm (H a), fun H => funext <| by simp, + fun H => funext <| by simp⟩ #align equiv.Pi_congr_right Equiv.piCongrRight /-- Given `φ : α → β → Sort*`, we have an equivalence between `Π a b, φ a b` and `Π b a, φ a b`. @@ -634,7 +657,8 @@ theorem Pi_comm_symm {α β} {φ : α → β → Sort _} : (piComm φ).symm = (P to the type of dependent functions of two arguments (i.e., functions to the space of functions). This is `sigma.curry` and `sigma.uncurry` together as an equiv. -/ -def piCurry {α} {β : α → Sort _} (γ : ∀ a, β a → Sort _) : (∀ x : Σi, β i, γ x.1 x.2) ≃ ∀ a b, γ a b where +def piCurry {α} {β : α → Sort _} (γ : ∀ a, β a → Sort _) : + (∀ x : Σi, β i, γ x.1 x.2) ≃ ∀ a b, γ a b where toFun := Sigma.curry invFun := Sigma.uncurry left_inv := Sigma.uncurry_curry @@ -665,7 +689,8 @@ theorem prod_congr_left_apply (b : β₁) (a : α₁) : prodCongrLeft e (b, a) = rfl #align equiv.prod_congr_left_apply Equiv.prod_congr_left_apply -theorem prod_congr_refl_right (e : β₁ ≃ β₂) : prodCongr e (Equiv.refl α₁) = prodCongrLeft fun _ => e := by +theorem prod_congr_refl_right (e : β₁ ≃ β₂) : + prodCongr e (Equiv.refl α₁) = prodCongrLeft fun _ => e := by ext ⟨a, b⟩ : 1 simp #align equiv.prod_congr_refl_right Equiv.prod_congr_refl_right @@ -688,7 +713,8 @@ theorem prod_congr_right_apply (a : α₁) (b : β₁) : prodCongrRight e (a, b) rfl #align equiv.prod_congr_right_apply Equiv.prod_congr_right_apply -theorem prod_congr_refl_left (e : β₁ ≃ β₂) : prodCongr (Equiv.refl α₁) e = prodCongrRight fun _ => e := by +theorem prod_congr_refl_left (e : β₁ ≃ β₂) : + prodCongr (Equiv.refl α₁) e = prodCongrRight fun _ => e := by ext ⟨a, b⟩ : 1 simp #align equiv.prod_congr_refl_left Equiv.prod_congr_refl_left @@ -708,13 +734,15 @@ theorem prod_congr_right_trans_prod_comm : #align equiv.prod_congr_right_trans_prod_comm Equiv.prod_congr_right_trans_prod_comm theorem sigma_congr_right_sigma_equiv_prod : - (sigmaCongrRight e).trans (sigmaEquivProd α₁ β₂) = (sigmaEquivProd α₁ β₁).trans (prodCongrRight e) := by + (sigmaCongrRight e).trans (sigmaEquivProd α₁ β₂) + = (sigmaEquivProd α₁ β₁).trans (prodCongrRight e) := by ext ⟨a, b⟩ : 1 simp #align equiv.sigma_congr_right_sigma_equiv_prod Equiv.sigma_congr_right_sigma_equiv_prod theorem sigma_equiv_prod_sigma_congr_right : - (sigmaEquivProd α₁ β₁).symm.trans (sigmaCongrRight e) = (prodCongrRight e).trans (sigmaEquivProd α₁ β₂).symm := by + (sigmaEquivProd α₁ β₁).symm.trans (sigmaCongrRight e) + = (prodCongrRight e).trans (sigmaEquivProd α₁ β₂).symm := by ext ⟨a, b⟩ : 1 simp #align equiv.sigma_equiv_prod_sigma_congr_right Equiv.sigma_equiv_prod_sigma_congr_right @@ -722,12 +750,13 @@ theorem sigma_equiv_prod_sigma_congr_right : -- See also `equiv.of_preimage_equiv`. /-- A family of equivalences between fibers gives an equivalence between domains. -/ @[simps] -def ofFiberEquiv {α β γ : Type _} {f : α → γ} {g : β → γ} (e : ∀ c, { a // f a = c } ≃ { b // g b = c }) : α ≃ β := +def ofFiberEquiv {α β γ : Type _} {f : α → γ} {g : β → γ} + (e : ∀ c, { a // f a = c } ≃ { b // g b = c }) : α ≃ β := (sigmaFiberEquiv f).symm.trans <| (Equiv.sigmaCongrRight e).trans (sigmaFiberEquiv g) #align equiv.of_fiber_equiv Equiv.ofFiberEquiv -theorem of_fiber_equiv_map {α β γ} {f : α → γ} {g : β → γ} (e : ∀ c, { a // f a = c } ≃ { b // g b = c }) (a : α) : - g (ofFiberEquiv e a) = f a := +theorem of_fiber_equiv_map {α β γ} {f : α → γ} {g : β → γ} + (e : ∀ c, { a // f a = c } ≃ { b // g b = c }) (a : α) : g (ofFiberEquiv e a) = f a := (_ : { b // g b = _ }).Prop #align equiv.of_fiber_equiv_map Equiv.of_fiber_equiv_map @@ -772,12 +801,13 @@ theorem prod_extend_right_apply_eq (b : β₁) : prodExtendRight a e (a, b) = (a if_pos rfl #align equiv.perm.prod_extend_right_apply_eq Equiv.Perm.prod_extend_right_apply_eq -theorem prod_extend_right_apply_ne {a a' : α₁} (h : a' ≠ a) (b : β₁) : prodExtendRight a e (a', b) = (a', b) := +theorem prod_extend_right_apply_ne {a a' : α₁} (h : a' ≠ a) (b : β₁) : + prodExtendRight a e (a', b) = (a', b) := if_neg h #align equiv.perm.prod_extend_right_apply_ne Equiv.Perm.prod_extend_right_apply_ne -theorem eq_of_prod_extend_right_ne {e : Perm β₁} {a a' : α₁} {b : β₁} (h : prodExtendRight a e (a', b) ≠ (a', b)) : - a' = a := by +theorem eq_of_prod_extend_right_ne {e : Perm β₁} {a a' : α₁} {b : β₁} + (h : prodExtendRight a e (a', b) ≠ (a', b)) : a' = a := by contrapose! h exact prod_extend_right_apply_ne _ h _ #align equiv.perm.eq_of_prod_extend_right_ne Equiv.Perm.eq_of_prod_extend_right_ne @@ -799,10 +829,11 @@ section /-- The type of functions to a product `α × β` is equivalent to the type of pairs of functions `γ → α` and `γ → β`. -/ def arrowProdEquivProdArrow (α β γ : Type _) : (γ → α × β) ≃ (γ → α) × (γ → β) := - ⟨fun f => (fun c => (f c).1, fun c => (f c).2), fun p c => (p.1 c, p.2 c), fun f => funext fun c => Prod.mk.eta, - fun p => by - cases p - rfl⟩ + ⟨fun f => (fun c => (f c).1, fun c => (f c).2), fun p c => (p.1 c, p.2 c), + fun f => funext fun c => Prod.mk.eta, + fun p => by + cases p + rfl⟩ #align equiv.arrow_prod_equiv_prod_arrow Equiv.arrowProdEquivProdArrow open Sum @@ -841,27 +872,32 @@ theorem sum_arrow_equiv_prod_arrow_symm_apply_inr {α β γ} (f : α → γ) (g /-- Type product is right distributive with respect to type sum up to an equivalence. -/ def sumProdDistrib (α β γ : Sort _) : Sum α β × γ ≃ Sum (α × γ) (β × γ) := - ⟨fun p => p.1.map (fun x => (x, p.2)) fun x => (x, p.2), fun s => s.elim (Prod.map inl id) (Prod.map inr id), by - rintro ⟨_ | _, _⟩ <;> rfl, by rintro (⟨_, _⟩ | ⟨_, _⟩) <;> rfl⟩ + ⟨fun p => p.1.map (fun x => (x, p.2)) fun x => (x, p.2), + fun s => s.elim (Prod.map inl id) (Prod.map inr id), by + rintro ⟨_ | _, _⟩ <;> rfl, by rintro (⟨_, _⟩ | ⟨_, _⟩) <;> rfl⟩ #align equiv.sum_prod_distrib Equiv.sumProdDistrib @[simp] -theorem sum_prod_distrib_apply_left {α β γ} (a : α) (c : γ) : sumProdDistrib α β γ (Sum.inl a, c) = Sum.inl (a, c) := +theorem sum_prod_distrib_apply_left {α β γ} (a : α) (c : γ) : + sumProdDistrib α β γ (Sum.inl a, c) = Sum.inl (a, c) := rfl #align equiv.sum_prod_distrib_apply_left Equiv.sum_prod_distrib_apply_left @[simp] -theorem sum_prod_distrib_apply_right {α β γ} (b : β) (c : γ) : sumProdDistrib α β γ (Sum.inr b, c) = Sum.inr (b, c) := +theorem sum_prod_distrib_apply_right {α β γ} (b : β) (c : γ) : + sumProdDistrib α β γ (Sum.inr b, c) = Sum.inr (b, c) := rfl #align equiv.sum_prod_distrib_apply_right Equiv.sum_prod_distrib_apply_right @[simp] -theorem sum_prod_distrib_symm_apply_left {α β γ} (a : α × γ) : (sumProdDistrib α β γ).symm (inl a) = (inl a.1, a.2) := +theorem sum_prod_distrib_symm_apply_left {α β γ} (a : α × γ) : + (sumProdDistrib α β γ).symm (inl a) = (inl a.1, a.2) := rfl #align equiv.sum_prod_distrib_symm_apply_left Equiv.sum_prod_distrib_symm_apply_left @[simp] -theorem sum_prod_distrib_symm_apply_right {α β γ} (b : β × γ) : (sumProdDistrib α β γ).symm (inr b) = (inr b.1, b.2) := +theorem sum_prod_distrib_symm_apply_right {α β γ} (b : β × γ) : + (sumProdDistrib α β γ).symm (inr b) = (inr b.1, b.2) := rfl #align equiv.sum_prod_distrib_symm_apply_right Equiv.sum_prod_distrib_symm_apply_right @@ -875,28 +911,33 @@ def prodSumDistrib (α β γ : Sort _) : α × Sum β γ ≃ Sum (α × β) (α #align equiv.prod_sum_distrib Equiv.prodSumDistrib @[simp] -theorem prod_sum_distrib_apply_left {α β γ} (a : α) (b : β) : prodSumDistrib α β γ (a, Sum.inl b) = Sum.inl (a, b) := +theorem prod_sum_distrib_apply_left {α β γ} (a : α) (b : β) : + prodSumDistrib α β γ (a, Sum.inl b) = Sum.inl (a, b) := rfl #align equiv.prod_sum_distrib_apply_left Equiv.prod_sum_distrib_apply_left @[simp] -theorem prod_sum_distrib_apply_right {α β γ} (a : α) (c : γ) : prodSumDistrib α β γ (a, Sum.inr c) = Sum.inr (a, c) := +theorem prod_sum_distrib_apply_right {α β γ} (a : α) (c : γ) : + prodSumDistrib α β γ (a, Sum.inr c) = Sum.inr (a, c) := rfl #align equiv.prod_sum_distrib_apply_right Equiv.prod_sum_distrib_apply_right @[simp] -theorem prod_sum_distrib_symm_apply_left {α β γ} (a : α × β) : (prodSumDistrib α β γ).symm (inl a) = (a.1, inl a.2) := +theorem prod_sum_distrib_symm_apply_left {α β γ} (a : α × β) : + (prodSumDistrib α β γ).symm (inl a) = (a.1, inl a.2) := rfl #align equiv.prod_sum_distrib_symm_apply_left Equiv.prod_sum_distrib_symm_apply_left @[simp] -theorem prod_sum_distrib_symm_apply_right {α β γ} (a : α × γ) : (prodSumDistrib α β γ).symm (inr a) = (a.1, inr a.2) := +theorem prod_sum_distrib_symm_apply_right {α β γ} (a : α × γ) : + (prodSumDistrib α β γ).symm (inr a) = (a.1, inr a.2) := rfl #align equiv.prod_sum_distrib_symm_apply_right Equiv.prod_sum_distrib_symm_apply_right /-- An indexed sum of disjoint sums of types is equivalent to the sum of the indexed sums. -/ @[simps] -def sigmaSumDistrib {ι : Type _} (α β : ι → Type _) : (Σi, Sum (α i) (β i)) ≃ Sum (Σi, α i) (Σi, β i) := +def sigmaSumDistrib {ι : Type _} (α β : ι → Type _) : + (Σi, Sum (α i) (β i)) ≃ Sum (Σi, α i) (Σi, β i) := ⟨fun p => p.2.map (Sigma.mk p.1) (Sigma.mk p.1), Sum.elim (Sigma.map id fun _ => Sum.inl) (Sigma.map id fun _ => Sum.inr), fun p => by rcases p with ⟨i, a | b⟩ <;> rfl, fun p => by rcases p with (⟨i, a⟩ | ⟨i, b⟩) <;> rfl⟩ @@ -1028,12 +1069,14 @@ theorem subtype_equiv_symm {p : α → Prop} {q : β → Prop} (e : α ≃ β) ( @[simp] theorem subtype_equiv_trans {p : α → Prop} {q : β → Prop} {r : γ → Prop} (e : α ≃ β) (f : β ≃ γ) (h : ∀ a : α, p a ↔ q (e a)) (h' : ∀ b : β, q b ↔ r (f b)) : - (e.subtypeEquiv h).trans (f.subtypeEquiv h') = (e.trans f).subtypeEquiv fun a => (h a).trans (h' <| e a) := + (e.subtypeEquiv h).trans (f.subtypeEquiv h') + = (e.trans f).subtypeEquiv fun a => (h a).trans (h' <| e a) := rfl #align equiv.subtype_equiv_trans Equiv.subtype_equiv_trans @[simp] -theorem subtype_equiv_apply {p : α → Prop} {q : β → Prop} (e : α ≃ β) (h : ∀ a : α, p a ↔ q (e a)) (x : { x // p x }) : +theorem subtype_equiv_apply {p : α → Prop} {q : β → Prop} + (e : α ≃ β) (h : ∀ a : α, p a ↔ q (e a)) (x : { x // p x }) : e.subtypeEquiv h x = ⟨e x, (h _).1 x.2⟩ := rfl #align equiv.subtype_equiv_apply Equiv.subtype_equiv_apply @@ -1053,7 +1096,8 @@ def subtypeEquivOfSubtype {p : β → Prop} (e : α ≃ β) : { a : α // p (e a /-- If `α ≃ β`, then for any predicate `p : α → Prop` the subtype `{a // p a}` is equivalent to the subtype `{b // p (e.symm b)}`. This version is used by `equiv_rw`. -/ -def subtypeEquivOfSubtype' {p : α → Prop} (e : α ≃ β) : { a : α // p a } ≃ { b : β // p (e.symm b) } := +def subtypeEquivOfSubtype' {p : α → Prop} (e : α ≃ β) : + { a : α // p a } ≃ { b : β // p (e.symm b) } := e.symm.subtypeEquivOfSubtype.symm #align equiv.subtype_equiv_of_subtype' Equiv.subtypeEquivOfSubtype' @@ -1097,8 +1141,10 @@ def subtypeUnivEquiv {α : Type u} {p : α → Prop} (h : ∀ x, p x) : Subtype #align equiv.subtype_univ_equiv Equiv.subtypeUnivEquiv /-- A subtype of a sigma-type is a sigma-type over a subtype. -/ -def subtypeSigmaEquiv {α : Type u} (p : α → Type v) (q : α → Prop) : { y : Sigma p // q y.1 } ≃ Σx : Subtype q, p x.1 := - ⟨fun x => ⟨⟨x.1.1, x.2⟩, x.1.2⟩, fun x => ⟨⟨x.1.1, x.2⟩, x.1.2⟩, fun ⟨⟨x, h⟩, y⟩ => rfl, fun ⟨⟨x, y⟩, h⟩ => rfl⟩ +def subtypeSigmaEquiv {α : Type u} (p : α → Type v) (q : α → Prop) : { y : Sigma p // q y.1 } ≃ Σx : + Subtype q, p x.1 := + ⟨fun x => ⟨⟨x.1.1, x.2⟩, x.1.2⟩, fun x => ⟨⟨x.1.1, x.2⟩, x.1.2⟩, fun ⟨⟨x, h⟩, y⟩ => rfl, + fun ⟨⟨x, y⟩, h⟩ => rfl⟩ #align equiv.subtype_sigma_equiv Equiv.subtypeSigmaEquiv /-- A sigma type over a subtype is equivalent to the sigma set over the original type, @@ -1120,17 +1166,19 @@ def sigmaSubtypeFiberEquiv {α : Type u} {β : Type v} (f : α → β) (p : β /-- If for each `x` we have `p x ↔ q (f x)`, then `Σ y : {y // q y}, f ⁻¹' {y}` is equivalent to `{x // p x}`. -/ -def sigmaSubtypeFiberEquivSubtype {α : Type u} {β : Type v} (f : α → β) {p : α → Prop} {q : β → Prop} - (h : ∀ x, p x ↔ q (f x)) : (Σy : Subtype q, { x : α // f x = y }) ≃ Subtype p := +def sigmaSubtypeFiberEquivSubtype {α : Type u} {β : Type v} (f : α → β) {p : α → Prop} + {q : β → Prop} (h : ∀ x, p x ↔ q (f x)) : (Σy : Subtype q, { x : α // f x = y }) ≃ Subtype p := calc - (Σy : Subtype q, { x : α // f x = y }) ≃ Σy : Subtype q, { x : Subtype p // Subtype.mk (f x) ((h x).1 x.2) = y } := + (Σy : Subtype q, { x : α // f x = y }) ≃ Σy : + Subtype q, { x : Subtype p // Subtype.mk (f x) ((h x).1 x.2) = y } := by apply sigma_congr_right intro y symm refine' (subtype_subtype_equiv_subtype_exists _ _).trans (subtype_equiv_right _) intro x - exact ⟨fun ⟨hp, h'⟩ => congr_arg Subtype.val h', fun h' => ⟨(h x).2 (h'.symm ▸ y.2), Subtype.eq h'⟩⟩ + exact ⟨fun ⟨hp, h'⟩ => congr_arg Subtype.val h', fun h' => ⟨(h x).2 (h'.symm ▸ y.2), + Subtype.eq h'⟩⟩ _ ≃ Subtype p := sigmaFiberEquiv fun x : Subtype p => (⟨f x, (h x).1 x.property⟩ : Subtype q) #align equiv.sigma_subtype_fiber_equiv_subtype Equiv.sigmaSubtypeFiberEquivSubtype @@ -1154,12 +1202,14 @@ def sigmaOptionEquivOfSome {α : Type u} (p : Option α → Type v) (h : p none /-- The `pi`-type `Π i, π i` is equivalent to the type of sections `f : ι → Σ i, π i` of the `sigma` type such that for all `i` we have `(f i).fst = i`. -/ -def piEquivSubtypeSigma (ι : Type _) (π : ι → Type _) : (∀ i, π i) ≃ { f : ι → Σi, π i // ∀ i, (f i).1 = i } := +def piEquivSubtypeSigma (ι : Type _) (π : ι → Type _) : + (∀ i, π i) ≃ { f : ι → Σi, π i // ∀ i, (f i).1 = i } := ⟨fun f => ⟨fun i => ⟨i, f i⟩, fun i => rfl⟩, fun f i => by rw [← f.2 i] exact (f.1 i).2, fun f => funext fun i => rfl, fun ⟨f, hf⟩ => Subtype.eq <| - funext fun i => Sigma.eq (hf i).symm <| eq_of_heq <| rec_heq_of_heq _ <| rec_heq_of_heq _ <| HEq.refl _⟩ + funext fun i => + Sigma.eq (hf i).symm <| eq_of_heq <| rec_heq_of_heq _ <| rec_heq_of_heq _ <| HEq.refl _⟩ #align equiv.pi_equiv_subtype_sigma Equiv.piEquivSubtypeSigma /-- The set of functions `f : Π a, β a` such that for all `a` we have `p a (f a)` is equivalent @@ -1178,8 +1228,8 @@ def subtypePiEquivPi {α : Sort u} {β : α → Sort v} {p : ∀ a, β a → Pro is equivalent to a product of subtypes. -/ def subtypeProdEquivProd {α : Type u} {β : Type v} {p : α → Prop} {q : β → Prop} : { c : α × β // p c.1 ∧ q c.2 } ≃ { a // p a } × { b // q b } := - ⟨fun x => ⟨⟨x.1.1, x.2.1⟩, ⟨x.1.2, x.2.2⟩⟩, fun x => ⟨⟨x.1.1, x.2.1⟩, ⟨x.1.2, x.2.2⟩⟩, fun ⟨⟨_, _⟩, ⟨_, _⟩⟩ => rfl, - fun ⟨⟨_, _⟩, ⟨_, _⟩⟩ => rfl⟩ + ⟨fun x => ⟨⟨x.1.1, x.2.1⟩, ⟨x.1.2, x.2.2⟩⟩, fun x => ⟨⟨x.1.1, x.2.1⟩, ⟨x.1.2, x.2.2⟩⟩, + fun ⟨⟨_, _⟩, ⟨_, _⟩⟩ => rfl, fun ⟨⟨_, _⟩, ⟨_, _⟩⟩ => rfl⟩ #align equiv.subtype_prod_equiv_prod Equiv.subtypeProdEquivProd /-- A subtype of a `prod` is equivalent to a sigma type whose fibers are subtypes. -/ @@ -1217,7 +1267,8 @@ def piEquivPiSubtypeProd {α : Type _} (p : α → Prop) (β : α → Type _) [D /-- A product of types can be split as the binary product of one of the types and the product of all the remaining types. -/ @[simps] -def piSplitAt {α : Type _} [DecidableEq α] (i : α) (β : α → Type _) : (∀ j, β j) ≃ β i × ∀ j : { j // j ≠ i }, β j where +def piSplitAt {α : Type _} [DecidableEq α] (i : α) (β : α → Type _) : + (∀ j, β j) ≃ β i × ∀ j : { j // j ≠ i }, β j where toFun f := ⟨f i, fun j => f j⟩ invFun f j := if h : j = i then h.symm.rec f.1 else f.2 ⟨j, h⟩ right_inv f := by @@ -1236,7 +1287,8 @@ def piSplitAt {α : Type _} [DecidableEq α] (i : α) (β : α → Type _) : ( /-- A product of copies of a type can be split as the binary product of one copy and the product of all the remaining copies. -/ @[simps] -def funSplitAt {α : Type _} [DecidableEq α] (i : α) (β : Type _) : (α → β) ≃ β × ({ j // j ≠ i } → β) := +def funSplitAt {α : Type _} [DecidableEq α] (i : α) (β : Type _) : + (α → β) ≃ β × ({ j // j ≠ i } → β) := piSplitAt i _ #align equiv.fun_split_at Equiv.funSplitAt @@ -1253,7 +1305,8 @@ def subtypeEquivCodomain (f : { x' // x' ≠ x } → Y) : { g : X → Y // g ∘ @funUnique { x' // ¬x' ≠ x } _ <| show Unique { x' // ¬x' ≠ x } from @Equiv.unique _ _ - (show Unique { x' // x' = x } from { default := ⟨x, rfl⟩, uniq := fun ⟨x', h⟩ => Subtype.val_injective h }) + (show Unique { x' // x' = x } from { + default := ⟨x, rfl⟩, uniq := fun ⟨x', h⟩ => Subtype.val_injective h }) (subtype_equiv_right fun a => not_not) #align equiv.subtype_equiv_codomain Equiv.subtypeEquivCodomain @@ -1290,7 +1343,8 @@ theorem subtype_equiv_codomain_symm_apply_eq (f : { x' // x' ≠ x } → Y) (y : dif_neg (not_not.mpr rfl) #align equiv.subtype_equiv_codomain_symm_apply_eq Equiv.subtype_equiv_codomain_symm_apply_eq -theorem subtype_equiv_codomain_symm_apply_ne (f : { x' // x' ≠ x } → Y) (y : Y) (x' : X) (h : x' ≠ x) : +theorem subtype_equiv_codomain_symm_apply_ne + (f : { x' // x' ≠ x } → Y) (y : Y) (x' : X) (h : x' ≠ x) : ((subtypeEquivCodomain f).symm y : X → Y) x' = f ⟨x', h⟩ := dif_pos h #align equiv.subtype_equiv_codomain_symm_apply_ne Equiv.subtype_equiv_codomain_symm_apply_ne @@ -1306,12 +1360,14 @@ noncomputable def ofBijective (f : α → β) (hf : Bijective f) : α ≃ β whe right_inv := Function.right_inverse_surj_inv _ #align equiv.of_bijective Equiv.ofBijective -theorem of_bijective_apply_symm_apply (f : α → β) (hf : Bijective f) (x : β) : f ((ofBijective f hf).symm x) = x := +theorem of_bijective_apply_symm_apply (f : α → β) (hf : Bijective f) (x : β) : + f ((ofBijective f hf).symm x) = x := (ofBijective f hf).apply_symm_apply x #align equiv.of_bijective_apply_symm_apply Equiv.of_bijective_apply_symm_apply @[simp] -theorem of_bijective_symm_apply_apply (f : α → β) (hf : Bijective f) (x : α) : (ofBijective f hf).symm (f x) = x := +theorem of_bijective_symm_apply_apply (f : α → β) (hf : Bijective f) (x : α) : + (ofBijective f hf).symm (f x) = x := (ofBijective f hf).symm_apply_apply x #align equiv.of_bijective_symm_apply_apply Equiv.of_bijective_symm_apply_apply @@ -1333,10 +1389,12 @@ def Perm.extendDomain : Perm β' := #align equiv.perm.extend_domain Equiv.Perm.extendDomain @[simp] -theorem Perm.extend_domain_apply_image (a : α') : e.extendDomain f (f a) = f (e a) := by simp [perm.extend_domain] +theorem Perm.extend_domain_apply_image (a : α') : e.extendDomain f (f a) = f (e a) := by + simp [perm.extend_domain] #align equiv.perm.extend_domain_apply_image Equiv.Perm.extend_domain_apply_image -theorem Perm.extend_domain_apply_subtype {b : β'} (h : p b) : e.extendDomain f b = f (e (f.symm ⟨b, h⟩)) := by +theorem Perm.extend_domain_apply_subtype {b : β'} (h : p b) : + e.extendDomain f b = f (e (f.symm ⟨b, h⟩)) := by simp [perm.extend_domain, h] #align equiv.perm.extend_domain_apply_subtype Equiv.Perm.extend_domain_apply_subtype @@ -1345,7 +1403,8 @@ theorem Perm.extend_domain_apply_not_subtype {b : β'} (h : ¬p b) : e.extendDom #align equiv.perm.extend_domain_apply_not_subtype Equiv.Perm.extend_domain_apply_not_subtype @[simp] -theorem Perm.extend_domain_refl : Perm.extendDomain (Equiv.refl _) f = Equiv.refl _ := by simp [perm.extend_domain] +theorem Perm.extend_domain_refl : Perm.extendDomain (Equiv.refl _) f = Equiv.refl _ := by + simp [perm.extend_domain] #align equiv.perm.extend_domain_refl Equiv.Perm.extend_domain_refl @[simp] @@ -1365,11 +1424,13 @@ equivalence relation `~`. Let `p₂` be a predicate on the quotient type `α/~`, of this predicate to `α`: `p₁ a ↔ p₂ ⟦a⟧`. Let `~₂` be the restriction of `~` to `{x // p₁ x}`. Then `{x // p₂ x}` is equivalent to the quotient of `{x // p₁ x}` by `~₂`. -/ def subtypeQuotientEquivQuotientSubtype (p₁ : α → Prop) [s₁ : Setoid α] [s₂ : Setoid (Subtype p₁)] - (p₂ : Quotient s₁ → Prop) (hp₂ : ∀ a, p₁ a ↔ p₂ ⟦a⟧) (h : ∀ x y : Subtype p₁, @Setoid.r _ s₂ x y ↔ (x : α) ≈ y) : + (p₂ : Quotient s₁ → Prop) (hp₂ : ∀ a, p₁ a ↔ p₂ ⟦a⟧) + (h : ∀ x y : Subtype p₁, @Setoid.r _ s₂ x y ↔ (x : α) ≈ y) : { x // p₂ x } ≃ Quotient s₂ where toFun a := Quotient.hrecOn a.1 (fun a h => ⟦⟨a, (hp₂ _).2 h⟩⟧) - (fun a b hab => hfunext (by rw [Quotient.sound hab]) fun h₁ h₂ _ => heq_of_eq (Quotient.sound ((h _ _).2 hab))) + (fun a b hab => hfunext (by rw [Quotient.sound hab]) fun h₁ h₂ _ => + heq_of_eq (Quotient.sound ((h _ _).2 hab))) a.2 invFun a := Quotient.liftOn a (fun a => (⟨⟦a.1⟧, (hp₂ _).1 a.2⟩ : { x // p₂ x })) fun a b hab => @@ -1379,15 +1440,17 @@ def subtypeQuotientEquivQuotientSubtype (p₁ : α → Prop) [s₁ : Setoid α] #align equiv.subtype_quotient_equiv_quotient_subtype Equiv.subtypeQuotientEquivQuotientSubtype @[simp] -theorem subtype_quotient_equiv_quotient_subtype_mk (p₁ : α → Prop) [s₁ : Setoid α] [s₂ : Setoid (Subtype p₁)] - (p₂ : Quotient s₁ → Prop) (hp₂ : ∀ a, p₁ a ↔ p₂ ⟦a⟧) (h : ∀ x y : Subtype p₁, @Setoid.r _ s₂ x y ↔ (x : α) ≈ y) +theorem subtype_quotient_equiv_quotient_subtype_mk (p₁ : α → Prop) + [s₁ : Setoid α] [s₂ : Setoid (Subtype p₁)] (p₂ : Quotient s₁ → Prop) (hp₂ : ∀ a, p₁ a ↔ p₂ ⟦a⟧) + (h : ∀ x y : Subtype p₁, @Setoid.r _ s₂ x y ↔ (x : α) ≈ y) (x hx) : subtypeQuotientEquivQuotientSubtype p₁ p₂ hp₂ h ⟨⟦x⟧, hx⟩ = ⟦⟨x, (hp₂ _).2 hx⟩⟧ := rfl #align equiv.subtype_quotient_equiv_quotient_subtype_mk Equiv.subtype_quotient_equiv_quotient_subtype_mk @[simp] -theorem subtype_quotient_equiv_quotient_subtype_symm_mk (p₁ : α → Prop) [s₁ : Setoid α] [s₂ : Setoid (Subtype p₁)] - (p₂ : Quotient s₁ → Prop) (hp₂ : ∀ a, p₁ a ↔ p₂ ⟦a⟧) (h : ∀ x y : Subtype p₁, @Setoid.r _ s₂ x y ↔ (x : α) ≈ y) +theorem subtype_quotient_equiv_quotient_subtype_symm_mk (p₁ : α → Prop) + [s₁ : Setoid α] [s₂ : Setoid (Subtype p₁)] (p₂ : Quotient s₁ → Prop) (hp₂ : ∀ a, p₁ a ↔ p₂ ⟦a⟧) + (h : ∀ x y : Subtype p₁, @Setoid.r _ s₂ x y ↔ (x : α) ≈ y) (x) : (subtypeQuotientEquivQuotientSubtype p₁ p₂ hp₂ h).symm ⟦x⟧ = ⟨⟦x⟧, (hp₂ _).1 x.Prop⟩ := rfl #align equiv.subtype_quotient_equiv_quotient_subtype_symm_mk Equiv.subtype_quotient_equiv_quotient_subtype_symm_mk @@ -1419,7 +1482,8 @@ theorem swap_core_comm (r a b : α) : swapCore a b r = swapCore b a r := by /-- `swap a b` is the permutation that swaps `a` and `b` and leaves other values as is. -/ def swap (a b : α) : Perm α := - ⟨swapCore a b, swapCore a b, fun r => swap_core_swap_core r a b, fun r => swap_core_swap_core r a b⟩ + ⟨swapCore a b, swapCore a b, fun r => swap_core_swap_core r a b, + fun r => swap_core_swap_core r a b⟩ #align equiv.swap Equiv.swap @[simp] @@ -1441,7 +1505,8 @@ theorem swap_apply_left (a b : α) : swap a b a = b := #align equiv.swap_apply_left Equiv.swap_apply_left @[simp] -theorem swap_apply_right (a b : α) : swap a b b = a := by by_cases h:b = a <;> simp [swap_apply_def, h] +theorem swap_apply_right (a b : α) : swap a b b = a := by + by_cases h:b = a <;> simp [swap_apply_def, h] #align equiv.swap_apply_right Equiv.swap_apply_right theorem swap_apply_of_ne_of_ne {a b x : α} : x ≠ a → x ≠ b → swap a b x = x := by @@ -1474,7 +1539,8 @@ theorem swap_eq_update (i j : α) : (Equiv.swap i j : α → α) = update (updat funext fun x => by rw [update_apply _ i j, update_apply _ j i, Equiv.swap_apply_def, id.def] #align equiv.swap_eq_update Equiv.swap_eq_update -theorem comp_swap_eq_update (i j : α) (f : α → β) : f ∘ Equiv.swap i j = update (update f j (f i)) i (f j) := by +theorem comp_swap_eq_update (i j : α) (f : α → β) : + f ∘ Equiv.swap i j = update (update f j (f i)) i (f j) := by rw [swap_eq_update, comp_update, comp_update, comp.right_id] #align equiv.comp_swap_eq_update Equiv.comp_swap_eq_update @@ -1501,7 +1567,8 @@ theorem swap_apply_self (i j a : α) : swap i j (swap i j a) = a := by #align equiv.swap_apply_self Equiv.swap_apply_self /-- A function is invariant to a swap if it is equal at both elements -/ -theorem apply_swap_eq_self {v : α → β} {i j : α} (hv : v i = v j) (k : α) : v (swap i j k) = v k := by +theorem apply_swap_eq_self {v : α → β} {i j : α} (hv : v i = v j) (k : α) : + v (swap i j k) = v k := by by_cases hi:k = i · rw [hi, swap_apply_left, hv] @@ -1599,7 +1666,8 @@ theorem PLift.eq_up_iff_down_eq {x : PLift α} {y : α} : x = PLift.up y ↔ x.d #align plift.eq_up_iff_down_eq PLift.eq_up_iff_down_eq theorem Function.Injective.map_swap {α β : Type _} [DecidableEq α] [DecidableEq β] {f : α → β} - (hf : Function.Injective f) (x y z : α) : f (Equiv.swap x y z) = Equiv.swap (f x) (f y) (f z) := by + (hf : Function.Injective f) (x y z : α) : + f (Equiv.swap x y z) = Equiv.swap (f x) (f y) (f z) := by conv_rhs => rw [Equiv.swap_apply_def] split_ifs with h₁ h₂ · rw [hf h₁, Equiv.swap_apply_left] @@ -1662,11 +1730,13 @@ def piCongr : (∀ a, W a) ≃ ∀ b, Z b := #align equiv.Pi_congr Equiv.piCongr @[simp] -theorem coe_Pi_congr_symm : ((h₁.piCongr h₂).symm : (∀ b, Z b) → ∀ a, W a) = fun f a => (h₂ a).symm (f (h₁ a)) := +theorem coe_Pi_congr_symm : ((h₁.piCongr h₂).symm : + (∀ b, Z b) → ∀ a, W a) = fun f a => (h₂ a).symm (f (h₁ a)) := rfl #align equiv.coe_Pi_congr_symm Equiv.coe_Pi_congr_symm -theorem Pi_congr_symm_apply (f : ∀ b, Z b) : (h₁.piCongr h₂).symm f = fun a => (h₂ a).symm (f (h₁ a)) := +theorem Pi_congr_symm_apply (f : ∀ b, Z b) : + (h₁.piCongr h₂).symm f = fun a => (h₂ a).symm (f (h₁ a)) := rfl #align equiv.Pi_congr_symm_apply Equiv.Pi_congr_symm_apply @@ -1694,7 +1764,8 @@ def piCongr' : (∀ a, W a) ≃ ∀ b, Z b := #align equiv.Pi_congr' Equiv.piCongr' @[simp] -theorem coe_Pi_congr' : (h₁.piCongr' h₂ : (∀ a, W a) → ∀ b, Z b) = fun f b => h₂ b <| f <| h₁.symm b := +theorem coe_Pi_congr' : + (h₁.piCongr' h₂ : (∀ a, W a) → ∀ b, Z b) = fun f b => h₂ b <| f <| h₁.symm b := rfl #align equiv.coe_Pi_congr' Equiv.coe_Pi_congr' @@ -1742,8 +1813,9 @@ end BinaryOp end Equiv -theorem Function.Injective.swap_apply [DecidableEq α] [DecidableEq β] {f : α → β} (hf : Function.Injective f) - (x y z : α) : Equiv.swap (f x) (f y) (f z) = f (Equiv.swap x y z) := by +theorem Function.Injective.swap_apply + [DecidableEq α] [DecidableEq β] {f : α → β} (hf : Function.Injective f) (x y z : α) : + Equiv.swap (f x) (f y) (f z) = f (Equiv.swap x y z) := by by_cases hx:z = x · simp [hx] @@ -1753,7 +1825,8 @@ theorem Function.Injective.swap_apply [DecidableEq α] [DecidableEq β] {f : α rw [Equiv.swap_apply_of_ne_of_ne hx hy, Equiv.swap_apply_of_ne_of_ne (hf.ne hx) (hf.ne hy)] #align function.injective.swap_apply Function.Injective.swap_apply -theorem Function.Injective.swap_comp [DecidableEq α] [DecidableEq β] {f : α → β} (hf : Function.Injective f) (x y : α) : +theorem Function.Injective.swap_comp + [DecidableEq α] [DecidableEq β] {f : α → β} (hf : Function.Injective f) (x y : α) : Equiv.swap (f x) (f y) ∘ f = f ∘ Equiv.swap x y := funext fun z => hf.swap_apply _ _ _ #align function.injective.swap_comp Function.Injective.swap_comp @@ -1768,7 +1841,8 @@ def subsingletonProdSelfEquiv {α : Type _} [Subsingleton α] : α × α ≃ α /-- To give an equivalence between two subsingleton types, it is sufficient to give any two functions between them. -/ -def equivOfSubsingletonOfSubsingleton [Subsingleton α] [Subsingleton β] (f : α → β) (g : β → α) : α ≃ β where +def equivOfSubsingletonOfSubsingleton [Subsingleton α] [Subsingleton β] (f : α → β) (g : β → α) : + α ≃ β where toFun := f invFun := g left_inv _ := Subsingleton.elim _ _ @@ -1776,7 +1850,8 @@ def equivOfSubsingletonOfSubsingleton [Subsingleton α] [Subsingleton β] (f : #align equiv_of_subsingleton_of_subsingleton equivOfSubsingletonOfSubsingleton /-- A nonempty subsingleton type is (noncomputably) equivalent to `punit`. -/ -noncomputable def Equiv.punitOfNonemptyOfSubsingleton {α : Sort _} [h : Nonempty α] [Subsingleton α] : α ≃ PUnit.{v} := +noncomputable def Equiv.punitOfNonemptyOfSubsingleton + {α : Sort _} [h : Nonempty α] [Subsingleton α] : α ≃ PUnit.{v} := equivOfSubsingletonOfSubsingleton (fun _ => PUnit.unit) fun _ => h.some #align equiv.punit_of_nonempty_of_subsingleton Equiv.punitOfNonemptyOfSubsingleton @@ -1788,18 +1863,20 @@ def uniqueUniqueEquiv : Unique (Unique α) ≃ Unique α := namespace Function -theorem update_comp_equiv {α β α' : Sort _} [DecidableEq α'] [DecidableEq α] (f : α → β) (g : α' ≃ α) (a : α) (v : β) : +theorem update_comp_equiv {α β α' : Sort _} [DecidableEq α'] [DecidableEq α] (f : α → β) + (g : α' ≃ α) (a : α) (v : β) : update f a v ∘ g = update (f ∘ g) (g.symm a) v := by rw [← update_comp_eq_of_injective _ g.injective, g.apply_symm_apply] #align function.update_comp_equiv Function.update_comp_equiv -theorem update_apply_equiv_apply {α β α' : Sort _} [DecidableEq α'] [DecidableEq α] (f : α → β) (g : α' ≃ α) (a : α) - (v : β) (a' : α') : update f a v (g a') = update (f ∘ g) (g.symm a) v a' := +theorem update_apply_equiv_apply {α β α' : Sort _} [DecidableEq α'] [DecidableEq α] (f : α → β) + (g : α' ≃ α) (a : α) (v : β) (a' : α') : update f a v (g a') = update (f ∘ g) (g.symm a) v a' := congr_fun (update_comp_equiv f g a v) a' #align function.update_apply_equiv_apply Function.update_apply_equiv_apply -theorem Pi_congr_left'_update [DecidableEq α] [DecidableEq β] (P : α → Sort _) (e : α ≃ β) (f : ∀ a, P a) (b : β) - (x : P (e.symm b)) : e.piCongrLeft' P (update f (e.symm b) x) = update (e.piCongrLeft' P f) b x := by +theorem Pi_congr_left'_update [DecidableEq α] [DecidableEq β] (P : α → Sort _) (e : α ≃ β) + (f : ∀ a, P a) (b : β) (x : P (e.symm b)) : + e.piCongrLeft' P (update f (e.symm b) x) = update (e.piCongrLeft' P f) b x := by ext b' rcases eq_or_ne b' b with (rfl | h) · simp @@ -1808,8 +1885,8 @@ theorem Pi_congr_left'_update [DecidableEq α] [DecidableEq β] (P : α → Sort #align function.Pi_congr_left'_update Function.Pi_congr_left'_update -theorem Pi_congr_left'_symm_update [DecidableEq α] [DecidableEq β] (P : α → Sort _) (e : α ≃ β) (f : ∀ b, P (e.symm b)) - (b : β) (x : P (e.symm b)) : +theorem Pi_congr_left'_symm_update [DecidableEq α] [DecidableEq β] (P : α → Sort _) (e : α ≃ β) + (f : ∀ b, P (e.symm b)) (b : β) (x : P (e.symm b)) : (e.piCongrLeft' P).symm (update f b x) = update ((e.piCongrLeft' P).symm f) (e.symm b) x := by simp [(e.Pi_congr_left' P).symm_apply_eq, Pi_congr_left'_update] #align function.Pi_congr_left'_symm_update Function.Pi_congr_left'_symm_update From c18a129a4f1f43ea085199e31f414e0692ebd4b6 Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Fri, 18 Nov 2022 22:04:38 +0000 Subject: [PATCH 038/127] up to line 632 --- Mathlib/Logic/Equiv/Basic.lean | 253 +++++++++++++++++---------------- 1 file changed, 130 insertions(+), 123 deletions(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 0159acbfafdbf..9a34fc5df9de4 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -9,6 +9,8 @@ import Mathlib.Data.Sigma.Basic import Mathlib.Data.Subtype import Mathlib.Data.Sum.Basic import Mathlib.Logic.Function.Conjugate +import Mathlib.Tactic.SplitIfs +import Mathlib.Tactic.Convert -- **TODO** remove these later set_option autoImplicit false @@ -20,6 +22,7 @@ In this file we continue the work on equivalences begun in `logic/equiv/defs.lea * canonical isomorphisms between various types: e.g., +-- **TODO** fix names - `equiv.sum_equiv_sigma_bool` is the canonical equivalence between the sum of two types `α ⊕ β` and the sigma-type `Σ b : bool, cond b α β`; @@ -28,11 +31,11 @@ In this file we continue the work on equivalences begun in `logic/equiv/defs.lea * operations on equivalences: e.g., - - `equiv.prod_congr ea eb : α₁ × β₁ ≃ α₂ × β₂`: combine two equivalences `ea : α₁ ≃ α₂` and - `eb : β₁ ≃ β₂` using `prod.map`. + - `Equiv.prodCongr ea eb : α₁ × β₁ ≃ α₂ × β₂`: combine two equivalences `ea : α₁ ≃ α₂` and + `eb : β₁ ≃ β₂` using `Prod.Map`. - More definitions of this kind can be found in other files. E.g., `data/equiv/transfer_instance` - does it for many algebraic type classes like `group`, `module`, etc. + More definitions of this kind can be found in other files. E.g., `Data/Equiv/TransferInstance` + does it for many algebraic type classes like `Group`, `Module`, etc. ## Tags @@ -48,7 +51,7 @@ variable {α : Sort u} {β : Sort v} {γ : Sort w} namespace Equiv -/-- `pprod α β` is equivalent to `α × β` -/ +/-- `PProd α β` is equivalent to `α × β` -/ @[simps apply symmApply] def pprodEquivProd {α β : Type _} : PProd α β ≃ α × β where toFun x := (x.1, x.2) @@ -57,8 +60,8 @@ def pprodEquivProd {α β : Type _} : PProd α β ≃ α × β where right_inv := fun _ => rfl #align equiv.pprod_equiv_prod Equiv.pprodEquivProd -/-- Product of two equivalences, in terms of `pprod`. If `α ≃ β` and `γ ≃ δ`, then -`pprod α γ ≃ pprod β δ`. -/ +/-- Product of two equivalences, in terms of `PProd`. If `α ≃ β` and `γ ≃ δ`, then +`PProd α γ ≃ PProd β δ`. -/ -- porting note: in Lean 3 this had @[congr]` @[simps apply] def pprodCongr {δ : Sort z} (e₁ : α ≃ β) (e₂ : γ ≃ δ) : PProd α γ ≃ PProd β δ where @@ -68,23 +71,23 @@ def pprodCongr {δ : Sort z} (e₁ : α ≃ β) (e₂ : γ ≃ δ) : PProd α γ right_inv := fun ⟨x, y⟩ => by simp #align equiv.pprod_congr Equiv.pprodCongr -/-- Combine two equivalences using `pprod` in the domain and `prod` in the codomain. -/ +/-- Combine two equivalences using `PProd` in the domain and `prod` in the codomain. -/ @[simps apply symmApply] def pprodProd {α₁ β₁ : Sort _} {α₂ β₂ : Type _} (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) : PProd α₁ β₁ ≃ α₂ × β₂ := (ea.pprodCongr eb).trans pprodEquivProd #align equiv.pprod_prod Equiv.pprodProd -/-- Combine two equivalences using `pprod` in the codomain and `prod` in the domain. -/ +/-- Combine two equivalences using `PProd` in the codomain and `prod` in the domain. -/ @[simps apply symmApply] def prodPProd {α₁ β₁ : Type _} {α₂ β₂ : Sort _} (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) : α₁ × β₁ ≃ PProd α₂ β₂ := (ea.symm.pprodProd eb.symm).symm -#align equiv.prod_pprod Equiv.prodPprod +#align equiv.prod_pprod Equiv.prodPProd -/-- `pprod α β` is equivalent to `plift α × plift β` -/ +/-- `PProd α β` is equivalent to `Plift α × Plift β` -/ @[simps apply symmApply] -def pprodEquivProdPlift {α β : Sort _} : PProd α β ≃ PLift α × PLift β := +def pprodEquivProdPLift {α β : Sort _} : PProd α β ≃ PLift α × PLift β := Equiv.plift.symm.pprodProd Equiv.plift.symm -#align equiv.pprod_equiv_prod_plift Equiv.pprodEquivProdPlift +#align equiv.pprod_equiv_prod_plift Equiv.pprodEquivProdPLift /-- Product of two equivalences. If `α₁ ≃ α₂` and `β₁ ≃ β₂`, then `α₁ × β₁ ≃ α₂ × β₂`. This is `prod.map` as an equivalence. -/ @@ -95,10 +98,10 @@ def prodCongr {α₁ β₁ α₂ β₂ : Type _} (e₁ : α₁ ≃ α₂) (e₂ #align equiv.prod_congr Equiv.prodCongr @[simp] -theorem prod_congr_symm {α₁ β₁ α₂ β₂ : Type _} (e₁ : α₁ ≃ α₂) (e₂ : β₁ ≃ β₂) : +theorem prodCongr_symm {α₁ β₁ α₂ β₂ : Type _} (e₁ : α₁ ≃ α₂) (e₂ : β₁ ≃ β₂) : (prodCongr e₁ e₂).symm = prodCongr e₁.symm e₂.symm := rfl -#align equiv.prod_congr_symm Equiv.prod_congr_symm +#align equiv.prod_congr_symm Equiv.prodCongr_symm /-- Type product is commutative up to an equivalence: `α × β ≃ β × α`. This is `prod.swap` as an equivalence.-/ @@ -107,19 +110,19 @@ def prodComm (α β : Type _) : α × β ≃ β × α := #align equiv.prod_comm Equiv.prodComm @[simp] -theorem coe_prod_comm (α β : Type _) : (⇑(prodComm α β) : α × β → β × α) = Prod.swap := +theorem coe_prodComm (α β : Type _) : (⇑(prodComm α β) : α × β → β × α) = Prod.swap := rfl -#align equiv.coe_prod_comm Equiv.coe_prod_comm +#align equiv.coe_prod_comm Equiv.coe_prodComm @[simp] -theorem prod_comm_apply {α β : Type _} (x : α × β) : prodComm α β x = x.swap := +theorem prodComm_apply {α β : Type _} (x : α × β) : prodComm α β x = x.swap := rfl -#align equiv.prod_comm_apply Equiv.prod_comm_apply +#align equiv.prod_comm_apply Equiv.prodComm_apply @[simp] -theorem prod_comm_symm (α β) : (prodComm α β).symm = prodComm β α := +theorem prodComm_symm (α β) : (prodComm α β).symm = prodComm β α := rfl -#align equiv.prod_comm_symm Equiv.prod_comm_symm +#align equiv.prod_comm_symm Equiv.prodComm_symm /-- Type product is associative up to an equivalence. -/ @[simps] @@ -147,7 +150,7 @@ def prodPUnit (α : Type _) : α × PUnit.{u + 1} ≃ α := /-- `punit` is a left identity for type product up to an equivalence. -/ @[simps] -def pUnitProd (α : Type _) : PUnit.{u + 1} × α ≃ α := +def punitProd (α : Type _) : PUnit.{u + 1} × α ≃ α := calc PUnit × α ≃ α × PUnit := prodComm _ _ _ ≃ α := prodPUnit _ @@ -160,13 +163,13 @@ def prodUnique (α : Type u) (β : Type v) [Unique β] : α × β ≃ α := #align equiv.prod_unique Equiv.prodUnique @[simp] -theorem coe_prod_unique {α β : Type _} [Unique β] : (⇑(prodUnique α β) : α × β → α) = Prod.fst := +theorem coe_prodUnique {α β : Type _} [Unique β] : (⇑(prodUnique α β) : α × β → α) = Prod.fst := rfl -#align equiv.coe_prod_unique Equiv.coe_prod_unique +#align equiv.coe_prod_unique Equiv.coe_prodUnique -theorem prod_unique_apply {α β : Type _} [Unique β] (x : α × β) : prodUnique α β x = x.1 := +theorem prodUnique_apply {α β : Type _} [Unique β] (x : α × β) : prodUnique α β x = x.1 := rfl -#align equiv.prod_unique_apply Equiv.prod_unique_apply +#align equiv.prod_unique_apply Equiv.prodUnique_apply @[simp] theorem prod_unique_symm_apply {α β : Type _} [Unique β] (x : α) : (prodUnique α β).symm x = (x, default) := @@ -175,23 +178,23 @@ theorem prod_unique_symm_apply {α β : Type _} [Unique β] (x : α) : (prodUniq /-- Any `unique` type is a left identity for type product up to equivalence. -/ def uniqueProd (α β : Type _) [Unique β] : β × α ≃ α := - ((equivPUnit.{v+1,v+1} β).prodCongr <| Equiv.refl α).trans <| pUnitProd α + ((equivPUnit.{v+1,v+1} β).prodCongr <| Equiv.refl α).trans <| punitProd α #align equiv.unique_prod Equiv.uniqueProd @[simp] -theorem coe_unique_prod {α β : Type _} [Unique β] : (⇑(uniqueProd α β) : β × α → α) = Prod.snd := +theorem coe_uniqueProd {α β : Type _} [Unique β] : (⇑(uniqueProd α β) : β × α → α) = Prod.snd := rfl -#align equiv.coe_unique_prod Equiv.coe_unique_prod +#align equiv.coe_unique_prod Equiv.coe_uniqueProd -theorem unique_prod_apply {α β : Type _} [Unique β] (x : β × α) : uniqueProd α β x = x.2 := +theorem uniqueProd_apply {α β : Type _} [Unique β] (x : β × α) : uniqueProd α β x = x.2 := rfl -#align equiv.unique_prod_apply Equiv.unique_prod_apply +#align equiv.unique_prod_apply Equiv.uniqueProd_apply @[simp] -theorem unique_prod_symm_apply {α β : Type _} [Unique β] (x : α) : +theorem uniqueProd_symm_apply {α β : Type _} [Unique β] (x : α) : (uniqueProd α β).symm x = (default, x) := rfl -#align equiv.unique_prod_symm_apply Equiv.unique_prod_symm_apply +#align equiv.unique_prod_symm_apply Equiv.uniqueProd_symm_apply /-- `empty` type is a right absorbing element for type product up to an equivalence. -/ def prodEmpty (α : Type _) : α × Empty ≃ Empty := @@ -204,9 +207,9 @@ def emptyProd (α : Type _) : Empty × α ≃ Empty := #align equiv.empty_prod Equiv.emptyProd /-- `pempty` type is a right absorbing element for type product up to an equivalence. -/ -def prodPempty (α : Type _) : α × PEmpty ≃ PEmpty := +def prodPEmpty (α : Type _) : α × PEmpty ≃ PEmpty := equivPEmpty _ -#align equiv.prod_pempty Equiv.prodPempty +#align equiv.prod_pempty Equiv.prodPEmpty /-- `pempty` type is a left absorbing element for type product up to an equivalence. -/ def pemptyProd (α : Type _) : PEmpty × α ≃ PEmpty := @@ -252,23 +255,23 @@ def sumPsum {α₁ β₁ : Type _} {α₂ β₂ : Sort _} (ea : α₁ ≃ α₂) #align equiv.sum_psum Equiv.sumPsum @[simp] -theorem sum_congr_trans {α₁ α₂ β₁ β₂ γ₁ γ₂ : Sort _} (e : α₁ ≃ β₁) (f : α₂ ≃ β₂) (g : β₁ ≃ γ₁) (h : β₂ ≃ γ₂) : +theorem sumCongr_trans {α₁ α₂ β₁ β₂ γ₁ γ₂ : Sort _} (e : α₁ ≃ β₁) (f : α₂ ≃ β₂) (g : β₁ ≃ γ₁) (h : β₂ ≃ γ₂) : (Equiv.sumCongr e f).trans (Equiv.sumCongr g h) = Equiv.sumCongr (e.trans g) (f.trans h) := by ext i cases i <;> rfl -#align equiv.sum_congr_trans Equiv.sum_congr_trans +#align equiv.sum_congr_trans Equiv.sumCongr_trans @[simp] -theorem sum_congr_symm {α β γ δ : Sort _} (e : α ≃ β) (f : γ ≃ δ) : +theorem sumCongr_symm {α β γ δ : Sort _} (e : α ≃ β) (f : γ ≃ δ) : (Equiv.sumCongr e f).symm = Equiv.sumCongr e.symm f.symm := rfl -#align equiv.sum_congr_symm Equiv.sum_congr_symm +#align equiv.sum_congr_symm Equiv.sumCongr_symm @[simp] -theorem sum_congr_refl {α β : Sort _} : Equiv.sumCongr (Equiv.refl α) (Equiv.refl β) = Equiv.refl (Sum α β) := by +theorem sumCongr_refl {α β : Sort _} : Equiv.sumCongr (Equiv.refl α) (Equiv.refl β) = Equiv.refl (Sum α β) := by ext i cases i <;> rfl -#align equiv.sum_congr_refl Equiv.sum_congr_refl +#align equiv.sum_congr_refl Equiv.sumCongr_refl namespace Perm @@ -279,35 +282,35 @@ def sumCongr {α β : Type _} (ea : Equiv.Perm α) (eb : Equiv.Perm β) : Equiv. #align equiv.perm.sum_congr Equiv.Perm.sumCongr @[simp] -theorem sum_congr_apply {α β : Type _} (ea : Equiv.Perm α) (eb : Equiv.Perm β) (x : Sum α β) : +theorem sumCongr_apply {α β : Type _} (ea : Equiv.Perm α) (eb : Equiv.Perm β) (x : Sum α β) : sumCongr ea eb x = Sum.map (⇑ea) (⇑eb) x := Equiv.sumCongr_apply ea eb x -#align equiv.perm.sum_congr_apply Equiv.Perm.sum_congr_apply +#align equiv.perm.sum_congr_apply Equiv.Perm.sumCongr_apply @[simp] -theorem sum_congr_trans {α β : Sort _} (e : Equiv.Perm α) (f : Equiv.Perm β) (g : Equiv.Perm α) (h : Equiv.Perm β) : +theorem sumCongr_trans {α β : Sort _} (e : Equiv.Perm α) (f : Equiv.Perm β) (g : Equiv.Perm α) (h : Equiv.Perm β) : (sumCongr e f).trans (sumCongr g h) = sumCongr (e.trans g) (f.trans h) := - Equiv.sum_congr_trans e f g h -#align equiv.perm.sum_congr_trans Equiv.Perm.sum_congr_trans + Equiv.sumCongr_trans e f g h +#align equiv.perm.sum_congr_trans Equiv.Perm.sumCongr_trans @[simp] -theorem sum_congr_symm {α β : Sort _} (e : Equiv.Perm α) (f : Equiv.Perm β) : +theorem sumCongr_symm {α β : Sort _} (e : Equiv.Perm α) (f : Equiv.Perm β) : (sumCongr e f).symm = sumCongr e.symm f.symm := - Equiv.sum_congr_symm e f -#align equiv.perm.sum_congr_symm Equiv.Perm.sum_congr_symm + Equiv.sumCongr_symm e f +#align equiv.perm.sum_congr_symm Equiv.Perm.sumCongr_symm @[simp] -theorem sum_congr_refl {α β : Sort _} : sumCongr (Equiv.refl α) (Equiv.refl β) = Equiv.refl (Sum α β) := - Equiv.sum_congr_refl -#align equiv.perm.sum_congr_refl Equiv.Perm.sum_congr_refl +theorem sumCongr_refl {α β : Sort _} : sumCongr (Equiv.refl α) (Equiv.refl β) = Equiv.refl (Sum α β) := + Equiv.sumCongr_refl +#align equiv.perm.sum_congr_refl Equiv.Perm.sumCongr_refl end Perm /-- `bool` is equivalent the sum of two `punit`s. -/ -def boolEquivPunitSumPunit : Bool ≃ Sum PUnit.{u + 1} PUnit.{v + 1} := +def boolEquivPUnitSumPUnit : Bool ≃ Sum PUnit.{u + 1} PUnit.{v + 1} := ⟨fun b => cond b (inr PUnit.unit) (inl PUnit.unit), Sum.elim (fun _ => false) fun _ => true, fun b => by cases b <;> rfl, fun s => by rcases s with (⟨⟨⟩⟩ | ⟨⟨⟩⟩) <;> rfl⟩ -#align equiv.bool_equiv_punit_sum_punit Equiv.boolEquivPunitSumPunit +#align equiv.bool_equiv_punit_sum_punit Equiv.boolEquivPUnitSumPUnit /-- Sum of types is commutative up to an equivalence. This is `sum.swap` as an equivalence. -/ @[simps (config := { fullyApplied := false }) apply] @@ -316,9 +319,9 @@ def sumComm (α β : Type _) : Sum α β ≃ Sum β α := #align equiv.sum_comm Equiv.sumComm @[simp] -theorem sum_comm_symm (α β) : (sumComm α β).symm = sumComm β α := +theorem sumComm_symm (α β) : (sumComm α β).symm = sumComm β α := rfl -#align equiv.sum_comm_symm Equiv.sum_comm_symm +#align equiv.sum_comm_symm Equiv.sumComm_symm /-- Sum of types is associative up to an equivalence. -/ def sumAssoc (α β γ : Type _) : Sum (Sum α β) γ ≃ Sum α (Sum β γ) := @@ -328,34 +331,34 @@ def sumAssoc (α β γ : Type _) : Sum (Sum α β) γ ≃ Sum α (Sum β γ) := #align equiv.sum_assoc Equiv.sumAssoc @[simp] -theorem sum_assoc_apply_inl_inl {α β γ} (a) : sumAssoc α β γ (inl (inl a)) = inl a := +theorem sumAssoc_apply_inl_inl {α β γ} (a) : sumAssoc α β γ (inl (inl a)) = inl a := rfl -#align equiv.sum_assoc_apply_inl_inl Equiv.sum_assoc_apply_inl_inl +#align equiv.sum_assoc_apply_inl_inl Equiv.sumAssoc_apply_inl_inl @[simp] -theorem sum_assoc_apply_inl_inr {α β γ} (b) : sumAssoc α β γ (inl (inr b)) = inr (inl b) := +theorem sumAssoc_apply_inl_inr {α β γ} (b) : sumAssoc α β γ (inl (inr b)) = inr (inl b) := rfl -#align equiv.sum_assoc_apply_inl_inr Equiv.sum_assoc_apply_inl_inr +#align equiv.sum_assoc_apply_inl_inr Equiv.sumAssoc_apply_inl_inr @[simp] -theorem sum_assoc_apply_inr {α β γ} (c) : sumAssoc α β γ (inr c) = inr (inr c) := +theorem sumAssoc_apply_inr {α β γ} (c) : sumAssoc α β γ (inr c) = inr (inr c) := rfl -#align equiv.sum_assoc_apply_inr Equiv.sum_assoc_apply_inr +#align equiv.sum_assoc_apply_inr Equiv.sumAssoc_apply_inr @[simp] -theorem sum_assoc_symm_apply_inl {α β γ} (a) : (sumAssoc α β γ).symm (inl a) = inl (inl a) := +theorem sumAssoc_symm_apply_inl {α β γ} (a) : (sumAssoc α β γ).symm (inl a) = inl (inl a) := rfl -#align equiv.sum_assoc_symm_apply_inl Equiv.sum_assoc_symm_apply_inl +#align equiv.sum_assoc_symm_apply_inl Equiv.sumAssoc_symm_apply_inl @[simp] -theorem sum_assoc_symm_apply_inr_inl {α β γ} (b) : (sumAssoc α β γ).symm (inr (inl b)) = inl (inr b) := +theorem sumAssoc_symm_apply_inr_inl {α β γ} (b) : (sumAssoc α β γ).symm (inr (inl b)) = inl (inr b) := rfl -#align equiv.sum_assoc_symm_apply_inr_inl Equiv.sum_assoc_symm_apply_inr_inl +#align equiv.sum_assoc_symm_apply_inr_inl Equiv.sumAssoc_symm_apply_inr_inl @[simp] -theorem sum_assoc_symm_apply_inr_inr {α β γ} (c) : (sumAssoc α β γ).symm (inr (inr c)) = inr c := +theorem sumAssoc_symm_apply_inr_inr {α β γ} (c) : (sumAssoc α β γ).symm (inr (inr c)) = inr c := rfl -#align equiv.sum_assoc_symm_apply_inr_inr Equiv.sum_assoc_symm_apply_inr_inr +#align equiv.sum_assoc_symm_apply_inr_inr Equiv.sumAssoc_symm_apply_inr_inr /-- Sum with `empty` is equivalent to the original type. -/ @[simps symmApply] @@ -367,9 +370,9 @@ def sumEmpty (α β : Type _) [IsEmpty β] : Sum α β ≃ α := #align equiv.sum_empty Equiv.sumEmpty @[simp] -theorem sum_empty_apply_inl {α β : Type _} [IsEmpty β] (a : α) : sumEmpty α β (Sum.inl a) = a := +theorem sumEmpty_apply_inl {α β : Type _} [IsEmpty β] (a : α) : sumEmpty α β (Sum.inl a) = a := rfl -#align equiv.sum_empty_apply_inl Equiv.sum_empty_apply_inl +#align equiv.sum_empty_apply_inl Equiv.sumEmpty_apply_inl /-- The sum of `empty` with any `Sort*` is equivalent to the right summand. -/ @[simps symmApply] @@ -378,55 +381,55 @@ def emptySum (α β : Type _) [IsEmpty α] : Sum α β ≃ β := #align equiv.empty_sum Equiv.emptySum @[simp] -theorem empty_sum_apply_inr {α β : Type _} [IsEmpty α] (b : β) : emptySum α β (Sum.inr b) = b := +theorem emptySum_apply_inr {α β : Type _} [IsEmpty α] (b : β) : emptySum α β (Sum.inr b) = b := rfl -#align equiv.empty_sum_apply_inr Equiv.empty_sum_apply_inr +#align equiv.empty_sum_apply_inr Equiv.emptySum_apply_inr /-- `option α` is equivalent to `α ⊕ punit` -/ -def optionEquivSumPunit (α : Type _) : Option α ≃ Sum α PUnit.{u + 1} := +def optionEquivSumPUnit (α : Type _) : Option α ≃ Sum α PUnit.{u + 1} := ⟨fun o => o.elim (inr PUnit.unit) inl, fun s => s.elim some fun _ => none, fun o => by cases o <;> rfl, fun s => by rcases s with (_ | ⟨⟨⟩⟩) <;> rfl⟩ -#align equiv.option_equiv_sum_punit Equiv.optionEquivSumPunit +#align equiv.option_equiv_sum_punit Equiv.optionEquivSumPUnit @[simp] -theorem option_equiv_sum_punit_none {α} : optionEquivSumPunit α none = Sum.inr PUnit.unit := +theorem optionEquivSumPUnit_none {α} : optionEquivSumPUnit α none = Sum.inr PUnit.unit := rfl -#align equiv.option_equiv_sum_punit_none Equiv.option_equiv_sum_punit_none -#exit +#align equiv.option_equiv_sum_punit_none Equiv.optionEquivSumPUnit_none @[simp] -theorem option_equiv_sum_punit_some {α} (a) : optionEquivSumPunit α (some a) = Sum.inl a := +theorem optionEquivSumPUnit_some {α} (a) : optionEquivSumPUnit α (some a) = Sum.inl a := rfl -#align equiv.option_equiv_sum_punit_some Equiv.option_equiv_sum_punit_some +#align equiv.option_equiv_sum_punit_some Equiv.optionEquivSumPUnit_some @[simp] -theorem option_equiv_sum_punit_coe {α} (a : α) : optionEquivSumPunit α a = Sum.inl a := +theorem optionEquivSumPUnit_coe {α} (a : α) : optionEquivSumPUnit α a = Sum.inl a := rfl -#align equiv.option_equiv_sum_punit_coe Equiv.option_equiv_sum_punit_coe +#align equiv.option_equiv_sum_punit_coe Equiv.optionEquivSumPUnit_coe @[simp] -theorem option_equiv_sum_punit_symm_inl {α} (a) : (optionEquivSumPunit α).symm (Sum.inl a) = a := +theorem optionEquivSumPUnit_symm_inl {α} (a) : (optionEquivSumPUnit α).symm (Sum.inl a) = a := rfl -#align equiv.option_equiv_sum_punit_symm_inl Equiv.option_equiv_sum_punit_symm_inl +#align equiv.option_equiv_sum_punit_symm_inl Equiv.optionEquivSumPUnit_symm_inl @[simp] -theorem option_equiv_sum_punit_symm_inr {α} (a) : (optionEquivSumPunit α).symm (Sum.inr a) = none := +theorem optionEquivSumPUnit_symm_inr {α} (a) : (optionEquivSumPUnit α).symm (Sum.inr a) = none := rfl -#align equiv.option_equiv_sum_punit_symm_inr Equiv.option_equiv_sum_punit_symm_inr +#align equiv.option_equiv_sum_punit_symm_inr Equiv.optionEquivSumPUnit_symm_inr /-- The set of `x : option α` such that `is_some x` is equivalent to `α`. -/ @[simps] def optionIsSomeEquiv (α : Type _) : { x : Option α // x.isSome } ≃ α where - toFun o := Option.get o.2 - invFun x := ⟨some x, by decide⟩ - left_inv o := Subtype.eq <| Option.some_get _ - right_inv x := Option.get_some _ _ + toFun o := Option.get _ o.2 + invFun x := ⟨some x, rfl⟩ + left_inv _ := Subtype.eq <| Option.some_get _ + right_inv _ := Option.get_some _ _ #align equiv.option_is_some_equiv Equiv.optionIsSomeEquiv /-- The product over `option α` of `β a` is the binary product of the product over `α` of `β (some α)` and `β none` -/ @[simps] -def piOptionEquivProd {α : Type _} {β : Option α → Type _} : (∀ a : Option α, β a) ≃ β none × ∀ a : α, β (some a) where +def piOptionEquivProd {α : Type _} {β : Option α → Type _} : + (∀ a : Option α, β a) ≃ β none × ∀ a : α, β (some a) where toFun f := (f none, fun a => f (some a)) invFun x a := Option.casesOn a x.fst x.snd left_inv f := funext fun a => by cases a <;> rfl @@ -440,8 +443,8 @@ difficulty. -/ def sumEquivSigmaBool (α β : Type u) : Sum α β ≃ Σb : Bool, cond b α β := ⟨fun s => s.elim (fun x => ⟨true, x⟩) fun x => ⟨false, x⟩, fun s => match s with - | ⟨tt, a⟩ => inl a - | ⟨ff, b⟩ => inr b, + | ⟨true, a⟩ => inl a + | ⟨false, b⟩ => inr b, fun s => by cases s <;> rfl, fun s => by rcases s with ⟨_ | _, _⟩ <;> rfl⟩ #align equiv.sum_equiv_sigma_bool Equiv.sumEquivSigmaBool @@ -450,7 +453,7 @@ def sumEquivSigmaBool (α β : Type u) : Sum α β ≃ Σb : Bool, cond b α β the type of all fibres of `f` and the total space `α`. -/ @[simps] def sigmaFiberEquiv {α β : Type _} (f : α → β) : (Σy : β, { x // f x = y }) ≃ α := - ⟨fun x => ↑x.2, fun x => ⟨f x, x, rfl⟩, fun ⟨y, x, rfl⟩ => rfl, fun x => rfl⟩ + ⟨fun x => ↑x.2, fun x => ⟨f x, x, rfl⟩, fun ⟨_, _, rfl⟩ => rfl, fun _ => rfl⟩ #align equiv.sigma_fiber_equiv Equiv.sigmaFiberEquiv end @@ -464,42 +467,45 @@ is naturally equivalent to `α`. See `subtype_or_equiv` for sum types over subtypes `{x // p x}` and `{x // q x}` that are not necessarily `is_compl p q`. -/ def sumCompl {α : Type _} (p : α → Prop) [DecidablePred p] : Sum { a // p a } { a // ¬p a } ≃ α where - toFun := Sum.elim coe coe + toFun := Sum.elim Subtype.val Subtype.val invFun a := if h : p a then Sum.inl ⟨a, h⟩ else Sum.inr ⟨a, h⟩ - left_inv := by rintro (⟨x, hx⟩ | ⟨x, hx⟩) <;> dsimp <;> [rw [dif_pos], rw [dif_neg]] + left_inv := by + rintro (⟨x, hx⟩ | ⟨x, hx⟩) <;> dsimp; + { rw [dif_pos] } + { rw [dif_neg] } right_inv a := by dsimp split_ifs <;> rfl #align equiv.sum_compl Equiv.sumCompl @[simp] -theorem sum_compl_apply_inl {α : Type _} (p : α → Prop) [DecidablePred p] (x : { a // p a }) : +theorem sumCompl_apply_inl {α : Type _} (p : α → Prop) [DecidablePred p] (x : { a // p a }) : sumCompl p (Sum.inl x) = x := rfl -#align equiv.sum_compl_apply_inl Equiv.sum_compl_apply_inl +#align equiv.sum_compl_apply_inl Equiv.sumCompl_apply_inl @[simp] -theorem sum_compl_apply_inr {α : Type _} (p : α → Prop) [DecidablePred p] (x : { a // ¬p a }) : +theorem sumCompl_apply_inr {α : Type _} (p : α → Prop) [DecidablePred p] (x : { a // ¬p a }) : sumCompl p (Sum.inr x) = x := rfl -#align equiv.sum_compl_apply_inr Equiv.sum_compl_apply_inr +#align equiv.sum_compl_apply_inr Equiv.sumCompl_apply_inr @[simp] -theorem sum_compl_apply_symm_of_pos {α : Type _} (p : α → Prop) [DecidablePred p] (a : α) (h : p a) : +theorem sumCompl_apply_symm_of_pos {α : Type _} (p : α → Prop) [DecidablePred p] (a : α) (h : p a) : (sumCompl p).symm a = Sum.inl ⟨a, h⟩ := dif_pos h -#align equiv.sum_compl_apply_symm_of_pos Equiv.sum_compl_apply_symm_of_pos +#align equiv.sum_compl_apply_symm_of_pos Equiv.sumCompl_apply_symm_of_pos @[simp] -theorem sum_compl_apply_symm_of_neg {α : Type _} (p : α → Prop) [DecidablePred p] (a : α) (h : ¬p a) : - (sumCompl p).symm a = Sum.inr ⟨a, h⟩ := +theorem sumCompl_apply_symm_of_neg {α : Type _} (p : α → Prop) [DecidablePred p] (a : α) + (h : ¬p a) : (sumCompl p).symm a = Sum.inr ⟨a, h⟩ := dif_neg h -#align equiv.sum_compl_apply_symm_of_neg Equiv.sum_compl_apply_symm_of_neg +#align equiv.sum_compl_apply_symm_of_neg Equiv.sumCompl_apply_symm_of_neg /-- Combines an `equiv` between two subtypes with an `equiv` between their complements to form a permutation. -/ -def subtypeCongr {α : Type _} {p q : α → Prop} [DecidablePred p] [DecidablePred q] (e : { x // p x } ≃ { x // q x }) - (f : { x // ¬p x } ≃ { x // ¬q x }) : Perm α := +def subtypeCongr {α : Type _} {p q : α → Prop} [DecidablePred p] [DecidablePred q] + (e : { x // p x } ≃ { x // q x }) (f : { x // ¬p x } ≃ { x // ¬q x }) : Perm α := (sumCompl p).symm.trans ((sumCongr e f).trans (sumCompl q)) #align equiv.subtype_congr Equiv.subtypeCongr @@ -515,30 +521,29 @@ def Perm.subtypeCongr : Equiv.Perm ε := permCongr (sumCompl p) (sumCongr ep en) #align equiv.perm.subtype_congr Equiv.Perm.subtypeCongr -theorem Perm.subtypeCongr.apply (a : ε) : ep.subtypeCongr en a = if h : p a then ep ⟨a, h⟩ else en ⟨a, h⟩ := by - by_cases h:p a <;> simp [perm.subtype_congr, h] +theorem Perm.subtypeCongr.apply (a : ε) : ep.subtypeCongr en a = + if h : p a then (ep ⟨a, h⟩ : ε) else en ⟨a, h⟩ := by + by_cases h : p a <;> simp [Perm.subtypeCongr, h] #align equiv.perm.subtype_congr.apply Equiv.Perm.subtypeCongr.apply @[simp] theorem Perm.subtypeCongr.left_apply {a : ε} (h : p a) : ep.subtypeCongr en a = ep ⟨a, h⟩ := by - simp [perm.subtype_congr.apply, h] + simp [Perm.subtypeCongr.apply, h] #align equiv.perm.subtype_congr.left_apply Equiv.Perm.subtypeCongr.left_apply @[simp] -theorem Perm.subtypeCongr.left_apply_subtype (a : { a // p a }) : ep.subtypeCongr en a = ep a := by - convert perm.subtype_congr.left_apply _ _ a.property - simp +theorem Perm.subtypeCongr.left_apply_subtype (a : { a // p a }) : ep.subtypeCongr en a = ep a := + Perm.subtypeCongr.left_apply ep en a.property #align equiv.perm.subtype_congr.left_apply_subtype Equiv.Perm.subtypeCongr.left_apply_subtype @[simp] theorem Perm.subtypeCongr.right_apply {a : ε} (h : ¬p a) : ep.subtypeCongr en a = en ⟨a, h⟩ := by - simp [perm.subtype_congr.apply, h] + simp [Perm.subtypeCongr.apply, h] #align equiv.perm.subtype_congr.right_apply Equiv.Perm.subtypeCongr.right_apply @[simp] -theorem Perm.subtypeCongr.right_apply_subtype (a : { a // ¬p a }) : ep.subtypeCongr en a = en a := by - convert perm.subtype_congr.right_apply _ _ a.property - simp +theorem Perm.subtypeCongr.right_apply_subtype (a : { a // ¬p a }) : ep.subtypeCongr en a = en a := + Perm.subtypeCongr.right_apply ep en a.property #align equiv.perm.subtype_congr.right_apply_subtype Equiv.Perm.subtypeCongr.right_apply_subtype @[simp] @@ -553,10 +558,10 @@ theorem Perm.subtypeCongr.symm : (ep.subtypeCongr en).symm = Perm.subtypeCongr e ext x by_cases h:p x · have : p (ep.symm ⟨x, h⟩) := Subtype.property _ - simp [perm.subtype_congr.apply, h, symm_apply_eq, this] + simp [Perm.subtypeCongr.apply, h, symm_apply_eq, this] · have : ¬p (en.symm ⟨x, h⟩) := Subtype.property (en.symm _) - simp [perm.subtype_congr.apply, h, symm_apply_eq, this] + simp [Perm.subtypeCongr.apply, h, symm_apply_eq, this] #align equiv.perm.subtype_congr.symm Equiv.Perm.subtypeCongr.symm @@ -566,10 +571,10 @@ theorem Perm.subtypeCongr.trans : ext x by_cases h:p x · have : p (ep ⟨x, h⟩) := Subtype.property _ - simp [perm.subtype_congr.apply, h, this] + simp [Perm.subtypeCongr.apply, h, this] · have : ¬p (en ⟨x, h⟩) := Subtype.property (en _) - simp [perm.subtype_congr.apply, h, symm_apply_eq, this] + simp [Perm.subtypeCongr.apply, h, symm_apply_eq, this] #align equiv.perm.subtype_congr.trans Equiv.Perm.subtypeCongr.trans @@ -583,14 +588,16 @@ variable (p : α → Prop) [DecidablePred p] (x₀ : { a // p a } → β) the subtype of functions `x : α → β` that agree with `x₀` on the subtype `{a // p a}` is naturally equivalent to the type of functions `{a // ¬ p a} → β`. -/ @[simps] -def subtypePreimage : { x : α → β // x ∘ coe = x₀ } ≃ ({ a // ¬p a } → β) where - toFun (x : { x : α → β // x ∘ coe = x₀ }) a := (x : α → β) a +def subtypePreimage : { x : α → β // x ∘ Subtype.val = x₀ } ≃ ({ a // ¬p a } → β) where + toFun (x : { x : α → β // x ∘ Subtype.val = x₀ }) a := (x : α → β) a invFun x := ⟨fun a => if h : p a then x₀ ⟨a, h⟩ else x ⟨a, h⟩, funext fun ⟨a, h⟩ => dif_pos h⟩ left_inv := fun ⟨x, hx⟩ => Subtype.val_injective <| funext fun a => by dsimp - split_ifs <;> [rw [← hx], skip] <;> rfl + split_ifs + { rw [← hx]; rfl } + { rfl } right_inv x := funext fun ⟨a, h⟩ => show dite (p a) _ _ = _ by From 2ef04b462fdf17d0d45854f3f9b129046b7f2d5b Mon Sep 17 00:00:00 2001 From: ruben vorster Date: Fri, 18 Nov 2022 22:07:40 +0000 Subject: [PATCH 039/127] "fixed" weird universe errors in prodUnique, etc. --- Mathlib/Logic/Equiv/Basic.lean | 10 +++++----- 1 file changed, 5 insertions(+), 5 deletions(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index c3533573537df..fadbbc73f58fe 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -161,11 +161,11 @@ def pUnitProd (α : Type _) : PUnit.{u + 1} × α ≃ α := -- **TODO** fix weird universe error /-- Any `unique` type is a right identity for type product up to equivalence. -/ def prodUnique (α : Type u) (β : Type v) [Unique β] : α × β ≃ α := - ((Equiv.refl α).prodCongr <| equivPUnit β).trans <| prodPUnit α + ((Equiv.refl α).prodCongr <| equivPUnit.{v+1,v+1} β).trans <| prodPUnit α #align equiv.prod_unique Equiv.prodUnique @[simp] -theorem coe_prod_unique {α β : Type _} [Unique β] : (⇑(prodUnique α β) : α × β → α) = Prod.fst := +theorem coe_prod_unique {α β : Type _} [Unique β] : (⇑(prodUnique α β) : α × β → α) = Prod.fst := rfl #align equiv.coe_prod_unique Equiv.coe_prod_unique @@ -181,7 +181,7 @@ theorem prod_unique_symm_apply {α β : Type _} [Unique β] (x : α) /-- Any `unique` type is a left identity for type product up to equivalence. -/ def uniqueProd (α β : Type _) [Unique β] : β × α ≃ α := - ((equivPunit β).prodCongr <| Equiv.refl α).trans <| punitProd α + ((equivPUnit.{v+1,v+1} β).prodCongr <| Equiv.refl α).trans <| pUnitProd α #align equiv.unique_prod Equiv.uniqueProd @[simp] @@ -198,7 +198,7 @@ theorem unique_prod_symm_apply {α β : Type _} [Unique β] (x : α) : (uniqueProd α β).symm x = (default, x) := rfl #align equiv.unique_prod_symm_apply Equiv.unique_prod_symm_apply -#exit + /-- `empty` type is a right absorbing element for type product up to an equivalence. -/ def prodEmpty (α : Type _) : α × Empty ≃ Empty := equivEmpty _ @@ -279,7 +279,7 @@ theorem sum_congr_refl {α β : Sort _} : ext i cases i <;> rfl #align equiv.sum_congr_refl Equiv.sum_congr_refl - +#exit namespace Perm /-- Combine a permutation of `α` and of `β` into a permutation of `α ⊕ β`. -/ From cc9f085da9f573a697f1dfd0cd55e88e99c84bc0 Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Fri, 18 Nov 2022 22:14:26 +0000 Subject: [PATCH 040/127] up to line 650 --- Mathlib/Logic/Equiv/Basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 542ad331941df..e2185d1f9cc80 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -644,7 +644,7 @@ def piCongrRight {α} {β₁ β₂ : α → Sort _} (F : ∀ a, β₁ a ≃ β This is `function.swap` as an `equiv`. -/ @[simps apply] def piComm {α β} (φ : α → β → Sort _) : (∀ a b, φ a b) ≃ ∀ b a, φ a b := - ⟨swap, swap, fun x => rfl, fun y => rfl⟩ + ⟨swap, swap, fun _ => rfl, fun _ => rfl⟩ #align equiv.Pi_comm Equiv.piComm @[simp] From 92aff859b1f981eb5c5af5413f72082106c386f3 Mon Sep 17 00:00:00 2001 From: ruben vorster Date: Fri, 18 Nov 2022 22:26:31 +0000 Subject: [PATCH 041/127] renaming --- Mathlib/Logic/Equiv/Basic.lean | 24 ++++++++++++------------ 1 file changed, 12 insertions(+), 12 deletions(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 08c7556ef31de..0126b06a05c89 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -24,10 +24,10 @@ In this file we continue the work on equivalences begun in `Logic/Equiv/Defs.lea * canonical isomorphisms between various types: e.g., -- **TODO** fix names - - `equiv.sum_equiv_sigma_bool` is the canonical equivalence between the sum of two types `α ⊕ β` + - `Equiv.sum_equiv_sigma_bool` is the canonical equivalence between the sum of two types `α ⊕ β` and the sigma-type `Σ b : bool, cond b α β`; - - `equiv.prod_sum_distrib : α × (β ⊕ γ) ≃ (α × β) ⊕ (α × γ)` shows that type product and type sum + - `Equiv.prod_sum_distrib : α × (β ⊕ γ) ≃ (α × β) ⊕ (α × γ)` shows that type product and type sum satisfy the distributive law up to a canonical equivalence; * operations on equivalences: e.g., @@ -275,7 +275,7 @@ theorem sumCongr_symm {α β γ δ : Sort _} (e : α ≃ β) (f : γ ≃ δ) : #align equiv.sum_congr_symm Equiv.sumCongr_symm @[simp] -theorem sum_congr_refl {α β : Sort _} : +theorem sumCongr_refl {α β : Sort _} : Equiv.sumCongr (Equiv.refl α) (Equiv.refl β) = Equiv.refl (Sum α β) := by ext i cases i <;> rfl @@ -308,10 +308,10 @@ theorem sumCongr_symm {α β : Sort _} (e : Equiv.Perm α) (f : Equiv.Perm β) : #align equiv.perm.sum_congr_symm Equiv.Perm.sumCongr_symm @[simp] -theorem sum_congr_refl {α β : Sort _} : +theorem sumCongr_refl {α β : Sort _} : sumCongr (Equiv.refl α) (Equiv.refl β) = Equiv.refl (Sum α β) := - Equiv.sum_congr_refl -#align equiv.perm.sum_congr_refl Equiv.Perm.sum_congr_refl + Equiv.sumCongr_refl +#align equiv.perm.sum_congr_refl Equiv.Perm.sumCongr_refl end Perm @@ -649,9 +649,9 @@ def piComm {α β} (φ : α → β → Sort _) : (∀ a b, φ a b) ≃ ∀ b a, -- up to here **TODO** remove this @[simp] -theorem Pi_comm_symm {α β} {φ : α → β → Sort _} : (piComm φ).symm = (Pi_comm <| swap φ) := +theorem piComm_symm {α β} {φ : α → β → Sort _} : (piComm φ).symm = (piComm <| swap φ) := rfl -#align equiv.Pi_comm_symm Equiv.Pi_comm_symm +#align equiv.Pi_comm_symm Equiv.piComm_symm /-- Dependent `curry` equivalence: the type of dependent functions on `Σ i, β i` is equivalent to the type of dependent functions of two arguments (i.e., functions to the space of functions). @@ -1598,7 +1598,7 @@ theorem swap_apply_ne_self_iff {a b x : α} : swap a b x ≠ x ↔ a ≠ b ∧ ( namespace Perm @[simp] -theorem sum_congr_swap_refl {α β : Sort _} [DecidableEq α] [DecidableEq β] (i j : α) : +theorem sumCongr_swap_refl {α β : Sort _} [DecidableEq α] [DecidableEq β] (i j : α) : Equiv.Perm.sumCongr (Equiv.swap i j) (Equiv.refl β) = Equiv.swap (Sum.inl i) (Sum.inl j) := by ext x cases x @@ -1607,10 +1607,10 @@ theorem sum_congr_swap_refl {α β : Sort _} [DecidableEq α] [DecidableEq β] ( · simp [Sum.map, swap_apply_of_ne_of_ne] -#align equiv.perm.sum_congr_swap_refl Equiv.Perm.sum_congr_swap_refl +#align equiv.perm.sum_congr_swap_refl Equiv.Perm.sumCongr_swap_refl @[simp] -theorem sum_congr_refl_swap {α β : Sort _} [DecidableEq α] [DecidableEq β] (i j : β) : +theorem sumCongr_refl_swap {α β : Sort _} [DecidableEq α] [DecidableEq β] (i j : β) : Equiv.Perm.sumCongr (Equiv.refl α) (Equiv.swap i j) = Equiv.swap (Sum.inr i) (Sum.inr j) := by ext x cases x @@ -1619,7 +1619,7 @@ theorem sum_congr_refl_swap {α β : Sort _} [DecidableEq α] [DecidableEq β] ( · simp [Sum.map, swap_apply_def] split_ifs <;> rfl -#align equiv.perm.sum_congr_refl_swap Equiv.Perm.sum_congr_refl_swap +#align equiv.perm.sumCongr_refl_swap Equiv.Perm.sumCongr_refl_swap end Perm From ae8e4314f0f17c97e884bcfa25ec24f28245a819 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Sat, 19 Nov 2022 13:07:58 +1100 Subject: [PATCH 042/127] finish renames and aligns --- Mathlib/Logic/Equiv/Basic.lean | 392 ++++++++++++++++----------------- 1 file changed, 194 insertions(+), 198 deletions(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 0126b06a05c89..86bcd4927f6b0 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -23,11 +23,10 @@ In this file we continue the work on equivalences begun in `Logic/Equiv/Defs.lea * canonical isomorphisms between various types: e.g., --- **TODO** fix names - - `Equiv.sum_equiv_sigma_bool` is the canonical equivalence between the sum of two types `α ⊕ β` + - `Equiv.sumEquivSigmaBool` is the canonical equivalence between the sum of two types `α ⊕ β` and the sigma-type `Σ b : bool, cond b α β`; - - `Equiv.prod_sum_distrib : α × (β ⊕ γ) ≃ (α × β) ⊕ (α × γ)` shows that type product and type sum + - `Equiv.prodSumDistrib : α × (β ⊕ γ) ≃ (α × β) ⊕ (α × γ)` shows that type product and type sum satisfy the distributive law up to a canonical equivalence; * operations on equivalences: e.g., @@ -35,8 +34,9 @@ In this file we continue the work on equivalences begun in `Logic/Equiv/Defs.lea - `Equiv.prodCongr ea eb : α₁ × β₁ ≃ α₂ × β₂`: combine two equivalences `ea : α₁ ≃ α₂` and `eb : β₁ ≃ β₂` using `Prod.Map`. - More definitions of this kind can be found in other files. E.g., `Data/Equiv/TransferInstance` - does it for many algebraic type classes like `Group`, `Module`, etc. + More definitions of this kind can be found in other files. + E.g., `Data/Equiv/TransferInstance.lean` does it for many algebraic type classes like + `Group`, `Module`, etc. ## Tags @@ -74,8 +74,8 @@ def pprodCongr {δ : Sort z} (e₁ : α ≃ β) (e₂ : γ ≃ δ) : PProd α γ /-- Combine two equivalences using `PProd` in the domain and `prod` in the codomain. -/ @[simps apply symmApply] -def pprodProd {α₁ β₁ : Sort _} {α₂ β₂ : Type _} (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) -: PProd α₁ β₁ ≃ α₂ × β₂ := +def pprodProd {α₁ β₁ : Sort _} {α₂ β₂ : Type _} (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) : + PProd α₁ β₁ ≃ α₂ × β₂ := (ea.pprodCongr eb).trans pprodEquivProd #align equiv.pprod_prod Equiv.pprodProd @@ -149,7 +149,7 @@ section @[simps] def prodPUnit (α : Type _) : α × PUnit.{u + 1} ≃ α := ⟨fun p => p.1, fun a => (a, PUnit.unit), fun ⟨_, PUnit.unit⟩ => rfl, fun _ => rfl⟩ -#align equiv.prod_punit Equiv.prodPunit +#align equiv.prod_punit Equiv.prodPUnit /-- `punit` is a left identity for type product up to an equivalence. -/ @[simps] @@ -157,7 +157,6 @@ def punitProd (α : Type _) : PUnit.{u + 1} × α ≃ α := calc PUnit × α ≃ α × PUnit := prodComm _ _ _ ≃ α := prodPUnit _ - #align equiv.punit_prod Equiv.punitProd /-- Any `unique` type is a right identity for type product up to equivalence. -/ @@ -175,10 +174,10 @@ theorem prodUnique_apply {α β : Type _} [Unique β] (x : α × β) : prodUniqu #align equiv.prod_unique_apply Equiv.prodUnique_apply @[simp] -theorem prod_unique_symm_apply {α β : Type _} [Unique β] (x : α) -: (prodUnique α β).symm x = (x, default) := +theorem prodUnique_symm_apply {α β : Type _} [Unique β] (x : α) : + (prodUnique α β).symm x = (x, default) := rfl -#align equiv.prod_unique_symm_apply Equiv.prod_unique_symm_apply +#align equiv.prod_unique_symm_apply Equiv.prodUnique_symm_apply /-- Any `unique` type is a left identity for type product up to equivalence. -/ def uniqueProd (α β : Type _) [Unique β] : β × α ≃ α := @@ -255,10 +254,10 @@ def psumSum {α₁ β₁ : Sort _} {α₂ β₂ : Type _} (ea : α₁ ≃ α₂) #align equiv.psum_sum Equiv.psumSum /-- Combine two `equiv`s using `sum` in the domain and `psum` in the codomain. -/ -def sumPsum {α₁ β₁ : Type _} {α₂ β₂ : Sort _} (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) : +def sumPSum {α₁ β₁ : Type _} {α₂ β₂ : Sort _} (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) : Sum α₁ β₁ ≃ PSum α₂ β₂ := (ea.symm.psumSum eb.symm).symm -#align equiv.sum_psum Equiv.sumPsum +#align equiv.sum_psum Equiv.sumPSum @[simp] theorem sumCongr_trans {α₁ α₂ β₁ β₂ γ₁ γ₂ : Sort _} (e : α₁ ≃ β₁) (f : α₂ ≃ β₂) (g : β₁ ≃ γ₁) @@ -316,10 +315,10 @@ theorem sumCongr_refl {α β : Sort _} : end Perm /-- `bool` is equivalent the sum of two `punit`s. -/ -def boolEquivPunitSumPunit : Bool ≃ Sum PUnit.{u + 1} PUnit.{v + 1} := +def boolEquivPUnitSumPUnit : Bool ≃ Sum PUnit.{u + 1} PUnit.{v + 1} := ⟨fun b => cond b (inr PUnit.unit) (inl PUnit.unit), Sum.elim (fun _ => false) fun _ => true, fun b => by cases b <;> rfl, fun s => by rcases s with (⟨⟨⟩⟩ | ⟨⟨⟩⟩) <;> rfl⟩ -#align equiv.bool_equiv_punit_sum_punit Equiv.boolEquivPunitSumPunit +#align equiv.bool_equiv_punit_sum_punit Equiv.boolEquivPUnitSumPUnit /-- Sum of types is commutative up to an equivalence. This is `sum.swap` as an equivalence. -/ @[simps (config := { fullyApplied := false }) apply] @@ -361,7 +360,7 @@ theorem sumAssoc_symm_apply_inl {α β γ} (a) : (sumAssoc α β γ).symm (inl a #align equiv.sum_assoc_symm_apply_inl Equiv.sumAssoc_symm_apply_inl @[simp] -theorem sum_assoc_symm_apply_inr_inl {α β γ} (b) : +theorem sumAssoc_symm_apply_inr_inl {α β γ} (b) : (sumAssoc α β γ).symm (inr (inl b)) = inl (inr b) := rfl #align equiv.sum_assoc_symm_apply_inr_inl Equiv.sumAssoc_symm_apply_inr_inl @@ -401,7 +400,7 @@ def optionEquivSumPUnit (α : Type _) : Option α ≃ Sum α PUnit.{u + 1} := ⟨fun o => o.elim (inr PUnit.unit) inl, fun s => s.elim some fun _ => none, fun o => by cases o <;> rfl, fun s => by rcases s with (_ | ⟨⟨⟩⟩) <;> rfl⟩ -#align equiv.option_equiv_sum_punit Equiv.optionEquivSumPunit +#align equiv.option_equiv_sum_punit Equiv.optionEquivSumPUnit @[simp] theorem optionEquivSumPUnit_none {α} : optionEquivSumPUnit α none = Sum.inr PUnit.unit := @@ -470,7 +469,7 @@ def sigmaFiberEquiv {α β : Type _} (f : α → β) : (Σy : β, { x // f x = y end -section SumCompl +section sumCompl /-- For any predicate `p` on `α`, the sum of the two subtypes `{a // p a}` and its complement `{a // ¬ p a}` @@ -592,9 +591,9 @@ theorem Perm.subtypeCongr.trans : #align equiv.perm.subtype_congr.trans Equiv.Perm.subtypeCongr.trans -end SumCompl +end sumCompl -section SubtypePreimage +section subtypePreimage variable (p : α → Prop) [DecidablePred p] (x₀ : { a // p a } → β) @@ -619,17 +618,17 @@ def subtypePreimage : { x : α → β // x ∘ Subtype.val = x₀ } ≃ ({ a // rw [dif_neg h] #align equiv.subtype_preimage Equiv.subtypePreimage -theorem subtype_preimage_symm_apply_coe_pos (x : { a // ¬p a } → β) (a : α) (h : p a) : +theorem subtypePreimage_symm_apply_coe_pos (x : { a // ¬p a } → β) (a : α) (h : p a) : ((subtypePreimage p x₀).symm x : α → β) a = x₀ ⟨a, h⟩ := dif_pos h -#align equiv.subtype_preimage_symm_apply_coe_pos Equiv.subtype_preimage_symm_apply_coe_pos +#align equiv.subtype_preimage_symm_apply_coe_pos Equiv.subtypePreimage_symm_apply_coe_pos -theorem subtype_preimage_symm_apply_coe_neg (x : { a // ¬p a } → β) (a : α) (h : ¬p a) : +theorem subtypePreimage_symm_apply_coe_neg (x : { a // ¬p a } → β) (a : α) (h : ¬p a) : ((subtypePreimage p x₀).symm x : α → β) a = x ⟨a, h⟩ := dif_neg h -#align equiv.subtype_preimage_symm_apply_coe_neg Equiv.subtype_preimage_symm_apply_coe_neg +#align equiv.subtype_preimage_symm_apply_coe_neg Equiv.subtypePreimage_symm_apply_coe_neg -end SubtypePreimage +end subtypePreimage section @@ -667,7 +666,7 @@ def piCurry {α} {β : α → Sort _} (γ : ∀ a, β a → Sort _) : end -section ProdCongr +section prodCongr variable {α₁ β₁ β₂ : Type _} (e : α₁ → β₁ ≃ β₂) @@ -685,15 +684,15 @@ def prodCongrLeft : β₁ × α₁ ≃ β₂ × α₁ where #align equiv.prod_congr_left Equiv.prodCongrLeft @[simp] -theorem prod_congr_left_apply (b : β₁) (a : α₁) : prodCongrLeft e (b, a) = (e a b, a) := +theorem prodCongr_left_apply (b : β₁) (a : α₁) : prodCongrLeft e (b, a) = (e a b, a) := rfl -#align equiv.prod_congr_left_apply Equiv.prod_congr_left_apply +#align equiv.prod_congr_left_apply Equiv.prodCongr_left_apply -theorem prod_congr_refl_right (e : β₁ ≃ β₂) : +theorem prodCongr_refl_right (e : β₁ ≃ β₂) : prodCongr e (Equiv.refl α₁) = prodCongrLeft fun _ => e := by ext ⟨a, b⟩ : 1 simp -#align equiv.prod_congr_refl_right Equiv.prod_congr_refl_right +#align equiv.prod_congr_refl_right Equiv.prodCongr_refl_right /-- A family of equivalences `Π (a : α₁), β₁ ≃ β₂` generates an equivalence between `α₁ × β₁` and `α₁ × β₂`. -/ @@ -709,29 +708,29 @@ def prodCongrRight : α₁ × β₁ ≃ α₁ × β₂ where #align equiv.prod_congr_right Equiv.prodCongrRight @[simp] -theorem prod_congr_right_apply (a : α₁) (b : β₁) : prodCongrRight e (a, b) = (a, e a b) := +theorem prodCongr_right_apply (a : α₁) (b : β₁) : prodCongrRight e (a, b) = (a, e a b) := rfl -#align equiv.prod_congr_right_apply Equiv.prod_congr_right_apply +#align equiv.prod_congr_right_apply Equiv.prodCongr_right_apply -theorem prod_congr_refl_left (e : β₁ ≃ β₂) : +theorem prodCongr_refl_left (e : β₁ ≃ β₂) : prodCongr (Equiv.refl α₁) e = prodCongrRight fun _ => e := by ext ⟨a, b⟩ : 1 simp -#align equiv.prod_congr_refl_left Equiv.prod_congr_refl_left +#align equiv.prod_congr_refl_left Equiv.prodCongr_refl_left @[simp] -theorem prod_congr_left_trans_prod_comm : +theorem prodCongr_left_trans_prod_comm : (prodCongrLeft e).trans (prodComm _ _) = (prodComm _ _).trans (prodCongrRight e) := by ext ⟨a, b⟩ : 1 simp -#align equiv.prod_congr_left_trans_prod_comm Equiv.prod_congr_left_trans_prod_comm +#align equiv.prod_congr_left_trans_prod_comm Equiv.prodCongr_left_trans_prod_comm @[simp] -theorem prod_congr_right_trans_prod_comm : +theorem prodCongr_right_trans_prod_comm : (prodCongrRight e).trans (prodComm _ _) = (prodComm _ _).trans (prodCongrLeft e) := by ext ⟨a, b⟩ : 1 simp -#align equiv.prod_congr_right_trans_prod_comm Equiv.prod_congr_right_trans_prod_comm +#align equiv.prod_congr_right_trans_prod_comm Equiv.prodCongr_right_trans_prod_comm theorem sigma_congr_right_sigma_equiv_prod : (sigmaCongrRight e).trans (sigmaEquivProd α₁ β₂) @@ -740,14 +739,14 @@ theorem sigma_congr_right_sigma_equiv_prod : simp #align equiv.sigma_congr_right_sigma_equiv_prod Equiv.sigma_congr_right_sigma_equiv_prod -theorem sigma_equiv_prod_sigma_congr_right : +theorem sigmaEquivProd_sigmaCongrRight : (sigmaEquivProd α₁ β₁).symm.trans (sigmaCongrRight e) = (prodCongrRight e).trans (sigmaEquivProd α₁ β₂).symm := by ext ⟨a, b⟩ : 1 simp -#align equiv.sigma_equiv_prod_sigma_congr_right Equiv.sigma_equiv_prod_sigma_congr_right +#align equiv.sigma_equiv_prod_sigma_congr_right Equiv.sigmaEquivProd_sigmaCongrRight --- See also `equiv.of_preimage_equiv`. +-- See also `Equiv.ofPreimageEquiv`. /-- A family of equivalences between fibers gives an equivalence between domains. -/ @[simps] def ofFiberEquiv {α β γ : Type _} {f : α → γ} {g : β → γ} @@ -755,12 +754,12 @@ def ofFiberEquiv {α β γ : Type _} {f : α → γ} {g : β → γ} (sigmaFiberEquiv f).symm.trans <| (Equiv.sigmaCongrRight e).trans (sigmaFiberEquiv g) #align equiv.of_fiber_equiv Equiv.ofFiberEquiv -theorem of_fiber_equiv_map {α β γ} {f : α → γ} {g : β → γ} +theorem ofFiberEquiv_map {α β γ} {f : α → γ} {g : β → γ} (e : ∀ c, { a // f a = c } ≃ { b // g b = c }) (a : α) : g (ofFiberEquiv e a) = f a := - (_ : { b // g b = _ }).Prop -#align equiv.of_fiber_equiv_map Equiv.of_fiber_equiv_map + (_ : { b // g b = _ }).property +#align equiv.of_fiber_equiv_map Equiv.ofFiberEquiv_map -/-- A variation on `equiv.prod_congr` where the equivalence in the second component can depend +/-- A variation on `Equiv.prodCongr` where the equivalence in the second component can depend on the first component. A typical example is a shear mapping, explaining the name of this declaration. -/ @[simps (config := { fullyApplied := false })] @@ -775,7 +774,7 @@ def prodShear {α₁ β₁ α₂ β₂ : Type _} (e₁ : α₁ ≃ α₂) (e₂ simp only [apply_symm_apply] #align equiv.prod_shear Equiv.prodShear -end ProdCongr +end prodCongr namespace Perm @@ -797,30 +796,29 @@ def prodExtendRight : Perm (α₁ × β₁) where #align equiv.perm.prod_extend_right Equiv.Perm.prodExtendRight @[simp] -theorem prod_extend_right_apply_eq (b : β₁) : prodExtendRight a e (a, b) = (a, e b) := +theorem prodExtendRight_apply_eq (b : β₁) : prodExtendRight a e (a, b) = (a, e b) := if_pos rfl -#align equiv.perm.prod_extend_right_apply_eq Equiv.Perm.prod_extend_right_apply_eq +#align equiv.perm.prod_extend_right_apply_eq Equiv.Perm.prodExtendRight_apply_eq -theorem prod_extend_right_apply_ne {a a' : α₁} (h : a' ≠ a) (b : β₁) : +theorem prodExtendRight_apply_ne {a a' : α₁} (h : a' ≠ a) (b : β₁) : prodExtendRight a e (a', b) = (a', b) := if_neg h -#align equiv.perm.prod_extend_right_apply_ne Equiv.Perm.prod_extend_right_apply_ne +#align equiv.perm.prod_extend_right_apply_ne Equiv.Perm.prodExtendRight_apply_ne -theorem eq_of_prod_extend_right_ne {e : Perm β₁} {a a' : α₁} {b : β₁} +theorem eq_of_prodExtendRight_ne {e : Perm β₁} {a a' : α₁} {b : β₁} (h : prodExtendRight a e (a', b) ≠ (a', b)) : a' = a := by contrapose! h - exact prod_extend_right_apply_ne _ h _ -#align equiv.perm.eq_of_prod_extend_right_ne Equiv.Perm.eq_of_prod_extend_right_ne + exact prodExtendRight_apply_ne _ h _ +#align equiv.perm.eq_of_prod_extend_right_ne Equiv.Perm.eq_of_prodExtendRight_ne @[simp] -theorem fst_prod_extend_right (ab : α₁ × β₁) : (prodExtendRight a e ab).fst = ab.fst := by - rw [prod_extend_right, coe_fn_mk] +theorem fst_prodExtendRight (ab : α₁ × β₁) : (prodExtendRight a e ab).fst = ab.fst := by + rw [prodExtendRight, coe_fn_mk] split_ifs with h · rw [h] · rfl - -#align equiv.perm.fst_prod_extend_right Equiv.Perm.fst_prod_extend_right +#align equiv.perm.fst_prod_extend_right Equiv.Perm.fst_prodExtendRight end Perm @@ -847,28 +845,28 @@ def sumArrowEquivProdArrow (α β γ : Type _) : (Sum α β → γ) ≃ (α → #align equiv.sum_arrow_equiv_prod_arrow Equiv.sumArrowEquivProdArrow @[simp] -theorem sum_arrow_equiv_prod_arrow_apply_fst {α β γ} (f : Sum α β → γ) (a : α) : +theorem sumArrowEquivProdArrow_apply_fst {α β γ} (f : Sum α β → γ) (a : α) : (sumArrowEquivProdArrow α β γ f).1 a = f (inl a) := rfl -#align equiv.sum_arrow_equiv_prod_arrow_apply_fst Equiv.sum_arrow_equiv_prod_arrow_apply_fst +#align equiv.sum_arrow_equiv_prod_arrow_apply_fst Equiv.sumArrowEquivProdArrow_apply_fst @[simp] -theorem sum_arrow_equiv_prod_arrow_apply_snd {α β γ} (f : Sum α β → γ) (b : β) : +theorem sumArrowEquivProdArrow_apply_snd {α β γ} (f : Sum α β → γ) (b : β) : (sumArrowEquivProdArrow α β γ f).2 b = f (inr b) := rfl -#align equiv.sum_arrow_equiv_prod_arrow_apply_snd Equiv.sum_arrow_equiv_prod_arrow_apply_snd +#align equiv.sum_arrow_equiv_prod_arrow_apply_snd Equiv.sumArrowEquivProdArrow_apply_snd @[simp] -theorem sum_arrow_equiv_prod_arrow_symm_apply_inl {α β γ} (f : α → γ) (g : β → γ) (a : α) : +theorem sumArrowEquivProdArrow_symm_apply_inl {α β γ} (f : α → γ) (g : β → γ) (a : α) : ((sumArrowEquivProdArrow α β γ).symm (f, g)) (inl a) = f a := rfl -#align equiv.sum_arrow_equiv_prod_arrow_symm_apply_inl Equiv.sum_arrow_equiv_prod_arrow_symm_apply_inl +#align equiv.sum_arrow_equiv_prod_arrow_symm_apply_inl Equiv.sumArrowEquivProdArrow_symm_apply_inl @[simp] -theorem sum_arrow_equiv_prod_arrow_symm_apply_inr {α β γ} (f : α → γ) (g : β → γ) (b : β) : +theorem sumArrowEquivProdArrow_symm_apply_inr {α β γ} (f : α → γ) (g : β → γ) (b : β) : ((sumArrowEquivProdArrow α β γ).symm (f, g)) (inr b) = g b := rfl -#align equiv.sum_arrow_equiv_prod_arrow_symm_apply_inr Equiv.sum_arrow_equiv_prod_arrow_symm_apply_inr +#align equiv.sum_arrow_equiv_prod_arrow_symm_apply_inr Equiv.sumArrowEquivProdArrow_symm_apply_inr /-- Type product is right distributive with respect to type sum up to an equivalence. -/ def sumProdDistrib (α β γ : Sort _) : Sum α β × γ ≃ Sum (α × γ) (β × γ) := @@ -878,28 +876,28 @@ def sumProdDistrib (α β γ : Sort _) : Sum α β × γ ≃ Sum (α × γ) (β #align equiv.sum_prod_distrib Equiv.sumProdDistrib @[simp] -theorem sum_prod_distrib_apply_left {α β γ} (a : α) (c : γ) : +theorem sumProdDistrib_apply_left {α β γ} (a : α) (c : γ) : sumProdDistrib α β γ (Sum.inl a, c) = Sum.inl (a, c) := rfl -#align equiv.sum_prod_distrib_apply_left Equiv.sum_prod_distrib_apply_left +#align equiv.sum_prod_distrib_apply_left Equiv.sumProdDistrib_apply_left @[simp] -theorem sum_prod_distrib_apply_right {α β γ} (b : β) (c : γ) : +theorem sumProdDistrib_apply_right {α β γ} (b : β) (c : γ) : sumProdDistrib α β γ (Sum.inr b, c) = Sum.inr (b, c) := rfl -#align equiv.sum_prod_distrib_apply_right Equiv.sum_prod_distrib_apply_right +#align equiv.sum_prod_distrib_apply_right Equiv.sumProdDistrib_apply_right @[simp] -theorem sum_prod_distrib_symm_apply_left {α β γ} (a : α × γ) : +theorem sumProdDistrib_symm_apply_left {α β γ} (a : α × γ) : (sumProdDistrib α β γ).symm (inl a) = (inl a.1, a.2) := rfl -#align equiv.sum_prod_distrib_symm_apply_left Equiv.sum_prod_distrib_symm_apply_left +#align equiv.sum_prod_distrib_symm_apply_left Equiv.sumProdDistrib_symm_apply_left @[simp] -theorem sum_prod_distrib_symm_apply_right {α β γ} (b : β × γ) : +theorem sumProdDistrib_symm_apply_right {α β γ} (b : β × γ) : (sumProdDistrib α β γ).symm (inr b) = (inr b.1, b.2) := rfl -#align equiv.sum_prod_distrib_symm_apply_right Equiv.sum_prod_distrib_symm_apply_right +#align equiv.sum_prod_distrib_symm_apply_right Equiv.sumProdDistrib_symm_apply_right /-- Type product is left distributive with respect to type sum up to an equivalence. -/ def prodSumDistrib (α β γ : Sort _) : α × Sum β γ ≃ Sum (α × β) (α × γ) := @@ -911,28 +909,28 @@ def prodSumDistrib (α β γ : Sort _) : α × Sum β γ ≃ Sum (α × β) (α #align equiv.prod_sum_distrib Equiv.prodSumDistrib @[simp] -theorem prod_sum_distrib_apply_left {α β γ} (a : α) (b : β) : +theorem prodSumDistrib_apply_left {α β γ} (a : α) (b : β) : prodSumDistrib α β γ (a, Sum.inl b) = Sum.inl (a, b) := rfl -#align equiv.prod_sum_distrib_apply_left Equiv.prod_sum_distrib_apply_left +#align equiv.prod_sum_distrib_apply_left Equiv.prodSumDistrib_apply_left @[simp] -theorem prod_sum_distrib_apply_right {α β γ} (a : α) (c : γ) : +theorem prodSumDistrib_apply_right {α β γ} (a : α) (c : γ) : prodSumDistrib α β γ (a, Sum.inr c) = Sum.inr (a, c) := rfl -#align equiv.prod_sum_distrib_apply_right Equiv.prod_sum_distrib_apply_right +#align equiv.prod_sum_distrib_apply_right Equiv.prodSumDistrib_apply_right @[simp] -theorem prod_sum_distrib_symm_apply_left {α β γ} (a : α × β) : +theorem prodSumDistrib_symm_apply_left {α β γ} (a : α × β) : (prodSumDistrib α β γ).symm (inl a) = (a.1, inl a.2) := rfl -#align equiv.prod_sum_distrib_symm_apply_left Equiv.prod_sum_distrib_symm_apply_left +#align equiv.prod_sum_distrib_symm_apply_left Equiv.prodSumDistrib_symm_apply_left @[simp] -theorem prod_sum_distrib_symm_apply_right {α β γ} (a : α × γ) : +theorem prodSumDistrib_symm_apply_right {α β γ} (a : α × γ) : (prodSumDistrib α β γ).symm (inr a) = (a.1, inr a.2) := rfl -#align equiv.prod_sum_distrib_symm_apply_right Equiv.prod_sum_distrib_symm_apply_right +#align equiv.prod_sum_distrib_symm_apply_right Equiv.prodSumDistrib_symm_apply_right /-- An indexed sum of disjoint sums of types is equivalent to the sum of the indexed sums. -/ @[simps] @@ -1027,13 +1025,13 @@ def uniqueCongr (e : α ≃ β) : Unique α ≃ Unique β where #align equiv.unique_congr Equiv.uniqueCongr /-- If `α` is equivalent to `β`, then `is_empty α` is equivalent to `is_empty β`. -/ -theorem is_empty_congr (e : α ≃ β) : IsEmpty α ↔ IsEmpty β := - ⟨fun h => @function.isEmpty _ _ h e.symm, fun h => @function.isEmpty _ _ h e⟩ -#align equiv.is_empty_congr Equiv.is_empty_congr +theorem isEmpty_congr (e : α ≃ β) : IsEmpty α ↔ IsEmpty β := + ⟨fun h => @Function.isEmpty _ _ h e.symm, fun h => @Function.isEmpty _ _ h e⟩ +#align equiv.is_empty_congr Equiv.isEmpty_congr -protected theorem is_empty (e : α ≃ β) [IsEmpty β] : IsEmpty α := - e.is_empty_congr.mpr ‹_› -#align equiv.is_empty Equiv.is_empty +protected theorem isEmpty (e : α ≃ β) [IsEmpty β] : IsEmpty α := + e.isEmpty_congr.mpr ‹_› +#align equiv.is_empty Equiv.isEmpty section @@ -1044,42 +1042,42 @@ at corresponding points, then `{a // p a}` is equivalent to `{b // q b}`. For the statement where `α = β`, that is, `e : perm α`, see `perm.subtype_perm`. -/ def subtypeEquiv {p : α → Prop} {q : β → Prop} (e : α ≃ β) (h : ∀ a, p a ↔ q (e a)) : { a : α // p a } ≃ { b : β // q b } where - toFun a := ⟨e a, (h _).mp a.Prop⟩ - invFun b := ⟨e.symm b, (h _).mpr ((e.apply_symm_apply b).symm ▸ b.Prop)⟩ + toFun a := ⟨e a, (h _).mp a.property⟩ + invFun b := ⟨e.symm b, (h _).mpr ((e.apply_symm_apply b).symm ▸ b.property)⟩ left_inv a := Subtype.ext <| by simp right_inv b := Subtype.ext <| by simp #align equiv.subtype_equiv Equiv.subtypeEquiv @[simp] -theorem subtype_equiv_refl {p : α → Prop} (h : ∀ a, p a ↔ p (Equiv.refl _ a) := fun a => Iff.rfl) : +theorem subtypeEquiv_refl {p : α → Prop} (h : ∀ a, p a ↔ p (Equiv.refl _ a) := fun a => Iff.rfl) : (Equiv.refl α).subtypeEquiv h = Equiv.refl { a : α // p a } := by ext rfl -#align equiv.subtype_equiv_refl Equiv.subtype_equiv_refl +#align equiv.subtype_equiv_refl Equiv.subtypeEquiv_refl @[simp] -theorem subtype_equiv_symm {p : α → Prop} {q : β → Prop} (e : α ≃ β) (h : ∀ a : α, p a ↔ q (e a)) : +theorem subtypeEquiv_symm {p : α → Prop} {q : β → Prop} (e : α ≃ β) (h : ∀ a : α, p a ↔ q (e a)) : (e.subtypeEquiv h).symm = e.symm.subtypeEquiv fun a => by convert (h <| e.symm a).symm exact (e.apply_symm_apply a).symm := rfl -#align equiv.subtype_equiv_symm Equiv.subtype_equiv_symm +#align equiv.subtype_equiv_symm Equiv.subtypeEquiv_symm @[simp] -theorem subtype_equiv_trans {p : α → Prop} {q : β → Prop} {r : γ → Prop} (e : α ≃ β) (f : β ≃ γ) +theorem subtypeEquiv_trans {p : α → Prop} {q : β → Prop} {r : γ → Prop} (e : α ≃ β) (f : β ≃ γ) (h : ∀ a : α, p a ↔ q (e a)) (h' : ∀ b : β, q b ↔ r (f b)) : (e.subtypeEquiv h).trans (f.subtypeEquiv h') = (e.trans f).subtypeEquiv fun a => (h a).trans (h' <| e a) := rfl -#align equiv.subtype_equiv_trans Equiv.subtype_equiv_trans +#align equiv.subtype_equiv_trans Equiv.subtypeEquiv_trans @[simp] -theorem subtype_equiv_apply {p : α → Prop} {q : β → Prop} +theorem subtypeEquiv_apply {p : α → Prop} {q : β → Prop} (e : α ≃ β) (h : ∀ a : α, p a ↔ q (e a)) (x : { x // p x }) : e.subtypeEquiv h x = ⟨e x, (h _).1 x.2⟩ := rfl -#align equiv.subtype_equiv_apply Equiv.subtype_equiv_apply +#align equiv.subtype_equiv_apply Equiv.subtypeEquiv_apply /-- If two predicates `p` and `q` are pointwise equivalent, then `{x // p x}` is equivalent to `{x // q x}`. -/ @@ -1103,7 +1101,7 @@ def subtypeEquivOfSubtype' {p : α → Prop} (e : α ≃ β) : /-- If two predicates are equal, then the corresponding subtypes are equivalent. -/ def subtypeEquivProp {α : Type _} {p q : α → Prop} (h : p = q) : Subtype p ≃ Subtype q := - subtypeEquiv (Equiv.refl α) fun a => h ▸ Iff.rfl + subtypeEquiv (Equiv.refl α) fun _ => h ▸ Iff.rfl #align equiv.subtype_equiv_prop Equiv.subtypeEquivProp /-- A subtype of a subtype is equivalent to the subtype of elements satisfying both predicates. This @@ -1122,7 +1120,7 @@ def subtypeSubtypeEquivSubtypeExists {α : Type u} (p : α → Prop) (q : Subtyp @[simps] def subtypeSubtypeEquivSubtypeInter {α : Type u} (p q : α → Prop) : { x : Subtype p // q x.1 } ≃ Subtype fun x => p x ∧ q x := - (subtypeSubtypeEquivSubtypeExists p _).trans <| subtype_equiv_right fun x => exists_prop + (subtypeSubtypeEquivSubtypeExists p _).trans <| subtypeEquivRight fun x => exists_prop #align equiv.subtype_subtype_equiv_subtype_inter Equiv.subtypeSubtypeEquivSubtypeInter /-- If the outer subtype has more restrictive predicate than the inner one, @@ -1130,28 +1128,28 @@ then we can drop the latter. -/ @[simps] def subtypeSubtypeEquivSubtype {α : Type u} {p q : α → Prop} (h : ∀ {x}, q x → p x) : { x : Subtype p // q x.1 } ≃ Subtype q := - (subtypeSubtypeEquivSubtypeInter p _).trans <| subtype_equiv_right fun x => and_iff_right_of_imp h + (subtypeSubtypeEquivSubtypeInter p _).trans <| subtypeEquivRight fun _ => and_iff_right_of_imp h #align equiv.subtype_subtype_equiv_subtype Equiv.subtypeSubtypeEquivSubtype /-- If a proposition holds for all elements, then the subtype is equivalent to the original type. -/ @[simps apply symmApply] def subtypeUnivEquiv {α : Type u} {p : α → Prop} (h : ∀ x, p x) : Subtype p ≃ α := - ⟨fun x => x, fun x => ⟨x, h x⟩, fun x => Subtype.eq rfl, fun x => rfl⟩ + ⟨fun x => x, fun x => ⟨x, h x⟩, fun _ => Subtype.eq rfl, fun _ => rfl⟩ #align equiv.subtype_univ_equiv Equiv.subtypeUnivEquiv /-- A subtype of a sigma-type is a sigma-type over a subtype. -/ def subtypeSigmaEquiv {α : Type u} (p : α → Type v) (q : α → Prop) : { y : Sigma p // q y.1 } ≃ Σx : Subtype q, p x.1 := - ⟨fun x => ⟨⟨x.1.1, x.2⟩, x.1.2⟩, fun x => ⟨⟨x.1.1, x.2⟩, x.1.2⟩, fun ⟨⟨x, h⟩, y⟩ => rfl, - fun ⟨⟨x, y⟩, h⟩ => rfl⟩ + ⟨fun x => ⟨⟨x.1.1, x.2⟩, x.1.2⟩, fun x => ⟨⟨x.1.1, x.2⟩, x.1.2⟩, fun _ => rfl, + fun _ => rfl⟩ #align equiv.subtype_sigma_equiv Equiv.subtypeSigmaEquiv /-- A sigma type over a subtype is equivalent to the sigma set over the original type, if the fiber is empty outside of the subset -/ def sigmaSubtypeEquivOfSubset {α : Type u} (p : α → Type v) (q : α → Prop) (h : ∀ x, p x → q x) : (Σx : Subtype q, p x) ≃ Σx : α, p x := - (subtypeSigmaEquiv p q).symm.trans <| subtype_univ_equiv fun x => h x.1 x.2 + (subtypeSigmaEquiv p q).symm.trans <| subtypeUnivEquiv fun x => h x.1 x.2 #align equiv.sigma_subtype_equiv_of_subset Equiv.sigmaSubtypeEquivOfSubset /-- If a predicate `p : β → Prop` is true on the range of a map `f : α → β`, then @@ -1159,7 +1157,7 @@ def sigmaSubtypeEquivOfSubset {α : Type u} (p : α → Type v) (q : α → Prop def sigmaSubtypeFiberEquiv {α : Type u} {β : Type v} (f : α → β) (p : β → Prop) (h : ∀ x, p (f x)) : (Σy : Subtype p, { x : α // f x = y }) ≃ α := calc - _ ≃ Σy : β, { x : α // f x = y } := sigmaSubtypeEquivOfSubset _ p fun y ⟨x, h'⟩ => h' ▸ h x + _ ≃ Σy : β, { x : α // f x = y } := sigmaSubtypeEquivOfSubset _ p fun _ ⟨x, h'⟩ => h' ▸ h x _ ≃ α := sigmaFiberEquiv f #align equiv.sigma_subtype_fiber_equiv Equiv.sigmaSubtypeFiberEquiv @@ -1170,15 +1168,14 @@ def sigmaSubtypeFiberEquivSubtype {α : Type u} {β : Type v} (f : α → β) {p {q : β → Prop} (h : ∀ x, p x ↔ q (f x)) : (Σy : Subtype q, { x : α // f x = y }) ≃ Subtype p := calc (Σy : Subtype q, { x : α // f x = y }) ≃ Σy : - Subtype q, { x : Subtype p // Subtype.mk (f x) ((h x).1 x.2) = y } := - by - apply sigma_congr_right - intro y - symm - refine' (subtype_subtype_equiv_subtype_exists _ _).trans (subtype_equiv_right _) - intro x - exact ⟨fun ⟨hp, h'⟩ => congr_arg Subtype.val h', fun h' => ⟨(h x).2 (h'.symm ▸ y.2), - Subtype.eq h'⟩⟩ + Subtype q, { x : Subtype p // Subtype.mk (f x) ((h x).1 x.2) = y } := by { + apply sigmaCongrRight + intro y + symm + refine' (subtype_subtype_equiv_subtype_exists _ _).trans (subtype_equiv_right _) + intro x + exact ⟨fun ⟨hp, h'⟩ => congr_arg Subtype.val h', fun h' => ⟨(h x).2 (h'.symm ▸ y.2), + Subtype.eq h'⟩⟩ } _ ≃ Subtype p := sigmaFiberEquiv fun x : Subtype p => (⟨f x, (h x).1 x.property⟩ : Subtype q) #align equiv.sigma_subtype_fiber_equiv_subtype Equiv.sigmaSubtypeFiberEquivSubtype @@ -1194,10 +1191,10 @@ def sigmaOptionEquivOfSome {α : Type u} (p : Option α → Type v) (h : p none exfalso exact h n - · intro s + · intro _ exact rfl - (sigma_subtype_equiv_of_subset _ _ h').symm.trans (sigma_congr_left' (option_is_some_equiv α)) + (sigmaSubtypeEquivOfSubset _ _ h').symm.trans (sigmaCongrLeft' (optionIsSomeEquiv α)) #align equiv.sigma_option_equiv_of_some Equiv.sigmaOptionEquivOfSome /-- The `pi`-type `Π i, π i` is equivalent to the type of sections `f : ι → Σ i, π i` of the @@ -1235,8 +1232,8 @@ def subtypeProdEquivProd {α : Type u} {β : Type v} {p : α → Prop} {q : β /-- A subtype of a `prod` is equivalent to a sigma type whose fibers are subtypes. -/ def subtypeProdEquivSigmaSubtype {α β : Type _} (p : α → β → Prop) : { x : α × β // p x.1 x.2 } ≃ Σa, { b : β // p a b } where - toFun x := ⟨x.1.1, x.1.2, x.Prop⟩ - invFun x := ⟨⟨x.1, x.2⟩, x.2.Prop⟩ + toFun x := ⟨x.1.1, x.1.2, x.property⟩ + invFun x := ⟨⟨x.1, x.2⟩, x.2.property⟩ left_inv x := by ext <;> rfl right_inv := fun ⟨a, b, pab⟩ => rfl #align equiv.subtype_prod_equiv_sigma_subtype Equiv.subtypeProdEquivSigmaSubtype @@ -1252,15 +1249,12 @@ def piEquivPiSubtypeProd {α : Type _} (p : α → Prop) (β : α → Type _) [D rintro ⟨f, g⟩ ext1 <;> · ext y - rcases y with ⟨⟩ - simp only [y_property, dif_pos, dif_neg, not_false_iff, Subtype.coe_mk] - rfl - + rcases y with ⟨val, property⟩ + simp only [property, dif_pos, dif_neg, not_false_iff, Subtype.coe_mk] left_inv f := by ext x by_cases h:p x <;> · simp only [h, dif_neg, dif_pos, not_false_iff] - rfl #align equiv.pi_equiv_pi_subtype_prod Equiv.piEquivPiSubtypeProd @@ -1294,6 +1288,8 @@ def funSplitAt {α : Type _} [DecidableEq α] (i : α) (β : Type _) : end +-- Porting note: this section requires some thought about what the statements should actually be! +-- Haven't touched any align statements in this section yet. section SubtypeEquivCodomain variable {X : Type _} {Y : Type _} [DecidableEq X] {x : X} @@ -1355,21 +1351,21 @@ end SubtypeEquivCodomain @[simps apply] noncomputable def ofBijective (f : α → β) (hf : Bijective f) : α ≃ β where toFun := f - invFun := Function.surjInv hf.Surjective - left_inv := Function.left_inverse_surj_inv hf - right_inv := Function.right_inverse_surj_inv _ + invFun := Function.surjInv hf.surjective + left_inv := Function.leftInverse_surjInv hf + right_inv := Function.rightInverse_surjInv _ #align equiv.of_bijective Equiv.ofBijective -theorem of_bijective_apply_symm_apply (f : α → β) (hf : Bijective f) (x : β) : +theorem ofBijective_apply_symm_apply (f : α → β) (hf : Bijective f) (x : β) : f ((ofBijective f hf).symm x) = x := (ofBijective f hf).apply_symm_apply x -#align equiv.of_bijective_apply_symm_apply Equiv.of_bijective_apply_symm_apply +#align equiv.of_bijective_apply_symm_apply Equiv.ofBijective_apply_symm_apply @[simp] -theorem of_bijective_symm_apply_apply (f : α → β) (hf : Bijective f) (x : α) : +theorem ofBijective_symm_apply_apply (f : α → β) (hf : Bijective f) (x : α) : (ofBijective f hf).symm (f x) = x := (ofBijective f hf).symm_apply_apply x -#align equiv.of_bijective_symm_apply_apply Equiv.of_bijective_symm_apply_apply +#align equiv.of_bijective_symm_apply_apply Equiv.ofBijective_symm_apply_apply instance : CanLift (α → β) (α ≃ β) coeFn Bijective where prf f hf := ⟨ofBijective f hf, rfl⟩ @@ -1389,33 +1385,33 @@ def Perm.extendDomain : Perm β' := #align equiv.perm.extend_domain Equiv.Perm.extendDomain @[simp] -theorem Perm.extend_domain_apply_image (a : α') : e.extendDomain f (f a) = f (e a) := by - simp [perm.extend_domain] -#align equiv.perm.extend_domain_apply_image Equiv.Perm.extend_domain_apply_image +theorem Perm.extendDomain_apply_image (a : α') : e.extendDomain f (f a) = f (e a) := by + simp [Perm.extendDomain] +#align equiv.perm.extend_domain_apply_image Equiv.Perm.extendDomain_apply_image -theorem Perm.extend_domain_apply_subtype {b : β'} (h : p b) : +theorem Perm.extendDomain_apply_subtype {b : β'} (h : p b) : e.extendDomain f b = f (e (f.symm ⟨b, h⟩)) := by - simp [perm.extend_domain, h] -#align equiv.perm.extend_domain_apply_subtype Equiv.Perm.extend_domain_apply_subtype + simp [Perm.extendDomain, h] +#align equiv.perm.extend_domain_apply_subtype Equiv.Perm.extendDomain_apply_subtype -theorem Perm.extend_domain_apply_not_subtype {b : β'} (h : ¬p b) : e.extendDomain f b = b := by - simp [perm.extend_domain, h] -#align equiv.perm.extend_domain_apply_not_subtype Equiv.Perm.extend_domain_apply_not_subtype +theorem Perm.extendDomain_apply_not_subtype {b : β'} (h : ¬p b) : e.extendDomain f b = b := by + simp [Perm.extendDomain, h] +#align equiv.perm.extend_domain_apply_not_subtype Equiv.Perm.extendDomain_apply_not_subtype @[simp] -theorem Perm.extend_domain_refl : Perm.extendDomain (Equiv.refl _) f = Equiv.refl _ := by - simp [perm.extend_domain] -#align equiv.perm.extend_domain_refl Equiv.Perm.extend_domain_refl +theorem Perm.extendDomain_refl : Perm.extendDomain (Equiv.refl _) f = Equiv.refl _ := by + simp [Perm.extendDomain] +#align equiv.perm.extend_domain_refl Equiv.Perm.extendDomain_refl @[simp] -theorem Perm.extend_domain_symm : (e.extendDomain f).symm = Perm.extendDomain e.symm f := +theorem Perm.extendDomain_symm : (e.extendDomain f).symm = Perm.extendDomain e.symm f := rfl -#align equiv.perm.extend_domain_symm Equiv.Perm.extend_domain_symm +#align equiv.perm.extend_domain_symm Equiv.Perm.extendDomain_symm -theorem Perm.extend_domain_trans (e e' : Perm α') : +theorem Perm.extendDomain_trans (e e' : Perm α') : (e.extendDomain f).trans (e'.extendDomain f) = Perm.extendDomain (e.trans e') f := by - simp [perm.extend_domain, perm_congr_trans] -#align equiv.perm.extend_domain_trans Equiv.Perm.extend_domain_trans + simp [Perm.extendDomain, permCongr_trans] +#align equiv.perm.extend_domain_trans Equiv.Perm.extendDomain_trans end @@ -1440,20 +1436,20 @@ def subtypeQuotientEquivQuotientSubtype (p₁ : α → Prop) [s₁ : Setoid α] #align equiv.subtype_quotient_equiv_quotient_subtype Equiv.subtypeQuotientEquivQuotientSubtype @[simp] -theorem subtype_quotient_equiv_quotient_subtype_mk (p₁ : α → Prop) +theorem subtypeQuotientEquivQuotientSubtype_mk (p₁ : α → Prop) [s₁ : Setoid α] [s₂ : Setoid (Subtype p₁)] (p₂ : Quotient s₁ → Prop) (hp₂ : ∀ a, p₁ a ↔ p₂ ⟦a⟧) (h : ∀ x y : Subtype p₁, @Setoid.r _ s₂ x y ↔ (x : α) ≈ y) (x hx) : subtypeQuotientEquivQuotientSubtype p₁ p₂ hp₂ h ⟨⟦x⟧, hx⟩ = ⟦⟨x, (hp₂ _).2 hx⟩⟧ := rfl -#align equiv.subtype_quotient_equiv_quotient_subtype_mk Equiv.subtype_quotient_equiv_quotient_subtype_mk +#align equiv.subtype_quotient_equiv_quotient_subtype_mk Equiv.subtypeQuotientEquivQuotientSubtype_mk @[simp] -theorem subtype_quotient_equiv_quotient_subtype_symm_mk (p₁ : α → Prop) +theorem subtypeQuotientEquivQuotientSubtype_symm_mk (p₁ : α → Prop) [s₁ : Setoid α] [s₂ : Setoid (Subtype p₁)] (p₂ : Quotient s₁ → Prop) (hp₂ : ∀ a, p₁ a ↔ p₂ ⟦a⟧) (h : ∀ x y : Subtype p₁, @Setoid.r _ s₂ x y ↔ (x : α) ≈ y) - (x) : (subtypeQuotientEquivQuotientSubtype p₁ p₂ hp₂ h).symm ⟦x⟧ = ⟨⟦x⟧, (hp₂ _).1 x.Prop⟩ := + (x) : (subtypeQuotientEquivQuotientSubtype p₁ p₂ hp₂ h).symm ⟦x⟧ = ⟨⟦x⟧, (hp₂ _).1 x.property⟩ := rfl -#align equiv.subtype_quotient_equiv_quotient_subtype_symm_mk Equiv.subtype_quotient_equiv_quotient_subtype_symm_mk +#align equiv.subtype_quotient_equiv_quotient_subtype_symm_mk Equiv.subtypeQuotientEquivQuotientSubtype_symm_mk section Swap @@ -1464,35 +1460,35 @@ def swapCore (a b r : α) : α := if r = a then b else if r = b then a else r #align equiv.swap_core Equiv.swapCore -theorem swap_core_self (r a : α) : swapCore a a r = r := by +theorem swapCore_self (r a : α) : swapCore a a r = r := by unfold swap_core split_ifs <;> cc -#align equiv.swap_core_self Equiv.swap_core_self +#align equiv.swap_core_self Equiv.swapCore_self -theorem swap_core_swap_core (r a b : α) : swapCore a b (swapCore a b r) = r := by +theorem swapCore_swapCore (r a b : α) : swapCore a b (swapCore a b r) = r := by unfold swap_core split_ifs <;> cc -#align equiv.swap_core_swap_core Equiv.swap_core_swap_core +#align equiv.swap_core_swap_core Equiv.swapCore_swapCore -theorem swap_core_comm (r a b : α) : swapCore a b r = swapCore b a r := by +theorem swapCore_comm (r a b : α) : swapCore a b r = swapCore b a r := by unfold swap_core split_ifs <;> cc -#align equiv.swap_core_comm Equiv.swap_core_comm +#align equiv.swap_core_comm Equiv.swapCore_comm /-- `swap a b` is the permutation that swaps `a` and `b` and leaves other values as is. -/ def swap (a b : α) : Perm α := - ⟨swapCore a b, swapCore a b, fun r => swap_core_swap_core r a b, - fun r => swap_core_swap_core r a b⟩ + ⟨swapCore a b, swapCore a b, fun r => swapCore_swapCore r a b, + fun r => swapCore_swapCore r a b⟩ #align equiv.swap Equiv.swap @[simp] theorem swap_self (a : α) : swap a a = Equiv.refl _ := - ext fun r => swap_core_self r a + ext fun r => swapCore_self r a #align equiv.swap_self Equiv.swap_self theorem swap_comm (a b : α) : swap a b = swap b a := - ext fun r => swap_core_comm r _ _ + ext fun r => swapCore_comm r _ _ #align equiv.swap_comm Equiv.swap_comm theorem swap_apply_def (a b x : α) : swap a b x = if x = a then b else if x = b then a else x := @@ -1515,7 +1511,7 @@ theorem swap_apply_of_ne_of_ne {a b x : α} : x ≠ a → x ≠ b → swap a b x @[simp] theorem swap_swap (a b : α) : (swap a b).trans (swap a b) = Equiv.refl _ := - ext fun x => swap_core_swap_core _ _ _ + ext fun x => swapCore_swapCore _ _ _ #align equiv.swap_swap Equiv.swap_swap @[simp] @@ -1525,7 +1521,7 @@ theorem symm_swap (a b : α) : (swap a b).symm = swap a b := @[simp] theorem swap_eq_refl_iff {x y : α} : swap x y = Equiv.refl _ ↔ x = y := by - refine' ⟨fun h => (Equiv.refl _).Injective _, fun h => h ▸ swap_self _⟩ + refine' ⟨fun h => (Equiv.refl _).injective _, fun h => h ▸ swap_self _⟩ rw [← h, swap_apply_left, h, refl_apply] #align equiv.swap_eq_refl_iff Equiv.swap_eq_refl_iff @@ -1629,10 +1625,10 @@ def setValue (f : α ≃ β) (a : α) (b : β) : α ≃ β := #align equiv.set_value Equiv.setValue @[simp] -theorem set_value_eq (f : α ≃ β) (a : α) (b : β) : setValue f a b a = b := by - dsimp [set_value] +theorem setValue_eq (f : α ≃ β) (a : α) (b : β) : setValue f a b a = b := by + dsimp [setValue] simp [swap_apply_left] -#align equiv.set_value_eq Equiv.set_value_eq +#align equiv.set_value_eq Equiv.setValue_eq end Swap @@ -1642,22 +1638,22 @@ namespace Function.Involutive /-- Convert an involutive function `f` to a permutation with `to_fun = inv_fun = f`. -/ def toPerm (f : α → α) (h : Involutive f) : Equiv.Perm α := - ⟨f, f, h.LeftInverse, h.RightInverse⟩ + ⟨f, f, h.leftInverse, h.rightInverse⟩ #align function.involutive.to_perm Function.Involutive.toPerm @[simp] -theorem coe_to_perm {f : α → α} (h : Involutive f) : (h.toPerm f : α → α) = f := +theorem coe_toPerm {f : α → α} (h : Involutive f) : (h.toPerm f : α → α) = f := rfl -#align function.involutive.coe_to_perm Function.Involutive.coe_to_perm +#align function.involutive.coe_to_perm Function.Involutive.coe_toPerm @[simp] -theorem to_perm_symm {f : α → α} (h : Involutive f) : (h.toPerm f).symm = h.toPerm f := +theorem toPerm_symm {f : α → α} (h : Involutive f) : (h.toPerm f).symm = h.toPerm f := rfl -#align function.involutive.to_perm_symm Function.Involutive.to_perm_symm +#align function.involutive.to_perm_symm Function.Involutive.toPerm_symm -theorem to_perm_involutive {f : α → α} (h : Involutive f) : Involutive (h.toPerm f) := +theorem toPerm_involutive {f : α → α} (h : Involutive f) : Involutive (h.toPerm f) := h -#align function.involutive.to_perm_involutive Function.Involutive.to_perm_involutive +#align function.involutive.to_perm_involutive Function.Involutive.toPerm_involutive end Function.Involutive @@ -1730,24 +1726,24 @@ def piCongr : (∀ a, W a) ≃ ∀ b, Z b := #align equiv.Pi_congr Equiv.piCongr @[simp] -theorem coe_Pi_congr_symm : ((h₁.piCongr h₂).symm : +theorem coe_piCongr_symm : ((h₁.piCongr h₂).symm : (∀ b, Z b) → ∀ a, W a) = fun f a => (h₂ a).symm (f (h₁ a)) := rfl -#align equiv.coe_Pi_congr_symm Equiv.coe_Pi_congr_symm +#align equiv.coe_Pi_congr_symm Equiv.coe_piCongr_symm -theorem Pi_congr_symm_apply (f : ∀ b, Z b) : +theorem piCongr_symm_apply (f : ∀ b, Z b) : (h₁.piCongr h₂).symm f = fun a => (h₂ a).symm (f (h₁ a)) := rfl -#align equiv.Pi_congr_symm_apply Equiv.Pi_congr_symm_apply +#align equiv.Pi_congr_symm_apply Equiv.piCongr_symm_apply @[simp] -theorem Pi_congr_apply_apply (f : ∀ a, W a) (a : α) : h₁.piCongr h₂ f (h₁ a) = h₂ a (f a) := by +theorem piCongr_apply_apply (f : ∀ a, W a) (a : α) : h₁.piCongr h₂ f (h₁ a) = h₂ a (f a) := by change cast _ ((h₂ (h₁.symm (h₁ a))) (f (h₁.symm (h₁ a)))) = (h₂ a) (f a) generalize_proofs hZa revert hZa rw [h₁.symm_apply_apply a] simp -#align equiv.Pi_congr_apply_apply Equiv.Pi_congr_apply_apply +#align equiv.Pi_congr_apply_apply Equiv.piCongr_apply_apply end @@ -1764,17 +1760,17 @@ def piCongr' : (∀ a, W a) ≃ ∀ b, Z b := #align equiv.Pi_congr' Equiv.piCongr' @[simp] -theorem coe_Pi_congr' : +theorem coe_piCongr' : (h₁.piCongr' h₂ : (∀ a, W a) → ∀ b, Z b) = fun f b => h₂ b <| f <| h₁.symm b := rfl -#align equiv.coe_Pi_congr' Equiv.coe_Pi_congr' +#align equiv.coe_Pi_congr' Equiv.coe_piCongr' -theorem Pi_congr'_apply (f : ∀ a, W a) : h₁.piCongr' h₂ f = fun b => h₂ b <| f <| h₁.symm b := +theorem piCongr'_apply (f : ∀ a, W a) : h₁.piCongr' h₂ f = fun b => h₂ b <| f <| h₁.symm b := rfl -#align equiv.Pi_congr'_apply Equiv.Pi_congr'_apply +#align equiv.Pi_congr'_apply Equiv.piCongr'_apply @[simp] -theorem Pi_congr'_symm_apply_symm_apply (f : ∀ b, Z b) (b : β) : +theorem piCongr'_symm_apply_symm_apply (f : ∀ b, Z b) (b : β) : (h₁.piCongr' h₂).symm f (h₁.symm b) = (h₂ b).symm (f b) := by change cast _ ((h₂ (h₁ (h₁.symm b))).symm (f (h₁ (h₁.symm b)))) = (h₂ b).symm (f b) generalize_proofs hWb @@ -1783,7 +1779,7 @@ theorem Pi_congr'_symm_apply_symm_apply (f : ∀ b, Z b) (b : β) : rw [h₁.apply_symm_apply b] at hb subst hb simp -#align equiv.Pi_congr'_symm_apply_symm_apply Equiv.Pi_congr'_symm_apply_symm_apply +#align equiv.Pi_congr'_symm_apply_symm_apply Equiv.piCongr'_symm_apply_symm_apply end @@ -1798,16 +1794,16 @@ theorem semiconj₂_conj : Semiconj₂ e f (e.arrowCongr e.conj f) := fun x y => #align equiv.semiconj₂_conj Equiv.semiconj₂_conj instance [IsAssociative α₁ f] : IsAssociative β₁ (e.arrowCongr (e.arrowCongr e) f) := - (e.semiconj₂_conj f).is_associative_right e.Surjective + (e.semiconj₂_conj f).isAssociative_right e.surjective instance [IsIdempotent α₁ f] : IsIdempotent β₁ (e.arrowCongr (e.arrowCongr e) f) := - (e.semiconj₂_conj f).is_idempotent_right e.Surjective + (e.semiconj₂_conj f).isIdempotent_right e.surjective instance [IsLeftCancel α₁ f] : IsLeftCancel β₁ (e.arrowCongr (e.arrowCongr e) f) := - ⟨e.Surjective.forall₃.2 fun x y z => by simpa using @IsLeftCancel.left_cancel _ f _ x y z⟩ + ⟨e.surjective.forall₃.2 fun x y z => by simpa using @IsLeftCancel.left_cancel _ f _ x y z⟩ instance [IsRightCancel α₁ f] : IsRightCancel β₁ (e.arrowCongr (e.arrowCongr e) f) := - ⟨e.Surjective.forall₃.2 fun x y z => by simpa using @IsRightCancel.right_cancel _ f _ x y z⟩ + ⟨e.surjective.forall₃.2 fun x y z => by simpa using @IsRightCancel.right_cancel _ f _ x y z⟩ end BinaryOp @@ -1835,8 +1831,8 @@ theorem Function.Injective.swap_comp def subsingletonProdSelfEquiv {α : Type _} [Subsingleton α] : α × α ≃ α where toFun p := p.1 invFun a := (a, a) - left_inv p := Subsingleton.elim _ _ - right_inv p := Subsingleton.elim _ _ + left_inv _ := Subsingleton.elim _ _ + right_inv _ := Subsingleton.elim _ _ #align subsingleton_prod_self_equiv subsingletonProdSelfEquiv /-- To give an equivalence between two subsingleton types, it is sufficient to give any two @@ -1874,7 +1870,7 @@ theorem update_apply_equiv_apply {α β α' : Sort _} [DecidableEq α'] [Decidab congr_fun (update_comp_equiv f g a v) a' #align function.update_apply_equiv_apply Function.update_apply_equiv_apply -theorem Pi_congr_left'_update [DecidableEq α] [DecidableEq β] (P : α → Sort _) (e : α ≃ β) +theorem piCongrLeft'_update [DecidableEq α] [DecidableEq β] (P : α → Sort _) (e : α ≃ β) (f : ∀ a, P a) (b : β) (x : P (e.symm b)) : e.piCongrLeft' P (update f (e.symm b) x) = update (e.piCongrLeft' P f) b x := by ext b' @@ -1883,12 +1879,12 @@ theorem Pi_congr_left'_update [DecidableEq α] [DecidableEq β] (P : α → Sort · simp [h] -#align function.Pi_congr_left'_update Function.Pi_congr_left'_update +#align function.Pi_congr_left'_update Function.piCongrLeft'_update -theorem Pi_congr_left'_symm_update [DecidableEq α] [DecidableEq β] (P : α → Sort _) (e : α ≃ β) +theorem piCongrLeft''_symm_update [DecidableEq α] [DecidableEq β] (P : α → Sort _) (e : α ≃ β) (f : ∀ b, P (e.symm b)) (b : β) (x : P (e.symm b)) : (e.piCongrLeft' P).symm (update f b x) = update ((e.piCongrLeft' P).symm f) (e.symm b) x := by - simp [(e.Pi_congr_left' P).symm_apply_eq, Pi_congr_left'_update] -#align function.Pi_congr_left'_symm_update Function.Pi_congr_left'_symm_update + simp [(e.piCongrLeft' P).symm_apply_eq, piCongrLeft'_update] +#align function.Pi_congr_left'_symm_update Function.piCongrLeft'_symm_update end Function From d1cb507e01782fb959b9886e21b766ed9b2f21dc Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Sat, 19 Nov 2022 13:08:38 +1100 Subject: [PATCH 043/127] unused arguments --- Mathlib/Logic/Equiv/Basic.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 86bcd4927f6b0..ad8031ea70fe2 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -1511,7 +1511,7 @@ theorem swap_apply_of_ne_of_ne {a b x : α} : x ≠ a → x ≠ b → swap a b x @[simp] theorem swap_swap (a b : α) : (swap a b).trans (swap a b) = Equiv.refl _ := - ext fun x => swapCore_swapCore _ _ _ + ext fun _ => swapCore_swapCore _ _ _ #align equiv.swap_swap Equiv.swap_swap @[simp] @@ -1824,7 +1824,7 @@ theorem Function.Injective.swap_apply theorem Function.Injective.swap_comp [DecidableEq α] [DecidableEq β] {f : α → β} (hf : Function.Injective f) (x y : α) : Equiv.swap (f x) (f y) ∘ f = f ∘ Equiv.swap x y := - funext fun z => hf.swap_apply _ _ _ + funext fun _ => hf.swap_apply _ _ _ #align function.injective.swap_comp Function.Injective.swap_comp /-- If `α` is a subsingleton, then it is equivalent to `α × α`. -/ From ef22eca15b302eb3fd627c7672eaffe41f3dc6b7 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Sat, 19 Nov 2022 13:16:18 +1100 Subject: [PATCH 044/127] fix a proof --- Mathlib/Logic/Equiv/Basic.lean | 10 ++++++++-- 1 file changed, 8 insertions(+), 2 deletions(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index ad8031ea70fe2..84e32da3aa77a 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -788,11 +788,17 @@ def prodExtendRight : Perm (α₁ × β₁) where left_inv := by rintro ⟨k', x⟩ dsimp only - split_ifs with h <;> simp [h] + split_ifs with h₁ h₂ + · simp [h₁] + · simp at h₂ + · simp right_inv := by rintro ⟨k', x⟩ dsimp only - split_ifs with h <;> simp [h] + split_ifs with h₁ h₂ + · simp [h₁] + · simp at h₂ + · simp #align equiv.perm.prod_extend_right Equiv.Perm.prodExtendRight @[simp] From df011cfe78fde42ba57ef5e6e03a17d4077f2cec Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Sat, 19 Nov 2022 13:17:06 +1100 Subject: [PATCH 045/127] add import to fix proof --- Mathlib/Logic/Equiv/Basic.lean | 1 + 1 file changed, 1 insertion(+) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 84e32da3aa77a..3558f0a3488e3 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -12,6 +12,7 @@ import Mathlib.Data.Sum.Basic import Mathlib.Logic.Function.Conjugate import Mathlib.Tactic.SplitIfs import Mathlib.Tactic.Convert +import Mathlib.Tactic.Contrapose -- **TODO** remove these later set_option autoImplicit false From 654677af94c133b2302add90de78fdbcc1a9a1a2 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Sat, 19 Nov 2022 13:18:22 +1100 Subject: [PATCH 046/127] add import to fix proof --- Mathlib/Logic/Equiv/Basic.lean | 9 +++++---- 1 file changed, 5 insertions(+), 4 deletions(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 3558f0a3488e3..fe31d06d05eac 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -5,6 +5,7 @@ Authors: Leonardo de Moura, Mario Carneiro Ported by: Kevin Buzzard, Ruben Vorster -/ import Mathlib.Logic.Equiv.Defs +import Mathlib.Data.Bool.Basic import Mathlib.Data.Prod.Basic import Mathlib.Data.Sigma.Basic import Mathlib.Data.Subtype @@ -820,10 +821,10 @@ theorem eq_of_prodExtendRight_ne {e : Perm β₁} {a a' : α₁} {b : β₁} @[simp] theorem fst_prodExtendRight (ab : α₁ × β₁) : (prodExtendRight a e ab).fst = ab.fst := by - rw [prodExtendRight, coe_fn_mk] + rw [prodExtendRight] + dsimp only split_ifs with h · rw [h] - · rfl #align equiv.perm.fst_prod_extend_right Equiv.Perm.fst_prodExtendRight @@ -982,8 +983,8 @@ def boolProdEquivSum (α : Type u) : Bool × α ≃ Sum α α where def boolArrowEquivProd (α : Type u) : (Bool → α) ≃ α × α where toFun f := (f true, f false) invFun p b := cond b p.1 p.2 - left_inv f := funext <| Bool.forall_bool.2 ⟨rfl, rfl⟩ - right_inv := fun ⟨x, y⟩ => rfl + left_inv _ := funext <| Bool.forall_bool.2 ⟨rfl, rfl⟩ + right_inv := fun _ => rfl #align equiv.bool_arrow_equiv_prod Equiv.boolArrowEquivProd end From 785081ad45d96c5dd9097b9289a5f3c6ebf5443e Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Sat, 19 Nov 2022 13:20:51 +1100 Subject: [PATCH 047/127] fix some proofs --- Mathlib/Logic/Equiv/Basic.lean | 11 ++++++----- 1 file changed, 6 insertions(+), 5 deletions(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index fe31d06d05eac..7cdb536980250 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -1009,7 +1009,7 @@ def natSumPunitEquivNat : Sum ℕ PUnit.{u + 1} ≃ ℕ := /-- The type of integer numbers is equivalent to `ℕ ⊕ ℕ`. -/ def intEquivNatSumNat : ℤ ≃ Sum ℕ ℕ where toFun z := Int.casesOn z inl inr - invFun := Sum.elim coe Int.negSucc + invFun := Sum.elim Int.ofNat Int.negSucc left_inv := by rintro (m | n) <;> rfl right_inv := by rintro (m | n) <;> rfl #align equiv.int_equiv_nat_sum_nat Equiv.intEquivNatSumNat @@ -1118,7 +1118,7 @@ version allows the “inner” predicate to depend on `h : p a`. -/ def subtypeSubtypeEquivSubtypeExists {α : Type u} (p : α → Prop) (q : Subtype p → Prop) : Subtype q ≃ { a : α // ∃ h : p a, q ⟨a, h⟩ } := ⟨fun a => - ⟨a, a.1.2, by + ⟨a.1, a.1.2, by rcases a with ⟨⟨a, hap⟩, haq⟩ exact haq⟩, fun a => ⟨⟨a, a.2.fst⟩, a.2.snd⟩, fun ⟨⟨a, ha⟩, h⟩ => rfl, fun ⟨a, h₁, h₂⟩ => rfl⟩ @@ -1128,7 +1128,8 @@ def subtypeSubtypeEquivSubtypeExists {α : Type u} (p : α → Prop) (q : Subtyp @[simps] def subtypeSubtypeEquivSubtypeInter {α : Type u} (p q : α → Prop) : { x : Subtype p // q x.1 } ≃ Subtype fun x => p x ∧ q x := - (subtypeSubtypeEquivSubtypeExists p _).trans <| subtypeEquivRight fun x => exists_prop + (subtypeSubtypeEquivSubtypeExists p _).trans <| + subtypeEquivRight fun x => @exists_prop (q x) (p x) #align equiv.subtype_subtype_equiv_subtype_inter Equiv.subtypeSubtypeEquivSubtypeInter /-- If the outer subtype has more restrictive predicate than the inner one, @@ -1179,8 +1180,8 @@ def sigmaSubtypeFiberEquivSubtype {α : Type u} {β : Type v} (f : α → β) {p Subtype q, { x : Subtype p // Subtype.mk (f x) ((h x).1 x.2) = y } := by { apply sigmaCongrRight intro y - symm - refine' (subtype_subtype_equiv_subtype_exists _ _).trans (subtype_equiv_right _) + apply Equiv.symm + refine' (subtypeSubtypeEquivSubtypeExists _ _).trans (subtypeEquivRight _) intro x exact ⟨fun ⟨hp, h'⟩ => congr_arg Subtype.val h', fun h' => ⟨(h x).2 (h'.symm ▸ y.2), Subtype.eq h'⟩⟩ } From 24fac8e41b35229f57c76e1e77b1b2abec713a35 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Sat, 19 Nov 2022 13:33:05 +1100 Subject: [PATCH 048/127] candidate translation of subtypeEquivCodomain --- Mathlib/Logic/Equiv/Basic.lean | 52 ++++++++++++++++++---------------- 1 file changed, 27 insertions(+), 25 deletions(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 7cdb536980250..4bbb541936593 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -1275,10 +1275,10 @@ def piSplitAt {α : Type _} [DecidableEq α] (i : α) (β : α → Type _) : toFun f := ⟨f i, fun j => f j⟩ invFun f j := if h : j = i then h.symm.rec f.1 else f.2 ⟨j, h⟩ right_inv f := by - ext + ext x exacts[dif_pos rfl, (dif_neg x.2).trans (by cases x <;> rfl)] left_inv f := by - ext + ext x dsimp only split_ifs · subst h @@ -1297,64 +1297,66 @@ def funSplitAt {α : Type _} [DecidableEq α] (i : α) (β : Type _) : end --- Porting note: this section requires some thought about what the statements should actually be! --- Haven't touched any align statements in this section yet. -section SubtypeEquivCodomain +section subtypeEquivCodomain variable {X : Type _} {Y : Type _} [DecidableEq X] {x : X} /-- The type of all functions `X → Y` with prescribed values for all `x' ≠ x` is equivalent to the codomain `Y`. -/ -def subtypeEquivCodomain (f : { x' // x' ≠ x } → Y) : { g : X → Y // g ∘ coe = f } ≃ Y := +def subtypeEquivCodomain (f : { x' // x' ≠ x } → Y) : + { g : X → Y // (fun x : { x' // x' ≠ x } => g x) = f } ≃ Y := (subtypePreimage _ f).trans <| @funUnique { x' // ¬x' ≠ x } _ <| show Unique { x' // ¬x' ≠ x } from @Equiv.unique _ _ (show Unique { x' // x' = x } from { - default := ⟨x, rfl⟩, uniq := fun ⟨x', h⟩ => Subtype.val_injective h }) - (subtype_equiv_right fun a => not_not) + default := ⟨x, rfl⟩, uniq := fun ⟨_, h⟩ => Subtype.val_injective h }) + (subtypeEquivRight fun _ => not_not) #align equiv.subtype_equiv_codomain Equiv.subtypeEquivCodomain @[simp] -theorem coe_subtype_equiv_codomain (f : { x' // x' ≠ x } → Y) : - (subtypeEquivCodomain f : { g : X → Y // g ∘ coe = f } → Y) = fun g => (g : X → Y) x := +theorem coe_subtypeEquivCodomain (f : { x' // x' ≠ x } → Y) : + (subtypeEquivCodomain f : _ → Y) = + fun g : { g : X → Y // (fun x : { x' // x' ≠ x } => g x) = f } => (g : X → Y) x := rfl -#align equiv.coe_subtype_equiv_codomain Equiv.coe_subtype_equiv_codomain +#align equiv.coe_subtype_equiv_codomain Equiv.coe_subtypeEquivCodomain @[simp] -theorem subtype_equiv_codomain_apply (f : { x' // x' ≠ x } → Y) (g : { g : X → Y // g ∘ coe = f }) : +theorem subtypeEquivCodomain_apply (f : { x' // x' ≠ x } → Y) (g) : subtypeEquivCodomain f g = (g : X → Y) x := rfl -#align equiv.subtype_equiv_codomain_apply Equiv.subtype_equiv_codomain_apply +#align equiv.subtype_equiv_codomain_apply Equiv.subtypeEquivCodomain_apply -theorem coe_subtype_equiv_codomain_symm (f : { x' // x' ≠ x } → Y) : - ((subtypeEquivCodomain f).symm : Y → { g : X → Y // g ∘ coe = f }) = fun y => +theorem coe_subtypeEquivCodomain_symm (f : { x' // x' ≠ x } → Y) : + ((subtypeEquivCodomain f).symm : Y → _) = fun y => ⟨fun x' => if h : x' ≠ x then f ⟨x', h⟩ else y, by funext x' - dsimp - erw [dif_pos x'.2, Subtype.coe_eta]⟩ := + simp only [ne_eq, Subtype.coe_eta, dite_eq_ite, ite_not, ite_eq_right_iff] + intro w + exfalso + exact x'.property w⟩ := rfl -#align equiv.coe_subtype_equiv_codomain_symm Equiv.coe_subtype_equiv_codomain_symm +#align equiv.coe_subtype_equiv_codomain_symm Equiv.coe_subtypeEquivCodomain_symm @[simp] -theorem subtype_equiv_codomain_symm_apply (f : { x' // x' ≠ x } → Y) (y : Y) (x' : X) : +theorem subtypeEquivCodomain_symm_apply (f : { x' // x' ≠ x } → Y) (y : Y) (x' : X) : ((subtypeEquivCodomain f).symm y : X → Y) x' = if h : x' ≠ x then f ⟨x', h⟩ else y := rfl -#align equiv.subtype_equiv_codomain_symm_apply Equiv.subtype_equiv_codomain_symm_apply +#align equiv.subtype_equiv_codomain_symm_apply Equiv.subtypeEquivCodomain_symm_apply @[simp] -theorem subtype_equiv_codomain_symm_apply_eq (f : { x' // x' ≠ x } → Y) (y : Y) : +theorem subtypeEquivCodomain_symm_apply_eq (f : { x' // x' ≠ x } → Y) (y : Y) : ((subtypeEquivCodomain f).symm y : X → Y) x = y := dif_neg (not_not.mpr rfl) -#align equiv.subtype_equiv_codomain_symm_apply_eq Equiv.subtype_equiv_codomain_symm_apply_eq +#align equiv.subtype_equiv_codomain_symm_apply_eq Equiv.subtypeEquivCodomain_symm_apply_eq -theorem subtype_equiv_codomain_symm_apply_ne +theorem subtypeEquivCodomain_symm_apply_ne (f : { x' // x' ≠ x } → Y) (y : Y) (x' : X) (h : x' ≠ x) : ((subtypeEquivCodomain f).symm y : X → Y) x' = f ⟨x', h⟩ := dif_pos h -#align equiv.subtype_equiv_codomain_symm_apply_ne Equiv.subtype_equiv_codomain_symm_apply_ne +#align equiv.subtype_equiv_codomain_symm_apply_ne Equiv.subtypeEquivCodomain_symm_apply_ne -end SubtypeEquivCodomain +end subtypeEquivCodomain /-- If `f` is a bijective function, then its domain is equivalent to its codomain. -/ @[simps apply] From 22987fd370402e3f1d0a25ff6dd6330b0ece6507 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Sat, 19 Nov 2022 13:48:22 +1100 Subject: [PATCH 049/127] work on swapCore lemmas --- Mathlib/Logic/Equiv/Basic.lean | 16 +++++++++++----- 1 file changed, 11 insertions(+), 5 deletions(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 4bbb541936593..233c03cf92323 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -14,6 +14,7 @@ import Mathlib.Logic.Function.Conjugate import Mathlib.Tactic.SplitIfs import Mathlib.Tactic.Convert import Mathlib.Tactic.Contrapose +import Mathlib.Tactic.CasesM -- **TODO** remove these later set_option autoImplicit false @@ -1472,18 +1473,23 @@ def swapCore (a b r : α) : α := #align equiv.swap_core Equiv.swapCore theorem swapCore_self (r a : α) : swapCore a a r = r := by - unfold swap_core - split_ifs <;> cc + unfold swapCore + split_ifs <;> simp [*] #align equiv.swap_core_self Equiv.swapCore_self theorem swapCore_swapCore (r a b : α) : swapCore a b (swapCore a b r) = r := by - unfold swap_core + unfold swapCore + -- Porting note: cc missing. + -- `casesm` would work here, with `casesm _ = _, ¬ _ = _`, + -- if it would just continue past failures on hypotheses matching the pattern split_ifs <;> cc #align equiv.swap_core_swap_core Equiv.swapCore_swapCore theorem swapCore_comm (r a b : α) : swapCore a b r = swapCore b a r := by - unfold swap_core - split_ifs <;> cc + unfold swapCore + -- Porting note: whatever solution works for `swapCore_swapCore` will work here too. + split_ifs with h₁ h₂ h₃ <;> simp + · cases h₁; cases h₂; rfl #align equiv.swap_core_comm Equiv.swapCore_comm /-- `swap a b` is the permutation that swaps `a` and `b` and From e096b32ed2f53dcbaa3c5eb0adcb977d0f04250f Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Sat, 19 Nov 2022 14:04:10 +1100 Subject: [PATCH 050/127] report issue --- Mathlib/Logic/Equiv/Basic.lean | 2 ++ 1 file changed, 2 insertions(+) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 233c03cf92323..aa2913698ab2a 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -1280,6 +1280,8 @@ def piSplitAt {α : Type _} [DecidableEq α] (i : α) (β : α → Type _) : exacts[dif_pos rfl, (dif_neg x.2).trans (by cases x <;> rfl)] left_inv f := by ext x + -- Porting note: this is broken, waiting on a Lean 4 bugfix: + -- https://github.com/leanprover/lean4/pull/1856 dsimp only split_ifs · subst h From 93b69f98835c92f72a850412d6b27062e9114467 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Sat, 19 Nov 2022 14:05:35 +1100 Subject: [PATCH 051/127] add note about subtypeEquivCodomain --- Mathlib/Logic/Equiv/Basic.lean | 2 ++ 1 file changed, 2 insertions(+) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index aa2913698ab2a..d3a23f82ebc79 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -1300,6 +1300,8 @@ def funSplitAt {α : Type _} [DecidableEq α] (i : α) (β : Type _) : end +-- Porting note (@semorrison): We need to redo these, I've changed the translation incorrectly. +-- Please see https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/logic.2Eequiv.2Ebasic/near/310956459 section subtypeEquivCodomain variable {X : Type _} {Y : Type _} [DecidableEq X] {x : X} From b0f817a46033e2be62cc06f8e9f96c53ba51e985 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Sat, 19 Nov 2022 18:07:01 +1100 Subject: [PATCH 052/127] fix subtypeEquivCodomain after gebner 's coe fix --- Mathlib/Logic/Equiv/Basic.lean | 8 +++----- 1 file changed, 3 insertions(+), 5 deletions(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index d3a23f82ebc79..b9bc61f61536a 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -1300,8 +1300,6 @@ def funSplitAt {α : Type _} [DecidableEq α] (i : α) (β : Type _) : end --- Porting note (@semorrison): We need to redo these, I've changed the translation incorrectly. --- Please see https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/logic.2Eequiv.2Ebasic/near/310956459 section subtypeEquivCodomain variable {X : Type _} {Y : Type _} [DecidableEq X] {x : X} @@ -1309,7 +1307,7 @@ variable {X : Type _} {Y : Type _} [DecidableEq X] {x : X} /-- The type of all functions `X → Y` with prescribed values for all `x' ≠ x` is equivalent to the codomain `Y`. -/ def subtypeEquivCodomain (f : { x' // x' ≠ x } → Y) : - { g : X → Y // (fun x : { x' // x' ≠ x } => g x) = f } ≃ Y := + { g : X → Y // g ∘ (↑) = f } ≃ Y := (subtypePreimage _ f).trans <| @funUnique { x' // ¬x' ≠ x } _ <| show Unique { x' // ¬x' ≠ x } from @@ -1322,7 +1320,7 @@ def subtypeEquivCodomain (f : { x' // x' ≠ x } → Y) : @[simp] theorem coe_subtypeEquivCodomain (f : { x' // x' ≠ x } → Y) : (subtypeEquivCodomain f : _ → Y) = - fun g : { g : X → Y // (fun x : { x' // x' ≠ x } => g x) = f } => (g : X → Y) x := + fun g : { g : X → Y // g ∘ (↑) = f } => (g : X → Y) x := rfl #align equiv.coe_subtype_equiv_codomain Equiv.coe_subtypeEquivCodomain @@ -1336,7 +1334,7 @@ theorem coe_subtypeEquivCodomain_symm (f : { x' // x' ≠ x } → Y) : ((subtypeEquivCodomain f).symm : Y → _) = fun y => ⟨fun x' => if h : x' ≠ x then f ⟨x', h⟩ else y, by funext x' - simp only [ne_eq, Subtype.coe_eta, dite_eq_ite, ite_not, ite_eq_right_iff] + simp only [ne_eq, dite_not, comp_apply, Subtype.coe_eta, dite_eq_ite, ite_eq_right_iff] intro w exfalso exact x'.property w⟩ := From 1f91792ec11c36945462ecf166aa5cd130f437ad Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Sat, 19 Nov 2022 18:31:21 +1100 Subject: [PATCH 053/127] add import --- Mathlib/Logic/Equiv/Basic.lean | 2 ++ 1 file changed, 2 insertions(+) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index b9bc61f61536a..a2f353b478468 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -15,6 +15,8 @@ import Mathlib.Tactic.SplitIfs import Mathlib.Tactic.Convert import Mathlib.Tactic.Contrapose import Mathlib.Tactic.CasesM +import Mathlib.Tactic.GeneralizeProofs +-- **TODO** review tactic imports to see what we actually used. -- **TODO** remove these later set_option autoImplicit false From b2a7cc3f26e6f588ff642778f18c63b4d7e4a71f Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Sat, 19 Nov 2022 19:25:52 +1100 Subject: [PATCH 054/127] piCongrLeft' --- Mathlib/Logic/Equiv/Basic.lean | 19 +++++++------------ 1 file changed, 7 insertions(+), 12 deletions(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index a2f353b478468..4e316f2eb0756 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -1704,19 +1704,14 @@ variable (P : α → Sort w) (e : α ≃ β) /-- Transport dependent functions through an equivalence of the base space. -/ @[simps] -def piCongrLeft' : (∀ a, P a) ≃ ∀ b, P (e.symm b) where +def piCongrLeft' (P : α → Sort _) (e : α ≃ β) : (∀ a, P a) ≃ ∀ b, P (e.symm b) where toFun f x := f (e.symm x) - invFun f x := by - rw [← e.symm_apply_apply x] - exact f (e x) - left_inv f := - funext fun x => - eq_of_heq - ((eq_rec_heq _ _).trans - (by - dsimp - rw [e.symm_apply_apply])) - right_inv f := funext fun x => eq_of_heq ((eq_rec_heq _ _).trans (by rw [e.apply_symm_apply])) + invFun f x := (e.symm_apply_apply x).ndrec (f (e x)) + left_inv f := funext fun x => + (by rintro _ rfl; rfl : ∀ {y} (h : y = x), h.ndrec (f y) = f x) (e.symm_apply_apply x) + right_inv f := funext fun x => + (by rintro _ rfl; rfl : ∀ {y} (h : y = x), (congr_arg e.symm h).ndrec (f y) = f x) + (e.apply_symm_apply x) #align equiv.Pi_congr_left' Equiv.piCongrLeft' end From faab506cfb3cb1bf3959f7b6f280d75ae6de5185 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Sat, 19 Nov 2022 19:32:46 +1100 Subject: [PATCH 055/127] fix some proofs --- Mathlib/Logic/Equiv/Basic.lean | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 4e316f2eb0756..1bf2b1a413553 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -1754,11 +1754,11 @@ theorem piCongr_symm_apply (f : ∀ b, Z b) : @[simp] theorem piCongr_apply_apply (f : ∀ a, W a) (a : α) : h₁.piCongr h₂ f (h₁ a) = h₂ a (f a) := by - change cast _ ((h₂ (h₁.symm (h₁ a))) (f (h₁.symm (h₁ a)))) = (h₂ a) (f a) + change Eq.ndrec _ _ = _ generalize_proofs hZa revert hZa rw [h₁.symm_apply_apply a] - simp + simp; rfl #align equiv.Pi_congr_apply_apply Equiv.piCongr_apply_apply end @@ -1788,13 +1788,13 @@ theorem piCongr'_apply (f : ∀ a, W a) : h₁.piCongr' h₂ f = fun b => h₂ b @[simp] theorem piCongr'_symm_apply_symm_apply (f : ∀ b, Z b) (b : β) : (h₁.piCongr' h₂).symm f (h₁.symm b) = (h₂ b).symm (f b) := by - change cast _ ((h₂ (h₁ (h₁.symm b))).symm (f (h₁ (h₁.symm b)))) = (h₂ b).symm (f b) + change Eq.ndrec _ _ = _ generalize_proofs hWb revert hWb generalize hb : h₁ (h₁.symm b) = b' rw [h₁.apply_symm_apply b] at hb subst hb - simp + simp; rfl #align equiv.Pi_congr'_symm_apply_symm_apply Equiv.piCongr'_symm_apply_symm_apply end From eab84842cea2295e2289988c1c9de9dfbd4bfd3d Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Sat, 19 Nov 2022 19:59:45 +1100 Subject: [PATCH 056/127] port missing init file for Sigma.eq --- Mathlib/Init/Data/Sigma/Basic.lean | 31 ++++++++++++++++++++++++++++++ Mathlib/Logic/Equiv/Basic.lean | 4 +++- 2 files changed, 34 insertions(+), 1 deletion(-) create mode 100644 Mathlib/Init/Data/Sigma/Basic.lean diff --git a/Mathlib/Init/Data/Sigma/Basic.lean b/Mathlib/Init/Data/Sigma/Basic.lean new file mode 100644 index 0000000000000..c6c6905401e4a --- /dev/null +++ b/Mathlib/Init/Data/Sigma/Basic.lean @@ -0,0 +1,31 @@ +/- +Copyright (c) 2014 Microsoft Corporation. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Author: Leonardo de Moura, Jeremy Avigad, Floris van Doorn +-/ +import Mathlib.Init.Logic + +universe u v + +theorem ex_of_psig {α : Type u} {p : α → Prop} : (Σ'x, p x) → ∃ x, p x + | ⟨x, hx⟩ => ⟨x, hx⟩ + +section + +variable {α : Type u} {β : α → Type v} + +protected theorem Sigma.eq : ∀ {p₁ p₂ : Σa : α, β a} (h₁ : p₁.1 = p₂.1), + (Eq.recOn h₁ p₁.2 : β p₂.1) = p₂.2 → p₁ = p₂ + | ⟨_, _⟩, _, rfl, rfl => rfl + +end + +section + +variable {α : Sort u} {β : α → Sort v} + +protected theorem PSigma.eq : ∀ {p₁ p₂ : PSigma β} (h₁ : p₁.1 = p₂.1), + (Eq.recOn h₁ p₁.2 : β p₂.1) = p₂.2 → p₁ = p₂ + | ⟨_, _⟩, _, rfl, rfl => rfl + +end diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 1bf2b1a413553..7edfd3160731b 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -16,6 +16,7 @@ import Mathlib.Tactic.Convert import Mathlib.Tactic.Contrapose import Mathlib.Tactic.CasesM import Mathlib.Tactic.GeneralizeProofs +import Mathlib.Init.Data.Sigma.Basic -- **TODO** review tactic imports to see what we actually used. -- **TODO** remove these later @@ -1383,7 +1384,8 @@ theorem ofBijective_symm_apply_apply (f : α → β) (hf : Bijective f) (x : α) (ofBijective f hf).symm_apply_apply x #align equiv.of_bijective_symm_apply_apply Equiv.ofBijective_symm_apply_apply -instance : CanLift (α → β) (α ≃ β) coeFn Bijective where prf f hf := ⟨ofBijective f hf, rfl⟩ +-- Porting note: `lift` tactic is not implemented yet. +-- instance : CanLift (α → β) (α ≃ β) coeFn Bijective where prf f hf := ⟨ofBijective f hf, rfl⟩ section From e31828372b33ed3105bb46e29551a85a2b314bc7 Mon Sep 17 00:00:00 2001 From: ruben vorster Date: Sat, 19 Nov 2022 14:02:12 +0000 Subject: [PATCH 057/127] fix semiconj2_conj --- Mathlib/Logic/Equiv/Basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 7edfd3160731b..f0d6858720280 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -1808,7 +1808,7 @@ variable {α₁ β₁ : Type _} (e : α₁ ≃ β₁) (f : α₁ → α₁ → theorem semiconj_conj (f : α₁ → α₁) : Semiconj e f (e.conj f) := fun x => by simp #align equiv.semiconj_conj Equiv.semiconj_conj -theorem semiconj₂_conj : Semiconj₂ e f (e.arrowCongr e.conj f) := fun x y => by simp +theorem semiconj₂_conj : Semiconj₂ e f (e.arrowCongr e.conj f) := fun x y => by simp [arrowCongr] #align equiv.semiconj₂_conj Equiv.semiconj₂_conj instance [IsAssociative α₁ f] : IsAssociative β₁ (e.arrowCongr (e.arrowCongr e) f) := From 8d808d51be0c8862b60b6c0e97cc2058d26330bc Mon Sep 17 00:00:00 2001 From: ruben vorster Date: Sat, 19 Nov 2022 16:33:38 +0000 Subject: [PATCH 058/127] Quotient.inductionOn, not Quotient.induction_on --- Mathlib/Logic/Equiv/Basic.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index f0d6858720280..2e83c2448be79 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -1449,8 +1449,8 @@ def subtypeQuotientEquivQuotientSubtype (p₁ : α → Prop) [s₁ : Setoid α] invFun a := Quotient.liftOn a (fun a => (⟨⟦a.1⟧, (hp₂ _).1 a.2⟩ : { x // p₂ x })) fun a b hab => Subtype.ext_val (Quotient.sound ((h _ _).1 hab)) - left_inv := fun ⟨a, ha⟩ => Quotient.induction_on a (fun a ha => rfl) ha - right_inv a := Quotient.induction_on a fun ⟨a, ha⟩ => rfl + left_inv := fun ⟨a, ha⟩ => Quotient.inductionOn a (fun a ha => rfl) ha + right_inv a := Quotient.inductionOn a fun ⟨a, ha⟩ => rfl #align equiv.subtype_quotient_equiv_quotient_subtype Equiv.subtypeQuotientEquivQuotientSubtype @[simp] From 90d6eba0036cf4809c73469b5484be6ce0d9a002 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Sun, 20 Nov 2022 10:48:36 +1100 Subject: [PATCH 059/127] cleanup after nightly-2022-11-19 --- Mathlib/Logic/Equiv/Basic.lean | 5 +---- 1 file changed, 1 insertion(+), 4 deletions(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 2e83c2448be79..f59ab9b0860d5 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -1283,12 +1283,9 @@ def piSplitAt {α : Type _} [DecidableEq α] (i : α) (β : α → Type _) : exacts[dif_pos rfl, (dif_neg x.2).trans (by cases x <;> rfl)] left_inv f := by ext x - -- Porting note: this is broken, waiting on a Lean 4 bugfix: - -- https://github.com/leanprover/lean4/pull/1856 dsimp only - split_ifs + split_ifs with h · subst h - · rfl #align equiv.pi_split_at Equiv.piSplitAt From 4400ecd41cffd0bd6e468bf3e2dc47b346141d20 Mon Sep 17 00:00:00 2001 From: Eric Rodriguez Date: Sun, 20 Nov 2022 01:18:09 +0000 Subject: [PATCH 060/127] mostly fixed --- Mathlib/Logic/Equiv/Basic.lean | 53 +++++++++++++++++++++++++--------- 1 file changed, 39 insertions(+), 14 deletions(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index f59ab9b0860d5..58e436e0123a8 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -749,7 +749,8 @@ theorem sigmaEquivProd_sigmaCongrRight : (sigmaEquivProd α₁ β₁).symm.trans (sigmaCongrRight e) = (prodCongrRight e).trans (sigmaEquivProd α₁ β₂).symm := by ext ⟨a, b⟩ : 1 - simp + simp only [trans_apply, sigmaCongrRight_apply, prodCongr_right_apply] + rfl #align equiv.sigma_equiv_prod_sigma_congr_right Equiv.sigmaEquivProd_sigmaCongrRight -- See also `Equiv.ofPreimageEquiv`. @@ -1210,16 +1211,17 @@ def sigmaOptionEquivOfSome {α : Type u} (p : Option α → Type v) (h : p none (sigmaSubtypeEquivOfSubset _ _ h').symm.trans (sigmaCongrLeft' (optionIsSomeEquiv α)) #align equiv.sigma_option_equiv_of_some Equiv.sigmaOptionEquivOfSome +-- ericr: this definition doesn't seem nice indentation-wise; is this just the Lean4 style? /-- The `pi`-type `Π i, π i` is equivalent to the type of sections `f : ι → Σ i, π i` of the `sigma` type such that for all `i` we have `(f i).fst = i`. -/ def piEquivSubtypeSigma (ι : Type _) (π : ι → Type _) : (∀ i, π i) ≃ { f : ι → Σi, π i // ∀ i, (f i).1 = i } := - ⟨fun f => ⟨fun i => ⟨i, f i⟩, fun i => rfl⟩, fun f i => by - rw [← f.2 i] - exact (f.1 i).2, fun f => funext fun i => rfl, fun ⟨f, hf⟩ => - Subtype.eq <| - funext fun i => - Sigma.eq (hf i).symm <| eq_of_heq <| rec_heq_of_heq _ <| rec_heq_of_heq _ <| HEq.refl _⟩ + ⟨fun f => ⟨fun i => ⟨i, f i⟩, fun i => rfl⟩, + fun f i => by rw [← f.2 i]; exact (f.1 i).2, + fun f => funext fun i => rfl, + fun ⟨f, hf⟩ => + Subtype.eq <| funext fun i => + Sigma.eq (hf i).symm <| eq_of_heq <| rec_heq_of_heq _ <| by simp⟩ #align equiv.pi_equiv_subtype_sigma Equiv.piEquivSubtypeSigma /-- The set of functions `f : Π a, β a` such that for all `a` we have `p a (f a)` is equivalent @@ -1280,12 +1282,12 @@ def piSplitAt {α : Type _} [DecidableEq α] (i : α) (β : α → Type _) : invFun f j := if h : j = i then h.symm.rec f.1 else f.2 ⟨j, h⟩ right_inv f := by ext x - exacts[dif_pos rfl, (dif_neg x.2).trans (by cases x <;> rfl)] + exacts[dif_pos rfl, (dif_neg x.2).trans (by cases x; rfl)] left_inv f := by ext x dsimp only split_ifs with h - · subst h + · subst h; rfl · rfl #align equiv.pi_split_at Equiv.piSplitAt @@ -1485,7 +1487,14 @@ theorem swapCore_swapCore (r a b : α) : swapCore a b (swapCore a b r) = r := by -- Porting note: cc missing. -- `casesm` would work here, with `casesm _ = _, ¬ _ = _`, -- if it would just continue past failures on hypotheses matching the pattern - split_ifs <;> cc + split_ifs with h₁ h₂ h₃ h₄ h₅ + · subst h₁; exact h₂ + · subst h₁; rfl + · cases h₃ rfl + · exact h₄.symm + · cases h₅ rfl + · cases h₅ rfl + · rfl #align equiv.swap_core_swap_core Equiv.swapCore_swapCore theorem swapCore_comm (r a b : α) : swapCore a b r = swapCore b a r := by @@ -1814,11 +1823,20 @@ instance [IsAssociative α₁ f] : IsAssociative β₁ (e.arrowCongr (e.arrowCon instance [IsIdempotent α₁ f] : IsIdempotent β₁ (e.arrowCongr (e.arrowCongr e) f) := (e.semiconj₂_conj f).isIdempotent_right e.surjective +-- porting note: the coe changes make `EmbeddingLike.apply_eq_iff_eq` almost impossible to use instance [IsLeftCancel α₁ f] : IsLeftCancel β₁ (e.arrowCongr (e.arrowCongr e) f) := - ⟨e.surjective.forall₃.2 fun x y z => by simpa using @IsLeftCancel.left_cancel _ f _ x y z⟩ + ⟨e.surjective.forall₃.2 fun x y z => by + simp only [arrowCongr, coe_fn_symm_mk, comp_apply, symm_apply_apply] + intro h + rw [@IsLeftCancel.left_cancel _ f _ x y z _] + exact e.injective h⟩ instance [IsRightCancel α₁ f] : IsRightCancel β₁ (e.arrowCongr (e.arrowCongr e) f) := - ⟨e.surjective.forall₃.2 fun x y z => by simpa using @IsRightCancel.right_cancel _ f _ x y z⟩ + ⟨e.surjective.forall₃.2 fun x y z => by + simp only [arrowCongr, coe_fn_symm_mk, comp_apply, symm_apply_apply] + intro h + rw [@IsRightCancel.right_cancel _ f _ x y z _] + exact e.injective h⟩ end BinaryOp @@ -1885,14 +1903,21 @@ theorem update_apply_equiv_apply {α β α' : Sort _} [DecidableEq α'] [Decidab congr_fun (update_comp_equiv f g a v) a' #align function.update_apply_equiv_apply Function.update_apply_equiv_apply +-- porting note: EmbeddingLike.apply_eq_iff_eq broken here too theorem piCongrLeft'_update [DecidableEq α] [DecidableEq β] (P : α → Sort _) (e : α ≃ β) (f : ∀ a, P a) (b : β) (x : P (e.symm b)) : e.piCongrLeft' P (update f (e.symm b) x) = update (e.piCongrLeft' P f) b x := by ext b' rcases eq_or_ne b' b with (rfl | h) · simp - - · simp [h] + · simp only [Equiv.piCongrLeft'_apply, ne_eq, h, not_false_iff, update_noteq] + rw [update_noteq _] + rw [ne_eq] + intro h' + /- an example of something that should work, or also putting `EmbeddingLike.apply_eq_iff_eq` + in the `simp` should too: + have := (EmbeddingLike.apply_eq_iff_eq e).mp h' -/ + cases e.symm.injective h' |> h #align function.Pi_congr_left'_update Function.piCongrLeft'_update From 9e8b9107b3206a0b10c50de7202026e82724edab Mon Sep 17 00:00:00 2001 From: Eric Rodriguez Date: Sun, 20 Nov 2022 01:47:28 +0000 Subject: [PATCH 061/127] fix big error --- Mathlib/Logic/Equiv/Basic.lean | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 58e436e0123a8..b2d95abf2435d 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -1448,7 +1448,8 @@ def subtypeQuotientEquivQuotientSubtype (p₁ : α → Prop) [s₁ : Setoid α] invFun a := Quotient.liftOn a (fun a => (⟨⟦a.1⟧, (hp₂ _).1 a.2⟩ : { x // p₂ x })) fun a b hab => Subtype.ext_val (Quotient.sound ((h _ _).1 hab)) - left_inv := fun ⟨a, ha⟩ => Quotient.inductionOn a (fun a ha => rfl) ha + -- for some reason, doing this as a `fun ⟨a, ha⟩ => ` block breaks things. + left_inv t := by rcases t with ⟨a, ha⟩; exact Quotient.inductionOn a (fun b hb => rfl) ha right_inv a := Quotient.inductionOn a fun ⟨a, ha⟩ => rfl #align equiv.subtype_quotient_equiv_quotient_subtype Equiv.subtypeQuotientEquivQuotientSubtype From a17bbf9f8d603c1803aeac2bcb798ca627e0ac26 Mon Sep 17 00:00:00 2001 From: Eric Rodriguez Date: Sun, 20 Nov 2022 01:51:03 +0000 Subject: [PATCH 062/127] some linting errors still but I don't understand them --- Mathlib/Logic/Equiv/Basic.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index b2d95abf2435d..8255d2ceb3d45 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -300,19 +300,19 @@ theorem sumCongr_apply {α β : Type _} (ea : Equiv.Perm α) (eb : Equiv.Perm β Equiv.sumCongr_apply ea eb x #align equiv.perm.sum_congr_apply Equiv.Perm.sumCongr_apply +-- porting note: it seems the general theorem about `Equiv` is now applied, so there's no need +-- to have this version also have `@[simp]`. Similarly for below. @[simp] theorem sumCongr_trans {α β : Sort _} (e : Equiv.Perm α) (f : Equiv.Perm β) (g : Equiv.Perm α) (h : Equiv.Perm β) : (sumCongr e f).trans (sumCongr g h) = sumCongr (e.trans g) (f.trans h) := Equiv.sumCongr_trans e f g h #align equiv.perm.sum_congr_trans Equiv.Perm.sumCongr_trans -@[simp] theorem sumCongr_symm {α β : Sort _} (e : Equiv.Perm α) (f : Equiv.Perm β) : (sumCongr e f).symm = sumCongr e.symm f.symm := Equiv.sumCongr_symm e f #align equiv.perm.sum_congr_symm Equiv.Perm.sumCongr_symm -@[simp] theorem sumCongr_refl {α β : Sort _} : sumCongr (Equiv.refl α) (Equiv.refl β) = Equiv.refl (Sum α β) := Equiv.sumCongr_refl @@ -1349,7 +1349,6 @@ theorem subtypeEquivCodomain_symm_apply (f : { x' // x' ≠ x } → Y) (y : Y) ( rfl #align equiv.subtype_equiv_codomain_symm_apply Equiv.subtypeEquivCodomain_symm_apply -@[simp] theorem subtypeEquivCodomain_symm_apply_eq (f : { x' // x' ≠ x } → Y) (y : Y) : ((subtypeEquivCodomain f).symm y : X → Y) x = y := dif_neg (not_not.mpr rfl) @@ -1929,3 +1928,4 @@ theorem piCongrLeft''_symm_update [DecidableEq α] [DecidableEq β] (P : α → #align function.Pi_congr_left'_symm_update Function.piCongrLeft'_symm_update end Function +#lint From 327b6d8a909cd95be21bb9216f2397778ea68aab Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Sun, 20 Nov 2022 13:42:42 +1100 Subject: [PATCH 063/127] cleanup imports --- Mathlib/Logic/Equiv/Basic.lean | 11 ++++------- 1 file changed, 4 insertions(+), 7 deletions(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 8255d2ceb3d45..7746f7bfc7b76 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -4,23 +4,18 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Mario Carneiro Ported by: Kevin Buzzard, Ruben Vorster -/ -import Mathlib.Logic.Equiv.Defs import Mathlib.Data.Bool.Basic import Mathlib.Data.Prod.Basic import Mathlib.Data.Sigma.Basic import Mathlib.Data.Subtype import Mathlib.Data.Sum.Basic +import Mathlib.Init.Data.Sigma.Basic +import Mathlib.Logic.Equiv.Defs import Mathlib.Logic.Function.Conjugate import Mathlib.Tactic.SplitIfs import Mathlib.Tactic.Convert import Mathlib.Tactic.Contrapose -import Mathlib.Tactic.CasesM import Mathlib.Tactic.GeneralizeProofs -import Mathlib.Init.Data.Sigma.Basic --- **TODO** review tactic imports to see what we actually used. - --- **TODO** remove these later -set_option autoImplicit false /-! # Equivalence between types @@ -49,6 +44,8 @@ In this file we continue the work on equivalences begun in `Logic/Equiv/Defs.lea equivalence, congruence, bijective map -/ +-- **TODO** remove these later +set_option autoImplicit false open Function From 0079d8372c5966f93ff5e2c0bb37b63b86c29d79 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Mon, 21 Nov 2022 08:38:41 +1100 Subject: [PATCH 064/127] linter --- Mathlib/Init/Data/Sigma/Basic.lean | 3 ++- Mathlib/Logic/Equiv/Basic.lean | 2 +- 2 files changed, 3 insertions(+), 2 deletions(-) diff --git a/Mathlib/Init/Data/Sigma/Basic.lean b/Mathlib/Init/Data/Sigma/Basic.lean index c6c6905401e4a..9d785d49db90b 100644 --- a/Mathlib/Init/Data/Sigma/Basic.lean +++ b/Mathlib/Init/Data/Sigma/Basic.lean @@ -1,7 +1,8 @@ /- Copyright (c) 2014 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. -Author: Leonardo de Moura, Jeremy Avigad, Floris van Doorn +Authors: Leonardo de Moura, Jeremy Avigad, Floris van Doorn +Ported by: Scott Morrison -/ import Mathlib.Init.Logic diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 7746f7bfc7b76..dd1e4bc7127c5 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -2,7 +2,7 @@ Copyright (c) 2015 Microsoft Corporation. All rights reserved. Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura, Mario Carneiro -Ported by: Kevin Buzzard, Ruben Vorster +Ported by: Kevin Buzzard, Ruben Vorster, Scott Morrison, Eric Rodriguez -/ import Mathlib.Data.Bool.Basic import Mathlib.Data.Prod.Basic From 7795164030c24f7c9596d54ced68c3162eaaf02a Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Mon, 21 Nov 2022 08:41:10 +1100 Subject: [PATCH 065/127] linter --- Mathlib/Init/Data/Sigma/Basic.lean | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/Mathlib/Init/Data/Sigma/Basic.lean b/Mathlib/Init/Data/Sigma/Basic.lean index 9d785d49db90b..9b66f8aeb6cc2 100644 --- a/Mathlib/Init/Data/Sigma/Basic.lean +++ b/Mathlib/Init/Data/Sigma/Basic.lean @@ -6,6 +6,10 @@ Ported by: Scott Morrison -/ import Mathlib.Init.Logic +/-! +# Lemmas about `Sigma` from Lean 3 core. +-/ + universe u v theorem ex_of_psig {α : Type u} {p : α → Prop} : (Σ'x, p x) → ∃ x, p x From 67a714119d73a92ee47e1c8178280ed54441362b Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Mon, 21 Nov 2022 08:42:13 +1100 Subject: [PATCH 066/127] long lines --- Mathlib/Logic/Equiv/Basic.lean | 7 ++++--- 1 file changed, 4 insertions(+), 3 deletions(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index dd1e4bc7127c5..2c276b0727f1a 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -1460,10 +1460,11 @@ theorem subtypeQuotientEquivQuotientSubtype_mk (p₁ : α → Prop) @[simp] theorem subtypeQuotientEquivQuotientSubtype_symm_mk (p₁ : α → Prop) [s₁ : Setoid α] [s₂ : Setoid (Subtype p₁)] (p₂ : Quotient s₁ → Prop) (hp₂ : ∀ a, p₁ a ↔ p₂ ⟦a⟧) - (h : ∀ x y : Subtype p₁, @Setoid.r _ s₂ x y ↔ (x : α) ≈ y) - (x) : (subtypeQuotientEquivQuotientSubtype p₁ p₂ hp₂ h).symm ⟦x⟧ = ⟨⟦x⟧, (hp₂ _).1 x.property⟩ := + (h : ∀ x y : Subtype p₁, @Setoid.r _ s₂ x y ↔ (x : α) ≈ y) (x) : + (subtypeQuotientEquivQuotientSubtype p₁ p₂ hp₂ h).symm ⟦x⟧ = ⟨⟦x⟧, (hp₂ _).1 x.property⟩ := rfl -#align equiv.subtype_quotient_equiv_quotient_subtype_symm_mk Equiv.subtypeQuotientEquivQuotientSubtype_symm_mk +#align equiv.subtype_quotient_equiv_quotient_subtype_symm_mk + Equiv.subtypeQuotientEquivQuotientSubtype_symm_mk section Swap From 8606881feb5a5e70c7bf8fb917dfd0bda4e80484 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Mon, 21 Nov 2022 08:49:17 +1100 Subject: [PATCH 067/127] import all files --- Mathlib.lean | 1 + 1 file changed, 1 insertion(+) diff --git a/Mathlib.lean b/Mathlib.lean index b6db869e7e45c..12aa78a56a365 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -93,6 +93,7 @@ import Mathlib.Init.Data.Ordering.Basic import Mathlib.Init.Data.Prod import Mathlib.Init.Data.Quot import Mathlib.Init.Data.Rat +import Mathlib.Init.Data.Sigma.Basic import Mathlib.Init.Function import Mathlib.Init.Logic import Mathlib.Init.Propext From ad8510d8d7d2e055256577fc1660e303ac088917 Mon Sep 17 00:00:00 2001 From: Arien Malec Date: Mon, 21 Nov 2022 18:48:54 -0800 Subject: [PATCH 068/127] feat: port Logic.Equiv.Option --- Mathlib/Logic/Equiv/Option.lean | 244 ++++++++++++++++++++++++++++++++ 1 file changed, 244 insertions(+) create mode 100644 Mathlib/Logic/Equiv/Option.lean diff --git a/Mathlib/Logic/Equiv/Option.lean b/Mathlib/Logic/Equiv/Option.lean new file mode 100644 index 0000000000000..0da1a4ee8f575 --- /dev/null +++ b/Mathlib/Logic/Equiv/Option.lean @@ -0,0 +1,244 @@ +/- +Copyright (c) 2021 Eric Wieser. All rights reserved. +Released under Apache 2.0 license as described in the file LICENSE. +Authors: Eric Wieser +-/ +import Mathlib.Control.EquivFunctor +import Mathlib.Data.Option.Basic +import Mathlib.Data.Subtype +import Mathlib.Logic.Equiv.Defs + +/-! +# Equivalences for `Option α` + + +We define +* `Equiv.option_congr`: the `Option α ≃ Option β` constructed from `e : α ≃ β` by sending `none` to + `none`, and applying a `e` elsewhere. +* `Equiv.remove_none`: the `α ≃ β` constructed from `Option α ≃ Option β` by removing `none` from + both sides. +-/ + + +namespace Equiv + +open Option + +variable {α β γ : Type _} + +section OptionCongr + +/-- A universe-polymorphic version of `EquivFunctor.map_equiv Option e`. -/ +@[simps apply] +def optionCongr (e : α ≃ β) : Option α ≃ Option β where + toFun := Option.map e + invFun := Option.map e.symm + left_inv x := (Option.map_map _ _ _).trans <| e.symm_comp_self.symm ▸ congr_fun Option.map_id x + right_inv x := (Option.map_map _ _ _).trans <| e.self_comp_symm.symm ▸ congr_fun Option.map_id x +#align equiv.option_congr Equiv.optionCongr + +@[simp] +theorem option_congr_refl : optionCongr (Equiv.refl α) = Equiv.refl _ := + ext <| congr_fun Option.map_id +#align equiv.option_congr_refl Equiv.option_congr_refl + +@[simp] +theorem option_congr_symm (e : α ≃ β) : (optionCongr e).symm = optionCongr e.symm := + rfl +#align equiv.option_congr_symm Equiv.option_congr_symm + +@[simp] +theorem option_congr_trans (e₁ : α ≃ β) (e₂ : β ≃ γ) : + (optionCongr e₁).trans (optionCongr e₂) = optionCongr (e₁.trans e₂) := + ext <| Option.map_map _ _ +#align equiv.option_congr_trans Equiv.option_congr_trans + +/-- When `α` and `β` are in the same universe, this is the same as the result of +`EquivFunctor.map_equiv`. -/ +theorem option_congr_eq_equiv_function_map_equiv {α β : Type _} (e : α ≃ β) : + optionCongr e = EquivFunctor.mapEquiv Option e := + rfl +#align equiv.option_congr_eq_equiv_function_map_equiv Equiv.option_congr_eq_equiv_function_map_equiv + +end OptionCongr + +section RemoveNone + +variable (e : Option α ≃ Option β) + +def remove_none_aux (x : α) : β := + if h : (e (some x)).isSome then Option.get _ h + else + Option.get _ <| + show (e none).isSome by + rw [← Option.ne_none_iff_isSome] + intro hn + rw [Option.not_isSome_iff_eq_none, ← hn] at h + exact Option.some_ne_none _ (e.injective h) + +#align equiv.remove_none_aux Equiv.remove_none_aux + +theorem remove_none_aux_some {x : α} (h : ∃ x', e (some x) = some x') : + some (remove_none_aux e x) = e (some x) := + by simp [remove_none_aux, Option.isSome_iff_exists.mpr h] +#align equiv.remove_none_aux_some Equiv.remove_none_aux_some + +theorem remove_none_aux_none {x : α} (h : e (some x) = none) : + some (remove_none_aux e x) = e none := by + simp [remove_none_aux, Option.not_isSome_iff_eq_none.mpr h] +#align equiv.remove_none_aux_none Equiv.remove_none_aux_none + +theorem remove_none_aux_inv (x : α) : remove_none_aux e.symm (remove_none_aux e x) = x := + Option.some_injective _ + (by + cases h1 : e.symm (some (remove_none_aux e x)) <;> cases h2 : e (some x) + · rw [remove_none_aux_none _ h1] + exact (e.eq_symm_apply.mpr h2).symm + + · rw [remove_none_aux_some _ ⟨_, h2⟩] at h1 + simp at h1 + + · rw [remove_none_aux_none _ h2] at h1 + simp at h1 + + · rw [remove_none_aux_some _ ⟨_, h1⟩] + rw [remove_none_aux_some _ ⟨_, h2⟩] + simp + ) +#align equiv.remove_none_aux_inv Equiv.remove_none_aux_inv + +/-- Given an equivalence between two `Option` types, eliminate `none` from that equivalence by +mapping `e.symm none` to `e none`. -/ +def removeNone : α ≃ β where + toFun := remove_none_aux e + invFun := remove_none_aux e.symm + left_inv := remove_none_aux_inv e + right_inv := remove_none_aux_inv e.symm +#align equiv.remove_none Equiv.removeNone + +@[simp] +theorem remove_none_symm : (removeNone e).symm = removeNone e.symm := + rfl +#align equiv.remove_none_symm Equiv.remove_none_symm + +theorem remove_none_some {x : α} (h : ∃ x', e (some x) = some x') : some (removeNone e x) = e (some x) := + remove_none_aux_some e h +#align equiv.remove_none_some Equiv.remove_none_some + +theorem remove_none_none {x : α} (h : e (some x) = none) : some (removeNone e x) = e none := + remove_none_aux_none e h +#align equiv.remove_none_none Equiv.remove_none_none + +@[simp] +theorem option_symm_apply_none_iff : e.symm none = none ↔ e none = none := + ⟨fun h => by simpa using (congr_arg e h).symm, fun h => by simpa using (congr_arg e.symm h).symm⟩ +#align equiv.option_symm_apply_none_iff Equiv.option_symm_apply_none_iff + +theorem some_remove_none_iff {x : α} : some (removeNone e x) = e none ↔ e.symm none = some x := by + cases' h : e (some x) with a + · rw [remove_none_none _ h] + simpa using (congr_arg e.symm h).symm + + · rw [remove_none_some _ ⟨a, h⟩] + have h1 := congr_arg e.symm h + rw [symm_apply_apply] at h1 + simp only [false_iff_iff, apply_eq_iff_eq] + simp [h1] + + +#align equiv.some_remove_none_iff Equiv.some_remove_none_iff + +@[simp] +theorem remove_none_option_congr (e : α ≃ β) : removeNone e.optionCongr = e := + Equiv.ext fun x => Option.some_injective _ <| remove_none_some _ ⟨e x, by simp [EquivFunctor.map]⟩ +#align equiv.remove_none_option_congr Equiv.remove_none_option_congr + +end RemoveNone + +theorem option_congr_injective : Function.Injective (optionCongr : α ≃ β → Option α ≃ Option β) := + Function.LeftInverse.injective remove_none_option_congr +#align equiv.option_congr_injective Equiv.option_congr_injective + +/-- Equivalences between `Option α` and `β` that send `none` to `x` are equivalent to +equivalences between `α` and `{y : β // y ≠ x}`. -/ +def option_subtype [DecidableEq β] (x : β) : { e : Option α ≃ β // e none = x } ≃ (α ≃ { y : β // y ≠ x }) where + toFun e := + { toFun := fun a => ⟨e a, ((EquivLike.injective _).ne_iff' e.property).2 (some_ne_none _)⟩, + invFun := fun b => + get + (ne_none_iff_is_some.1 + (((EquivLike.injective _).ne_iff' ((apply_eq_iff_eq_symm_apply _).1 e.property).symm).2 b.property)), + left_inv := fun a => by + rw [← some_inj, some_get, ← coe_def] + exact symm_apply_apply (e : Option α ≃ β) a, + right_inv := fun b => by + ext + simp + exact apply_symm_apply _ _ } + invFun e := + ⟨{ toFun := fun a => casesOn' a x (coe ∘ e), invFun := fun b => if h : b = x then none else e.symm ⟨b, h⟩, + left_inv := fun a => by + cases a + · simp + + simp only [cases_on'_some, Function.comp_apply, Subtype.coe_eta, symm_apply_apply, dite_eq_ite] + exact if_neg (e a).property, + right_inv := fun b => by by_cases h : b = x <;> simp [h] }, + rfl⟩ + left_inv e := by + ext a + cases a + · simpa using e.property.symm + + · simpa + + right_inv e := by + ext a + rfl +#align equiv.option_subtype Equiv.optionSubtype + +@[simp] +theorem option_subtype_apply_apply [DecidableEq β] (x : β) (e : { e : Option α ≃ β // e none = x }) (a : α) (h) : + option_subtype x e a = ⟨(e : Option α ≃ β) a, h⟩ := + rfl +#align equiv.option_subtype_apply_apply Equiv.option_subtype_apply_apply + +@[simp] +theorem coe_option_subtype_apply_apply [DecidableEq β] (x : β) (e : { e : Option α ≃ β // e none = x }) (a : α) : + ↑(option_subtype x e a) = (e : Option α ≃ β) a := + rfl +#align equiv.coe_option_subtype_apply_apply Equiv.coe_option_subtype_apply_apply + +@[simp] +theorem option_subtype_apply_symm_apply [DecidableEq β] (x : β) (e : { e : Option α ≃ β // e none = x }) + (b : { y : β // y ≠ x }) : ↑((option_subtype x e).symm b) = (e : Option α ≃ β).symm b := by + dsimp only [option_subtype] + simp +#align equiv.option_subtype_apply_symm_apply Equiv.option_subtype_apply_symm_apply + +@[simp] +theorem option_subtype_symm_apply_apply_coe [DecidableEq β] (x : β) (e : α ≃ { y : β // y ≠ x }) (a : α) : + (option_subtype x).symm e a = e a := + rfl +#align equiv.option_subtype_symm_apply_apply_coe Equiv.option_subtype_symm_apply_apply_coe + +@[simp] +theorem option_subtype_symm_apply_apply_some [DecidableEq β] (x : β) (e : α ≃ { y : β // y ≠ x }) (a : α) : + (option_subtype x).symm e (some a) = e a := + rfl +#align equiv.option_subtype_symm_apply_apply_some Equiv.option_subtype_symm_apply_apply_some + +@[simp] +theorem option_subtype_symm_apply_apply_none [DecidableEq β] (x : β) (e : α ≃ { y : β // y ≠ x }) : + (option_subtype x).symm e none = x := + rfl +#align equiv.option_subtype_symm_apply_apply_none Equiv.option_subtype_symm_apply_apply_none + +@[simp] +theorem option_subtype_symm_apply_symm_apply [DecidableEq β] (x : β) (e : α ≃ { y : β // y ≠ x }) + (b : { y : β // y ≠ x }) : ((option_subtype x).symm e : Option α ≃ β).symm b = e.symm b := by + simp only [option_subtype, coe_fn_symm_mk, Subtype.coe_mk, Subtype.coe_eta, dite_eq_ite, ite_eq_right_iff] + exact fun h => False.elim (b.property h) +#align equiv.option_subtype_symm_apply_symm_apply Equiv.option_subtype_symm_apply_symm_apply + +end Equiv From 2aae5fd989f2258e91aae8e8d0a388b44981b8e0 Mon Sep 17 00:00:00 2001 From: Arien Malec Date: Mon, 21 Nov 2022 21:19:54 -0800 Subject: [PATCH 069/127] Add to Mathlib.lean --- Mathlib.lean | 1 + 1 file changed, 1 insertion(+) diff --git a/Mathlib.lean b/Mathlib.lean index 45feb656393a2..582462f27a876 100644 --- a/Mathlib.lean +++ b/Mathlib.lean @@ -111,6 +111,7 @@ import Mathlib.Logic.Basic import Mathlib.Logic.Equiv.Defs import Mathlib.Logic.Equiv.LocalEquiv import Mathlib.Logic.Equiv.MfldSimpsAttr +import Mathlib.Logic.Equiv.Option import Mathlib.Logic.Function.Basic import Mathlib.Logic.Function.Conjugate import Mathlib.Logic.Function.Iterate From f4c9a49d8b746947544aa00a219b3a3ff86bbc7d Mon Sep 17 00:00:00 2001 From: Arien Malec Date: Mon, 21 Nov 2022 21:37:01 -0800 Subject: [PATCH 070/127] name fix --- Mathlib/Logic/Equiv/Option.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Logic/Equiv/Option.lean b/Mathlib/Logic/Equiv/Option.lean index 0da1a4ee8f575..def27994411e3 100644 --- a/Mathlib/Logic/Equiv/Option.lean +++ b/Mathlib/Logic/Equiv/Option.lean @@ -195,7 +195,7 @@ def option_subtype [DecidableEq β] (x : β) : { e : Option α ≃ β // e none right_inv e := by ext a rfl -#align equiv.option_subtype Equiv.optionSubtype +#align equiv.option_subtype Equiv.option_subtype @[simp] theorem option_subtype_apply_apply [DecidableEq β] (x : β) (e : { e : Option α ≃ β // e none = x }) (a : α) (h) : From 0cc1afd6a434e4b888522e394641b635582496f0 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?David=20W=C3=A4rn?= Date: Tue, 22 Nov 2022 21:32:23 +0100 Subject: [PATCH 071/127] cosmetic changes --- Mathlib/Init/Data/Sigma/Basic.lean | 20 +++----------------- 1 file changed, 3 insertions(+), 17 deletions(-) diff --git a/Mathlib/Init/Data/Sigma/Basic.lean b/Mathlib/Init/Data/Sigma/Basic.lean index 9b66f8aeb6cc2..c639f2a4960db 100644 --- a/Mathlib/Init/Data/Sigma/Basic.lean +++ b/Mathlib/Init/Data/Sigma/Basic.lean @@ -10,27 +10,13 @@ import Mathlib.Init.Logic # Lemmas about `Sigma` from Lean 3 core. -/ -universe u v - -theorem ex_of_psig {α : Type u} {p : α → Prop} : (Σ'x, p x) → ∃ x, p x +theorem ex_of_psig {p : α → Prop} : (Σ' x, p x) → ∃ x, p x | ⟨x, hx⟩ => ⟨x, hx⟩ -section - -variable {α : Type u} {β : α → Type v} - -protected theorem Sigma.eq : ∀ {p₁ p₂ : Σa : α, β a} (h₁ : p₁.1 = p₂.1), +protected theorem Sigma.eq {β : α → Type v} : ∀ {p₁ p₂ : Σ a, β a} (h₁ : p₁.1 = p₂.1), (Eq.recOn h₁ p₁.2 : β p₂.1) = p₂.2 → p₁ = p₂ | ⟨_, _⟩, _, rfl, rfl => rfl -end - -section - -variable {α : Sort u} {β : α → Sort v} - -protected theorem PSigma.eq : ∀ {p₁ p₂ : PSigma β} (h₁ : p₁.1 = p₂.1), +protected theorem PSigma.eq {β : α → Sort v} : ∀ {p₁ p₂ : Σ' a, β a} (h₁ : p₁.1 = p₂.1), (Eq.recOn h₁ p₁.2 : β p₂.1) = p₂.2 → p₁ = p₂ | ⟨_, _⟩, _, rfl, rfl => rfl - -end From ed432281355be11606371c9ea00f29d680da384f Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?David=20W=C3=A4rn?= Date: Tue, 22 Nov 2022 22:57:57 +0100 Subject: [PATCH 072/127] use autoimplicits + some casing + whitespace --- Mathlib/Logic/Equiv/Basic.lean | 440 ++++++++++++++++----------------- 1 file changed, 211 insertions(+), 229 deletions(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 2c276b0727f1a..266707e85fc51 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -25,7 +25,7 @@ In this file we continue the work on equivalences begun in `Logic/Equiv/Defs.lea * canonical isomorphisms between various types: e.g., - `Equiv.sumEquivSigmaBool` is the canonical equivalence between the sum of two types `α ⊕ β` - and the sigma-type `Σ b : bool, cond b α β`; + and the sigma-type `Σ b : Bool, cond b α β`; - `Equiv.prodSumDistrib : α × (β ⊕ γ) ≃ (α × β) ⊕ (α × γ)` shows that type product and type sum satisfy the distributive law up to a canonical equivalence; @@ -33,7 +33,7 @@ In this file we continue the work on equivalences begun in `Logic/Equiv/Defs.lea * operations on equivalences: e.g., - `Equiv.prodCongr ea eb : α₁ × β₁ ≃ α₂ × β₂`: combine two equivalences `ea : α₁ ≃ α₂` and - `eb : β₁ ≃ β₂` using `Prod.Map`. + `eb : β₁ ≃ β₂` using `Prod.map`. More definitions of this kind can be found in other files. E.g., `Data/Equiv/TransferInstance.lean` does it for many algebraic type classes like @@ -44,20 +44,13 @@ In this file we continue the work on equivalences begun in `Logic/Equiv/Defs.lea equivalence, congruence, bijective map -/ --- **TODO** remove these later -set_option autoImplicit false - open Function -universe u v w z - -variable {α : Sort u} {β : Sort v} {γ : Sort w} - namespace Equiv /-- `PProd α β` is equivalent to `α × β` -/ -@[simps apply symmApply] -def pprodEquivProd {α β : Type _} : PProd α β ≃ α × β where +@[simps apply symm_apply] +def pprodEquivProd : PProd α β ≃ α × β where toFun x := (x.1, x.2) invFun x := ⟨x.1, x.2⟩ left_inv := fun _ => rfl @@ -68,60 +61,60 @@ def pprodEquivProd {α β : Type _} : PProd α β ≃ α × β where `PProd α γ ≃ PProd β δ`. -/ -- porting note: in Lean 3 this had @[congr]` @[simps apply] -def pprodCongr {δ : Sort z} (e₁ : α ≃ β) (e₂ : γ ≃ δ) : PProd α γ ≃ PProd β δ where +def pprodCongr (e₁ : α ≃ β) (e₂ : γ ≃ δ) : PProd α γ ≃ PProd β δ where toFun x := ⟨e₁ x.1, e₂ x.2⟩ invFun x := ⟨e₁.symm x.1, e₂.symm x.2⟩ left_inv := fun ⟨x, y⟩ => by simp right_inv := fun ⟨x, y⟩ => by simp #align equiv.pprod_congr Equiv.pprodCongr -/-- Combine two equivalences using `PProd` in the domain and `prod` in the codomain. -/ -@[simps apply symmApply] -def pprodProd {α₁ β₁ : Sort _} {α₂ β₂ : Type _} (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) : +/-- Combine two equivalences using `PProd` in the domain and `Prod` in the codomain. -/ +@[simps apply symm_apply] +def pprodProd (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) : PProd α₁ β₁ ≃ α₂ × β₂ := (ea.pprodCongr eb).trans pprodEquivProd #align equiv.pprod_prod Equiv.pprodProd -/-- Combine two equivalences using `PProd` in the codomain and `prod` in the domain. -/ -@[simps apply symmApply] -def prodPProd {α₁ β₁ : Type _} {α₂ β₂ : Sort _} (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) : +/-- Combine two equivalences using `PProd` in the codomain and `Prod` in the domain. -/ +@[simps apply symm_apply] +def prodPProd (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) : α₁ × β₁ ≃ PProd α₂ β₂ := (ea.symm.pprodProd eb.symm).symm #align equiv.prod_pprod Equiv.prodPProd /-- `PProd α β` is equivalent to `Plift α × Plift β` -/ -@[simps apply symmApply] -def pprodEquivProdPLift {α β : Sort _} : PProd α β ≃ PLift α × PLift β := +@[simps apply symm_apply] +def pprodEquivProdPLift : PProd α β ≃ PLift α × PLift β := Equiv.plift.symm.pprodProd Equiv.plift.symm #align equiv.pprod_equiv_prod_plift Equiv.pprodEquivProdPLift /-- Product of two equivalences. If `α₁ ≃ α₂` and `β₁ ≃ β₂`, then `α₁ × β₁ ≃ α₂ × β₂`. This is -`prod.map` as an equivalence. -/ +`Prod.map` as an equivalence. -/ -- porting note: in Lean 3 there was also a @[congr] tag @[simps apply] -def prodCongr {α₁ β₁ α₂ β₂ : Type _} (e₁ : α₁ ≃ α₂) (e₂ : β₁ ≃ β₂) : α₁ × β₁ ≃ α₂ × β₂ := +def prodCongr (e₁ : α₁ ≃ α₂) (e₂ : β₁ ≃ β₂) : α₁ × β₁ ≃ α₂ × β₂ := ⟨Prod.map e₁ e₂, Prod.map e₁.symm e₂.symm, fun ⟨a, b⟩ => by simp, fun ⟨a, b⟩ => by simp⟩ #align equiv.prod_congr Equiv.prodCongr @[simp] -theorem prodCongr_symm {α₁ β₁ α₂ β₂ : Type _} (e₁ : α₁ ≃ α₂) (e₂ : β₁ ≃ β₂) : +theorem prodCongr_symm (e₁ : α₁ ≃ α₂) (e₂ : β₁ ≃ β₂) : (prodCongr e₁ e₂).symm = prodCongr e₁.symm e₂.symm := rfl #align equiv.prod_congr_symm Equiv.prodCongr_symm -/-- Type product is commutative up to an equivalence: `α × β ≃ β × α`. This is `prod.swap` as an +/-- Type product is commutative up to an equivalence: `α × β ≃ β × α`. This is `Prod.swap` as an equivalence.-/ -def prodComm (α β : Type _) : α × β ≃ β × α := +def prodComm (α β) : α × β ≃ β × α := ⟨Prod.swap, Prod.swap, Prod.swap_swap, Prod.swap_swap⟩ #align equiv.prod_comm Equiv.prodComm @[simp] -theorem coe_prodComm (α β : Type _) : (⇑(prodComm α β) : α × β → β × α) = Prod.swap := +theorem coe_prodComm (α β) : (⇑(prodComm α β) : α × β → β × α) = Prod.swap := rfl #align equiv.coe_prod_comm Equiv.coe_prodComm @[simp] -theorem prodComm_apply {α β : Type _} (x : α × β) : prodComm α β x = x.swap := +theorem prodComm_apply (x : α × β) : prodComm α β x = x.swap := rfl #align equiv.prod_comm_apply Equiv.prodComm_apply @@ -132,14 +125,14 @@ theorem prodComm_symm (α β) : (prodComm α β).symm = prodComm β α := /-- Type product is associative up to an equivalence. -/ @[simps] -def prodAssoc (α β γ : Sort _) : (α × β) × γ ≃ α × β × γ := +def prodAssoc (α β γ) : (α × β) × γ ≃ α × β × γ := ⟨fun p => (p.1.1, p.1.2, p.2), fun p => ((p.1, p.2.1), p.2.2), fun ⟨⟨_, _⟩, _⟩ => rfl, fun ⟨_, ⟨_, _⟩⟩ => rfl⟩ #align equiv.prod_assoc Equiv.prodAssoc /-- Functions on `α × β` are equivalent to functions `α → β → γ`. -/ @[simps (config := { fullyApplied := false })] -def curry (α β γ : Type _) : (α × β → γ) ≃ (α → β → γ) where +def curry (α β γ) : (α × β → γ) ≃ (α → β → γ) where toFun := Function.curry invFun := uncurry left_inv := uncurry_curry @@ -148,77 +141,77 @@ def curry (α β γ : Type _) : (α × β → γ) ≃ (α → β → γ) where section -/-- `punit` is a right identity for type product up to an equivalence. -/ +/-- `PUnit` is a right identity for type product up to an equivalence. -/ @[simps] -def prodPUnit (α : Type _) : α × PUnit.{u + 1} ≃ α := +def prodPUnit (α) : α × PUnit ≃ α := ⟨fun p => p.1, fun a => (a, PUnit.unit), fun ⟨_, PUnit.unit⟩ => rfl, fun _ => rfl⟩ #align equiv.prod_punit Equiv.prodPUnit -/-- `punit` is a left identity for type product up to an equivalence. -/ +/-- `PUnit` is a left identity for type product up to an equivalence. -/ @[simps] -def punitProd (α : Type _) : PUnit.{u + 1} × α ≃ α := +def punitProd (α) : PUnit × α ≃ α := calc PUnit × α ≃ α × PUnit := prodComm _ _ _ ≃ α := prodPUnit _ #align equiv.punit_prod Equiv.punitProd -/-- Any `unique` type is a right identity for type product up to equivalence. -/ -def prodUnique (α : Type u) (β : Type v) [Unique β] : α × β ≃ α := - ((Equiv.refl α).prodCongr <| equivPUnit.{v+1,v+1} β).trans <| prodPUnit α +/-- Any `Unique` type is a right identity for type product up to equivalence. -/ +def prodUnique (α β) [Unique β] : α × β ≃ α := + ((Equiv.refl α).prodCongr <| equivPUnit.{_,1} β).trans <| prodPUnit α #align equiv.prod_unique Equiv.prodUnique @[simp] -theorem coe_prodUnique {α β : Type _} [Unique β] : (⇑(prodUnique α β) : α × β → α) = Prod.fst := +theorem coe_prodUnique [Unique β] : (⇑(prodUnique α β) : α × β → α) = Prod.fst := rfl #align equiv.coe_prod_unique Equiv.coe_prodUnique -theorem prodUnique_apply {α β : Type _} [Unique β] (x : α × β) : prodUnique α β x = x.1 := +theorem prodUnique_apply [Unique β] (x : α × β) : prodUnique α β x = x.1 := rfl #align equiv.prod_unique_apply Equiv.prodUnique_apply @[simp] -theorem prodUnique_symm_apply {α β : Type _} [Unique β] (x : α) : +theorem prodUnique_symm_apply [Unique β] (x : α) : (prodUnique α β).symm x = (x, default) := rfl #align equiv.prod_unique_symm_apply Equiv.prodUnique_symm_apply -/-- Any `unique` type is a left identity for type product up to equivalence. -/ -def uniqueProd (α β : Type _) [Unique β] : β × α ≃ α := - ((equivPUnit.{v+1,v+1} β).prodCongr <| Equiv.refl α).trans <| punitProd α +/-- Any `Unique` type is a left identity for type product up to equivalence. -/ +def uniqueProd (α β) [Unique β] : β × α ≃ α := + ((equivPUnit.{_,1} β).prodCongr <| Equiv.refl α).trans <| punitProd α #align equiv.unique_prod Equiv.uniqueProd @[simp] -theorem coe_uniqueProd {α β : Type _} [Unique β] : (⇑(uniqueProd α β) : β × α → α) = Prod.snd := +theorem coe_uniqueProd [Unique β] : (⇑(uniqueProd α β) : β × α → α) = Prod.snd := rfl #align equiv.coe_unique_prod Equiv.coe_uniqueProd -theorem uniqueProd_apply {α β : Type _} [Unique β] (x : β × α) : uniqueProd α β x = x.2 := +theorem uniqueProd_apply [Unique β] (x : β × α) : uniqueProd α β x = x.2 := rfl #align equiv.unique_prod_apply Equiv.uniqueProd_apply @[simp] -theorem uniqueProd_symm_apply {α β : Type _} [Unique β] (x : α) : +theorem uniqueProd_symm_apply [Unique β] (x : α) : (uniqueProd α β).symm x = (default, x) := rfl #align equiv.unique_prod_symm_apply Equiv.uniqueProd_symm_apply -/-- `empty` type is a right absorbing element for type product up to an equivalence. -/ -def prodEmpty (α : Type _) : α × Empty ≃ Empty := +/-- `Empty` type is a right absorbing element for type product up to an equivalence. -/ +def prodEmpty (α) : α × Empty ≃ Empty := equivEmpty _ #align equiv.prod_empty Equiv.prodEmpty -/-- `empty` type is a left absorbing element for type product up to an equivalence. -/ -def emptyProd (α : Type _) : Empty × α ≃ Empty := +/-- `Empty` type is a left absorbing element for type product up to an equivalence. -/ +def emptyProd (α) : Empty × α ≃ Empty := equivEmpty _ #align equiv.empty_prod Equiv.emptyProd -/-- `pempty` type is a right absorbing element for type product up to an equivalence. -/ -def prodPEmpty (α : Type _) : α × PEmpty ≃ PEmpty := +/-- `PEmpty` type is a right absorbing element for type product up to an equivalence. -/ +def prodPEmpty (α) : α × PEmpty ≃ PEmpty := equivPEmpty _ #align equiv.prod_pempty Equiv.prodPEmpty -/-- `pempty` type is a left absorbing element for type product up to an equivalence. -/ -def pemptyProd (α : Type _) : PEmpty × α ≃ PEmpty := +/-- `PEmpty` type is a left absorbing element for type product up to an equivalence. -/ +def pemptyProd (α) : PEmpty × α ≃ PEmpty := equivPEmpty _ #align equiv.pempty_prod Equiv.pemptyProd @@ -228,57 +221,55 @@ section open Sum -/-- `psum` is equivalent to `sum`. -/ -def psumEquivSum (α β : Type _) : PSum α β ≃ Sum α β where +/-- `PSum` is equivalent to `Sum`. -/ +def psumEquivSum (α β) : PSum α β ≃ Sum α β where toFun s := PSum.casesOn s inl inr invFun := Sum.elim PSum.inl PSum.inr left_inv s := by cases s <;> rfl right_inv s := by cases s <;> rfl #align equiv.psum_equiv_sum Equiv.psumEquivSum -/-- If `α ≃ α'` and `β ≃ β'`, then `α ⊕ β ≃ α' ⊕ β'`. This is `sum.map` as an equivalence. -/ +/-- If `α ≃ α'` and `β ≃ β'`, then `α ⊕ β ≃ α' ⊕ β'`. This is `Sum.map` as an equivalence. -/ @[simps apply] -def sumCongr {α₁ β₁ α₂ β₂ : Type _} (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) : Sum α₁ β₁ ≃ Sum α₂ β₂ := +def sumCongr (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) : Sum α₁ β₁ ≃ Sum α₂ β₂ := ⟨Sum.map ea eb, Sum.map ea.symm eb.symm, fun x => by simp, fun x => by simp⟩ #align equiv.sum_congr Equiv.sumCongr -/-- If `α ≃ α'` and `β ≃ β'`, then `psum α β ≃ psum α' β'`. -/ -def psumCongr {δ : Sort z} (e₁ : α ≃ β) (e₂ : γ ≃ δ) : PSum α γ ≃ PSum β δ where +/-- If `α ≃ α'` and `β ≃ β'`, then `PSum α β ≃ PSum α' β'`. -/ +def psumCongr (e₁ : α ≃ β) (e₂ : γ ≃ δ) : PSum α γ ≃ PSum β δ where toFun x := PSum.casesOn x (PSum.inl ∘ e₁) (PSum.inr ∘ e₂) invFun x := PSum.casesOn x (PSum.inl ∘ e₁.symm) (PSum.inr ∘ e₂.symm) left_inv := by rintro (x | x) <;> simp right_inv := by rintro (x | x) <;> simp #align equiv.psum_congr Equiv.psumCongr -/-- Combine two `equiv`s using `psum` in the domain and `sum` in the codomain. -/ -def psumSum {α₁ β₁ : Sort _} {α₂ β₂ : Type _} (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) : +/-- Combine two `Equiv`s using `PSum` in the domain and `Sum` in the codomain. -/ +def psumSum (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) : PSum α₁ β₁ ≃ Sum α₂ β₂ := (ea.psumCongr eb).trans (psumEquivSum _ _) #align equiv.psum_sum Equiv.psumSum -/-- Combine two `equiv`s using `sum` in the domain and `psum` in the codomain. -/ -def sumPSum {α₁ β₁ : Type _} {α₂ β₂ : Sort _} (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) : +/-- Combine two `Equiv`s using `Sum` in the domain and `PSum` in the codomain. -/ +def sumPSum (ea : α₁ ≃ α₂) (eb : β₁ ≃ β₂) : Sum α₁ β₁ ≃ PSum α₂ β₂ := (ea.symm.psumSum eb.symm).symm #align equiv.sum_psum Equiv.sumPSum @[simp] -theorem sumCongr_trans {α₁ α₂ β₁ β₂ γ₁ γ₂ : Sort _} (e : α₁ ≃ β₁) (f : α₂ ≃ β₂) (g : β₁ ≃ γ₁) - (h : β₂ ≃ γ₂) : +theorem sumCongr_trans (e : α₁ ≃ β₁) (f : α₂ ≃ β₂) (g : β₁ ≃ γ₁) (h : β₂ ≃ γ₂) : (Equiv.sumCongr e f).trans (Equiv.sumCongr g h) = Equiv.sumCongr (e.trans g) (f.trans h) := by ext i cases i <;> rfl #align equiv.sum_congr_trans Equiv.sumCongr_trans @[simp] -theorem sumCongr_symm {α β γ δ : Sort _} (e : α ≃ β) (f : γ ≃ δ) : +theorem sumCongr_symm (e : α ≃ β) (f : γ ≃ δ) : (Equiv.sumCongr e f).symm = Equiv.sumCongr e.symm f.symm := rfl #align equiv.sum_congr_symm Equiv.sumCongr_symm @[simp] -theorem sumCongr_refl {α β : Sort _} : - Equiv.sumCongr (Equiv.refl α) (Equiv.refl β) = Equiv.refl (Sum α β) := by +theorem sumCongr_refl : Equiv.sumCongr (Equiv.refl α) (Equiv.refl β) = Equiv.refl (Sum α β) := by ext i cases i <;> rfl #align equiv.sum_congr_refl Equiv.sumCongr_refl @@ -287,12 +278,12 @@ namespace Perm /-- Combine a permutation of `α` and of `β` into a permutation of `α ⊕ β`. -/ @[reducible] -def sumCongr {α β : Type _} (ea : Equiv.Perm α) (eb : Equiv.Perm β) : Equiv.Perm (Sum α β) := +def sumCongr (ea : Equiv.Perm α) (eb : Equiv.Perm β) : Equiv.Perm (Sum α β) := Equiv.sumCongr ea eb #align equiv.perm.sum_congr Equiv.Perm.sumCongr @[simp] -theorem sumCongr_apply {α β : Type _} (ea : Equiv.Perm α) (eb : Equiv.Perm β) (x : Sum α β) : +theorem sumCongr_apply (ea : Equiv.Perm α) (eb : Equiv.Perm β) (x : Sum α β) : sumCongr ea eb x = Sum.map (⇑ea) (⇑eb) x := Equiv.sumCongr_apply ea eb x #align equiv.perm.sum_congr_apply Equiv.Perm.sumCongr_apply @@ -300,32 +291,31 @@ theorem sumCongr_apply {α β : Type _} (ea : Equiv.Perm α) (eb : Equiv.Perm β -- porting note: it seems the general theorem about `Equiv` is now applied, so there's no need -- to have this version also have `@[simp]`. Similarly for below. @[simp] -theorem sumCongr_trans {α β : Sort _} (e : Equiv.Perm α) (f : Equiv.Perm β) (g : Equiv.Perm α) +theorem sumCongr_trans (e : Equiv.Perm α) (f : Equiv.Perm β) (g : Equiv.Perm α) (h : Equiv.Perm β) : (sumCongr e f).trans (sumCongr g h) = sumCongr (e.trans g) (f.trans h) := Equiv.sumCongr_trans e f g h #align equiv.perm.sum_congr_trans Equiv.Perm.sumCongr_trans -theorem sumCongr_symm {α β : Sort _} (e : Equiv.Perm α) (f : Equiv.Perm β) : +theorem sumCongr_symm (e : Equiv.Perm α) (f : Equiv.Perm β) : (sumCongr e f).symm = sumCongr e.symm f.symm := Equiv.sumCongr_symm e f #align equiv.perm.sum_congr_symm Equiv.Perm.sumCongr_symm -theorem sumCongr_refl {α β : Sort _} : - sumCongr (Equiv.refl α) (Equiv.refl β) = Equiv.refl (Sum α β) := +theorem sumCongr_refl : sumCongr (Equiv.refl α) (Equiv.refl β) = Equiv.refl (Sum α β) := Equiv.sumCongr_refl #align equiv.perm.sum_congr_refl Equiv.Perm.sumCongr_refl end Perm -/-- `bool` is equivalent the sum of two `punit`s. -/ +/-- `Bool` is equivalent the sum of two `PUnit`s. -/ def boolEquivPUnitSumPUnit : Bool ≃ Sum PUnit.{u + 1} PUnit.{v + 1} := ⟨fun b => cond b (inr PUnit.unit) (inl PUnit.unit), Sum.elim (fun _ => false) fun _ => true, fun b => by cases b <;> rfl, fun s => by rcases s with (⟨⟨⟩⟩ | ⟨⟨⟩⟩) <;> rfl⟩ #align equiv.bool_equiv_punit_sum_punit Equiv.boolEquivPUnitSumPUnit -/-- Sum of types is commutative up to an equivalence. This is `sum.swap` as an equivalence. -/ +/-- Sum of types is commutative up to an equivalence. This is `Sum.swap` as an equivalence. -/ @[simps (config := { fullyApplied := false }) apply] -def sumComm (α β : Type _) : Sum α β ≃ Sum β α := +def sumComm (α β) : Sum α β ≃ Sum β α := ⟨Sum.swap, Sum.swap, Sum.swap_swap, Sum.swap_swap⟩ #align equiv.sum_comm Equiv.sumComm @@ -335,7 +325,7 @@ theorem sumComm_symm (α β) : (sumComm α β).symm = sumComm β α := #align equiv.sum_comm_symm Equiv.sumComm_symm /-- Sum of types is associative up to an equivalence. -/ -def sumAssoc (α β γ : Type _) : Sum (Sum α β) γ ≃ Sum α (Sum β γ) := +def sumAssoc (α β γ) : Sum (Sum α β) γ ≃ Sum α (Sum β γ) := ⟨Sum.elim (Sum.elim Sum.inl (Sum.inr ∘ Sum.inl)) (Sum.inr ∘ Sum.inr), Sum.elim (Sum.inl ∘ Sum.inl) <| Sum.elim (Sum.inl ∘ Sum.inr) Sum.inr, by rintro (⟨_ | _⟩ | _) <;> rfl, by @@ -343,17 +333,17 @@ def sumAssoc (α β γ : Type _) : Sum (Sum α β) γ ≃ Sum α (Sum β γ) := #align equiv.sum_assoc Equiv.sumAssoc @[simp] -theorem sumAssoc_apply_inl_inl {α β γ} (a) : sumAssoc α β γ (inl (inl a)) = inl a := +theorem sumAssoc_apply_inl_inl (a) : sumAssoc α β γ (inl (inl a)) = inl a := rfl #align equiv.sum_assoc_apply_inl_inl Equiv.sumAssoc_apply_inl_inl @[simp] -theorem sumAssoc_apply_inl_inr {α β γ} (b) : sumAssoc α β γ (inl (inr b)) = inr (inl b) := +theorem sumAssoc_apply_inl_inr (b) : sumAssoc α β γ (inl (inr b)) = inr (inl b) := rfl #align equiv.sum_assoc_apply_inl_inr Equiv.sumAssoc_apply_inl_inr @[simp] -theorem sumAssoc_apply_inr {α β γ} (c) : sumAssoc α β γ (inr c) = inr (inr c) := +theorem sumAssoc_apply_inr (c) : sumAssoc α β γ (inr c) = inr (inr c) := rfl #align equiv.sum_assoc_apply_inr Equiv.sumAssoc_apply_inr @@ -373,9 +363,9 @@ theorem sumAssoc_symm_apply_inr_inr {α β γ} (c) : (sumAssoc α β γ).symm (i rfl #align equiv.sum_assoc_symm_apply_inr_inr Equiv.sumAssoc_symm_apply_inr_inr -/-- Sum with `empty` is equivalent to the original type. -/ -@[simps symmApply] -def sumEmpty (α β : Type _) [IsEmpty β] : Sum α β ≃ α := +/-- Sum with `IsEmpty` is equivalent to the original type. -/ +@[simps symm_apply] +def sumEmpty (α β) [IsEmpty β] : Sum α β ≃ α := ⟨Sum.elim id isEmptyElim, inl, fun s => by rcases s with (_ | x) rfl @@ -383,66 +373,66 @@ def sumEmpty (α β : Type _) [IsEmpty β] : Sum α β ≃ α := #align equiv.sum_empty Equiv.sumEmpty @[simp] -theorem sumEmpty_apply_inl {α β : Type _} [IsEmpty β] (a : α) : sumEmpty α β (Sum.inl a) = a := +theorem sumEmpty_apply_inl [IsEmpty β] (a : α) : sumEmpty α β (Sum.inl a) = a := rfl #align equiv.sum_empty_apply_inl Equiv.sumEmpty_apply_inl -/-- The sum of `empty` with any `Sort*` is equivalent to the right summand. -/ -@[simps symmApply] -def emptySum (α β : Type _) [IsEmpty α] : Sum α β ≃ β := +/-- The sum of `IsEmpty` with any type is equivalent to that type. -/ +@[simps symm_apply] +def emptySum (α β) [IsEmpty α] : Sum α β ≃ β := (sumComm _ _).trans <| sumEmpty _ _ #align equiv.empty_sum Equiv.emptySum @[simp] -theorem emptySum_apply_inr {α β : Type _} [IsEmpty α] (b : β) : emptySum α β (Sum.inr b) = b := +theorem emptySum_apply_inr [IsEmpty α] (b : β) : emptySum α β (Sum.inr b) = b := rfl #align equiv.empty_sum_apply_inr Equiv.emptySum_apply_inr -/-- `option α` is equivalent to `α ⊕ punit` -/ -def optionEquivSumPUnit (α : Type _) : Option α ≃ Sum α PUnit.{u + 1} := +/-- `Option α` is equivalent to `α ⊕ punit` -/ +def optionEquivSumPUnit (α) : Option α ≃ Sum α PUnit := ⟨fun o => o.elim (inr PUnit.unit) inl, fun s => s.elim some fun _ => none, fun o => by cases o <;> rfl, fun s => by rcases s with (_ | ⟨⟨⟩⟩) <;> rfl⟩ #align equiv.option_equiv_sum_punit Equiv.optionEquivSumPUnit @[simp] -theorem optionEquivSumPUnit_none {α} : optionEquivSumPUnit α none = Sum.inr PUnit.unit := +theorem optionEquivSumPUnit_none : optionEquivSumPUnit α none = Sum.inr PUnit.unit := rfl #align equiv.option_equiv_sum_punit_none Equiv.optionEquivSumPUnit_none @[simp] -theorem optionEquivSumPUnit_some {α} (a) : optionEquivSumPUnit α (some a) = Sum.inl a := +theorem optionEquivSumPUnit_some (a) : optionEquivSumPUnit α (some a) = Sum.inl a := rfl #align equiv.option_equiv_sum_punit_some Equiv.optionEquivSumPUnit_some @[simp] -theorem optionEquivSumPUnit_coe {α} (a : α) : optionEquivSumPUnit α a = Sum.inl a := +theorem optionEquivSumPUnit_coe (a : α) : optionEquivSumPUnit α a = Sum.inl a := rfl #align equiv.option_equiv_sum_punit_coe Equiv.optionEquivSumPUnit_coe @[simp] -theorem optionEquivSumPUnit_symm_inl {α} (a) : (optionEquivSumPUnit α).symm (Sum.inl a) = a := +theorem optionEquivSumPUnit_symm_inl (a) : (optionEquivSumPUnit α).symm (Sum.inl a) = a := rfl #align equiv.option_equiv_sum_punit_symm_inl Equiv.optionEquivSumPUnit_symm_inl @[simp] -theorem optionEquivSumPUnit_symm_inr {α} (a) : (optionEquivSumPUnit α).symm (Sum.inr a) = none := +theorem optionEquivSumPUnit_symm_inr (a) : (optionEquivSumPUnit α).symm (Sum.inr a) = none := rfl #align equiv.option_equiv_sum_punit_symm_inr Equiv.optionEquivSumPUnit_symm_inr -/-- The set of `x : option α` such that `is_some x` is equivalent to `α`. -/ +/-- The set of `x : Option α` such that `isSome x` is equivalent to `α`. -/ @[simps] -def optionIsSomeEquiv (α : Type _) : { x : Option α // x.isSome } ≃ α where +def optionIsSomeEquiv (α) : { x : Option α // x.isSome } ≃ α where toFun o := Option.get _ o.2 invFun x := ⟨some x, rfl⟩ left_inv _ := Subtype.eq <| Option.some_get _ right_inv _ := Option.get_some _ _ #align equiv.option_is_some_equiv Equiv.optionIsSomeEquiv -/-- The product over `option α` of `β a` is the binary product of the +/-- The product over `Option α` of `β a` is the binary product of the product over `α` of `β (some α)` and `β none` -/ @[simps] -def piOptionEquivProd {α : Type _} {β : Option α → Type _} : +def piOptionEquivProd {β : Option α → Type _} : (∀ a : Option α, β a) ≃ β none × ∀ a : α, β (some a) where toFun f := (f none, fun a => f (some a)) invFun x a := Option.casesOn a x.fst x.snd @@ -450,11 +440,11 @@ def piOptionEquivProd {α : Type _} {β : Option α → Type _} : right_inv x := by simp #align equiv.pi_option_equiv_prod Equiv.piOptionEquivProd -/-- `α ⊕ β` is equivalent to a `sigma`-type over `bool`. Note that this definition assumes `α` and +/-- `α ⊕ β` is equivalent to a `Sigma`-type over `Bool`. Note that this definition assumes `α` and `β` to be types from the same universe, so it cannot by used directly to transfer theorems about sigma types to theorems about sum types. In many cases one can use `ulift` to work around this difficulty. -/ -def sumEquivSigmaBool (α β : Type u) : Sum α β ≃ Σb : Bool, cond b α β := +def sumEquivSigmaBool (α β : Type u) : Sum α β ≃ Σ b : Bool, cond b α β := ⟨fun s => s.elim (fun x => ⟨true, x⟩) fun x => ⟨false, x⟩, fun s => match s with | ⟨true, a⟩ => inl a @@ -462,11 +452,11 @@ def sumEquivSigmaBool (α β : Type u) : Sum α β ≃ Σb : Bool, cond b α β fun s => by cases s <;> rfl, fun s => by rcases s with ⟨_ | _, _⟩ <;> rfl⟩ #align equiv.sum_equiv_sigma_bool Equiv.sumEquivSigmaBool --- See also `equiv.sigma_preimage_equiv`. -/-- `sigma_fiber_equiv f` for `f : α → β` is the natural equivalence between +-- See also `Equiv.sigmaPreimageEquiv`. +/-- `sigmaFiberEquiv f` for `f : α → β` is the natural equivalence between the type of all fibres of `f` and the total space `α`. -/ @[simps] -def sigmaFiberEquiv {α β : Type _} (f : α → β) : (Σy : β, { x // f x = y }) ≃ α := +def sigmaFiberEquiv {α β : Type _} (f : α → β) : (Σ y : β, { x // f x = y }) ≃ α := ⟨fun x => ↑x.2, fun x => ⟨f x, x, rfl⟩, fun ⟨_, _, rfl⟩ => rfl, fun _ => rfl⟩ #align equiv.sigma_fiber_equiv Equiv.sigmaFiberEquiv @@ -478,8 +468,8 @@ section sumCompl the sum of the two subtypes `{a // p a}` and its complement `{a // ¬ p a}` is naturally equivalent to `α`. -See `subtype_or_equiv` for sum types over subtypes `{x // p x}` and `{x // q x}` -that are not necessarily `is_compl p q`. -/ +See `subtypeOrEquiv` for sum types over subtypes `{x // p x}` and `{x // q x}` +that are not necessarily `IsCompl p q`. -/ def sumCompl {α : Type _} (p : α → Prop) [DecidablePred p] : Sum { a // p a } { a // ¬p a } ≃ α where toFun := Sum.elim Subtype.val Subtype.val @@ -494,39 +484,39 @@ def sumCompl {α : Type _} (p : α → Prop) [DecidablePred p] : #align equiv.sum_compl Equiv.sumCompl @[simp] -theorem sumCompl_apply_inl {α : Type _} (p : α → Prop) [DecidablePred p] (x : { a // p a }) : +theorem sumCompl_apply_inl (p : α → Prop) [DecidablePred p] (x : { a // p a }) : sumCompl p (Sum.inl x) = x := rfl #align equiv.sum_compl_apply_inl Equiv.sumCompl_apply_inl @[simp] -theorem sumCompl_apply_inr {α : Type _} (p : α → Prop) [DecidablePred p] (x : { a // ¬p a }) : +theorem sumCompl_apply_inr (p : α → Prop) [DecidablePred p] (x : { a // ¬p a }) : sumCompl p (Sum.inr x) = x := rfl #align equiv.sum_compl_apply_inr Equiv.sumCompl_apply_inr @[simp] -theorem sumCompl_apply_symm_of_pos {α : Type _} (p : α → Prop) [DecidablePred p] (a : α) (h : p a) : +theorem sumCompl_apply_symm_of_pos (p : α → Prop) [DecidablePred p] (a : α) (h : p a) : (sumCompl p).symm a = Sum.inl ⟨a, h⟩ := dif_pos h #align equiv.sum_compl_apply_symm_of_pos Equiv.sumCompl_apply_symm_of_pos @[simp] -theorem sumCompl_apply_symm_of_neg {α : Type _} (p : α → Prop) [DecidablePred p] (a : α) +theorem sumCompl_apply_symm_of_neg (p : α → Prop) [DecidablePred p] (a : α) (h : ¬p a) : (sumCompl p).symm a = Sum.inr ⟨a, h⟩ := dif_neg h #align equiv.sum_compl_apply_symm_of_neg Equiv.sumCompl_apply_symm_of_neg -/-- Combines an `equiv` between two subtypes with an `equiv` between their complements to form a +/-- Combines an `Equiv` between two subtypes with an `Equiv` between their complements to form a permutation. -/ -def subtypeCongr {α : Type _} {p q : α → Prop} [DecidablePred p] [DecidablePred q] +def subtypeCongr {p q : α → Prop} [DecidablePred p] [DecidablePred q] (e : { x // p x } ≃ { x // q x }) (f : { x // ¬p x } ≃ { x // ¬q x }) : Perm α := (sumCompl p).symm.trans ((sumCongr e f).trans (sumCompl q)) #align equiv.subtype_congr Equiv.subtypeCongr open Equiv -variable {ε : Type _} {p : ε → Prop} [DecidablePred p] +variable {p : ε → Prop} [DecidablePred p] variable (ep ep' : Perm { a // p a }) (en en' : Perm { a // ¬p a }) @@ -635,32 +625,32 @@ end subtypePreimage section -/-- A family of equivalences `Π a, β₁ a ≃ β₂ a` generates an equivalence between `Π a, β₁ a` and -`Π a, β₂ a`. -/ -def piCongrRight {α} {β₁ β₂ : α → Sort _} (F : ∀ a, β₁ a ≃ β₂ a) : (∀ a, β₁ a) ≃ ∀ a, β₂ a := +/-- A family of equivalences `∀ a, β₁ a ≃ β₂ a` generates an equivalence between `∀ a, β₁ a` and +`∀ a, β₂ a`. -/ +def piCongrRight {β₁ β₂ : α → Sort _} (F : ∀ a, β₁ a ≃ β₂ a) : (∀ a, β₁ a) ≃ (∀ a, β₂ a) := ⟨fun H a => F a (H a), fun H a => (F a).symm (H a), fun H => funext <| by simp, fun H => funext <| by simp⟩ #align equiv.Pi_congr_right Equiv.piCongrRight -/-- Given `φ : α → β → Sort*`, we have an equivalence between `Π a b, φ a b` and `Π b a, φ a b`. -This is `function.swap` as an `equiv`. -/ +/-- Given `φ : α → β → Sort*`, we have an equivalence between `∀ a b, φ a b` and `∀ b a, φ a b`. +This is `Function.swap` as an `Equiv`. -/ @[simps apply] -def piComm {α β} (φ : α → β → Sort _) : (∀ a b, φ a b) ≃ ∀ b a, φ a b := +def piComm (φ : α → β → Sort _) : (∀ a b, φ a b) ≃ ∀ b a, φ a b := ⟨swap, swap, fun _ => rfl, fun _ => rfl⟩ #align equiv.Pi_comm Equiv.piComm -- up to here **TODO** remove this @[simp] -theorem piComm_symm {α β} {φ : α → β → Sort _} : (piComm φ).symm = (piComm <| swap φ) := +theorem piComm_symm {φ : α → β → Sort _} : (piComm φ).symm = (piComm <| swap φ) := rfl #align equiv.Pi_comm_symm Equiv.piComm_symm /-- Dependent `curry` equivalence: the type of dependent functions on `Σ i, β i` is equivalent to the type of dependent functions of two arguments (i.e., functions to the space of functions). -This is `sigma.curry` and `sigma.uncurry` together as an equiv. -/ -def piCurry {α} {β : α → Sort _} (γ : ∀ a, β a → Sort _) : - (∀ x : Σi, β i, γ x.1 x.2) ≃ ∀ a b, γ a b where +This is `Sigma.curry` and `Sigma.uncurry` together as an equiv. -/ +def piCurry {β : α → Sort _} (γ : ∀ a, β a → Sort _) : + (∀ x : Σ i, β i, γ x.1 x.2) ≃ ∀ a b, γ a b where toFun := Sigma.curry invFun := Sigma.uncurry left_inv := Sigma.uncurry_curry @@ -671,9 +661,9 @@ end section prodCongr -variable {α₁ β₁ β₂ : Type _} (e : α₁ → β₁ ≃ β₂) +variable (e : α₁ → β₁ ≃ β₂) -/-- A family of equivalences `Π (a : α₁), β₁ ≃ β₂` generates an equivalence +/-- A family of equivalences `∀ (a : α₁), β₁ ≃ β₂` generates an equivalence between `β₁ × α₁` and `β₂ × α₁`. -/ def prodCongrLeft : β₁ × α₁ ≃ β₂ × α₁ where toFun ab := ⟨e ab.2 ab.1, ab.2⟩ @@ -687,9 +677,9 @@ def prodCongrLeft : β₁ × α₁ ≃ β₂ × α₁ where #align equiv.prod_congr_left Equiv.prodCongrLeft @[simp] -theorem prodCongr_left_apply (b : β₁) (a : α₁) : prodCongrLeft e (b, a) = (e a b, a) := +theorem prodCongrLeft_apply (b : β₁) (a : α₁) : prodCongrLeft e (b, a) = (e a b, a) := rfl -#align equiv.prod_congr_left_apply Equiv.prodCongr_left_apply +#align equiv.prod_congr_left_apply Equiv.prodCongrLeft_apply theorem prodCongr_refl_right (e : β₁ ≃ β₂) : prodCongr e (Equiv.refl α₁) = prodCongrLeft fun _ => e := by @@ -697,7 +687,7 @@ theorem prodCongr_refl_right (e : β₁ ≃ β₂) : simp #align equiv.prod_congr_refl_right Equiv.prodCongr_refl_right -/-- A family of equivalences `Π (a : α₁), β₁ ≃ β₂` generates an equivalence +/-- A family of equivalences `∀ (a : α₁), β₁ ≃ β₂` generates an equivalence between `α₁ × β₁` and `α₁ × β₂`. -/ def prodCongrRight : α₁ × β₁ ≃ α₁ × β₂ where toFun ab := ⟨ab.1, e ab.1 ab.2⟩ @@ -711,9 +701,9 @@ def prodCongrRight : α₁ × β₁ ≃ α₁ × β₂ where #align equiv.prod_congr_right Equiv.prodCongrRight @[simp] -theorem prodCongr_right_apply (a : α₁) (b : β₁) : prodCongrRight e (a, b) = (a, e a b) := +theorem prodCongrRight_apply (a : α₁) (b : β₁) : prodCongrRight e (a, b) = (a, e a b) := rfl -#align equiv.prod_congr_right_apply Equiv.prodCongr_right_apply +#align equiv.prod_congr_right_apply Equiv.prodCongrRight_apply theorem prodCongr_refl_left (e : β₁ ≃ β₂) : prodCongr (Equiv.refl α₁) e = prodCongrRight fun _ => e := by @@ -722,39 +712,38 @@ theorem prodCongr_refl_left (e : β₁ ≃ β₂) : #align equiv.prod_congr_refl_left Equiv.prodCongr_refl_left @[simp] -theorem prodCongr_left_trans_prod_comm : +theorem prodCongrLeft_trans_prodComm : (prodCongrLeft e).trans (prodComm _ _) = (prodComm _ _).trans (prodCongrRight e) := by ext ⟨a, b⟩ : 1 simp -#align equiv.prod_congr_left_trans_prod_comm Equiv.prodCongr_left_trans_prod_comm +#align equiv.prod_congr_left_trans_prod_comm Equiv.prodCongrLeft_trans_prodComm @[simp] -theorem prodCongr_right_trans_prod_comm : +theorem prodCongrRight_trans_prodComm : (prodCongrRight e).trans (prodComm _ _) = (prodComm _ _).trans (prodCongrLeft e) := by ext ⟨a, b⟩ : 1 simp -#align equiv.prod_congr_right_trans_prod_comm Equiv.prodCongr_right_trans_prod_comm +#align equiv.prod_congr_right_trans_prod_comm Equiv.prodCongrRight_trans_prodComm -theorem sigma_congr_right_sigma_equiv_prod : +theorem sigmaCongrRight_sigmaEquivProd : (sigmaCongrRight e).trans (sigmaEquivProd α₁ β₂) = (sigmaEquivProd α₁ β₁).trans (prodCongrRight e) := by ext ⟨a, b⟩ : 1 simp -#align equiv.sigma_congr_right_sigma_equiv_prod Equiv.sigma_congr_right_sigma_equiv_prod +#align equiv.sigma_congr_right_sigma_equiv_prod Equiv.sigmaCongrRight_sigmaEquivProd theorem sigmaEquivProd_sigmaCongrRight : (sigmaEquivProd α₁ β₁).symm.trans (sigmaCongrRight e) = (prodCongrRight e).trans (sigmaEquivProd α₁ β₂).symm := by ext ⟨a, b⟩ : 1 - simp only [trans_apply, sigmaCongrRight_apply, prodCongr_right_apply] + simp only [trans_apply, sigmaCongrRight_apply, prodCongrRight_apply] rfl #align equiv.sigma_equiv_prod_sigma_congr_right Equiv.sigmaEquivProd_sigmaCongrRight -- See also `Equiv.ofPreimageEquiv`. /-- A family of equivalences between fibers gives an equivalence between domains. -/ @[simps] -def ofFiberEquiv {α β γ : Type _} {f : α → γ} {g : β → γ} - (e : ∀ c, { a // f a = c } ≃ { b // g b = c }) : α ≃ β := +def ofFiberEquiv {f : α → γ} {g : β → γ} (e : ∀ c, { a // f a = c } ≃ { b // g b = c }) : α ≃ β := (sigmaFiberEquiv f).symm.trans <| (Equiv.sigmaCongrRight e).trans (sigmaFiberEquiv g) #align equiv.of_fiber_equiv Equiv.ofFiberEquiv @@ -767,7 +756,7 @@ theorem ofFiberEquiv_map {α β γ} {f : α → γ} {g : β → γ} on the first component. A typical example is a shear mapping, explaining the name of this declaration. -/ @[simps (config := { fullyApplied := false })] -def prodShear {α₁ β₁ α₂ β₂ : Type _} (e₁ : α₁ ≃ α₂) (e₂ : α₁ → β₁ ≃ β₂) : α₁ × β₁ ≃ α₂ × β₂ where +def prodShear (e₁ : α₁ ≃ α₂) (e₂ : α₁ → β₁ ≃ β₂) : α₁ × β₁ ≃ α₂ × β₂ where toFun := fun x : α₁ × β₁ => (e₁ x.1, e₂ x.1 x.2) invFun := fun y : α₂ × β₂ => (e₁.symm y.1, (e₂ <| e₁.symm y.1).symm y.2) left_inv := by @@ -782,9 +771,9 @@ end prodCongr namespace Perm -variable {α₁ β₁ β₂ : Type _} [DecidableEq α₁] (a : α₁) (e : Perm β₁) +variable [DecidableEq α₁] (a : α₁) (e : Perm β₁) -/-- `prod_extend_right a e` extends `e : perm β` to `perm (α × β)` by sending `(a, b)` to +/-- `prodExtendRight a e` extends `e : Perm β` to `Perm (α × β)` by sending `(a, b)` to `(a, e b)` and keeping the other `(a', b)` fixed. -/ def prodExtendRight : Perm (α₁ × β₁) where toFun ab := if ab.fst = a then (a, e ab.snd) else ab @@ -855,105 +844,104 @@ def sumArrowEquivProdArrow (α β γ : Type _) : (Sum α β → γ) ≃ (α → #align equiv.sum_arrow_equiv_prod_arrow Equiv.sumArrowEquivProdArrow @[simp] -theorem sumArrowEquivProdArrow_apply_fst {α β γ} (f : Sum α β → γ) (a : α) : +theorem sumArrowEquivProdArrow_apply_fst (f : Sum α β → γ) (a : α) : (sumArrowEquivProdArrow α β γ f).1 a = f (inl a) := rfl #align equiv.sum_arrow_equiv_prod_arrow_apply_fst Equiv.sumArrowEquivProdArrow_apply_fst @[simp] -theorem sumArrowEquivProdArrow_apply_snd {α β γ} (f : Sum α β → γ) (b : β) : +theorem sumArrowEquivProdArrow_apply_snd (f : Sum α β → γ) (b : β) : (sumArrowEquivProdArrow α β γ f).2 b = f (inr b) := rfl #align equiv.sum_arrow_equiv_prod_arrow_apply_snd Equiv.sumArrowEquivProdArrow_apply_snd @[simp] -theorem sumArrowEquivProdArrow_symm_apply_inl {α β γ} (f : α → γ) (g : β → γ) (a : α) : +theorem sumArrowEquivProdArrow_symm_apply_inl (f : α → γ) (g : β → γ) (a : α) : ((sumArrowEquivProdArrow α β γ).symm (f, g)) (inl a) = f a := rfl #align equiv.sum_arrow_equiv_prod_arrow_symm_apply_inl Equiv.sumArrowEquivProdArrow_symm_apply_inl @[simp] -theorem sumArrowEquivProdArrow_symm_apply_inr {α β γ} (f : α → γ) (g : β → γ) (b : β) : +theorem sumArrowEquivProdArrow_symm_apply_inr (f : α → γ) (g : β → γ) (b : β) : ((sumArrowEquivProdArrow α β γ).symm (f, g)) (inr b) = g b := rfl #align equiv.sum_arrow_equiv_prod_arrow_symm_apply_inr Equiv.sumArrowEquivProdArrow_symm_apply_inr /-- Type product is right distributive with respect to type sum up to an equivalence. -/ -def sumProdDistrib (α β γ : Sort _) : Sum α β × γ ≃ Sum (α × γ) (β × γ) := +def sumProdDistrib (α β γ) : Sum α β × γ ≃ Sum (α × γ) (β × γ) := ⟨fun p => p.1.map (fun x => (x, p.2)) fun x => (x, p.2), fun s => s.elim (Prod.map inl id) (Prod.map inr id), by rintro ⟨_ | _, _⟩ <;> rfl, by rintro (⟨_, _⟩ | ⟨_, _⟩) <;> rfl⟩ #align equiv.sum_prod_distrib Equiv.sumProdDistrib @[simp] -theorem sumProdDistrib_apply_left {α β γ} (a : α) (c : γ) : +theorem sumProdDistrib_apply_left (a : α) (c : γ) : sumProdDistrib α β γ (Sum.inl a, c) = Sum.inl (a, c) := rfl #align equiv.sum_prod_distrib_apply_left Equiv.sumProdDistrib_apply_left @[simp] -theorem sumProdDistrib_apply_right {α β γ} (b : β) (c : γ) : +theorem sumProdDistrib_apply_right (b : β) (c : γ) : sumProdDistrib α β γ (Sum.inr b, c) = Sum.inr (b, c) := rfl #align equiv.sum_prod_distrib_apply_right Equiv.sumProdDistrib_apply_right @[simp] -theorem sumProdDistrib_symm_apply_left {α β γ} (a : α × γ) : +theorem sumProdDistrib_symm_apply_left (a : α × γ) : (sumProdDistrib α β γ).symm (inl a) = (inl a.1, a.2) := rfl #align equiv.sum_prod_distrib_symm_apply_left Equiv.sumProdDistrib_symm_apply_left @[simp] -theorem sumProdDistrib_symm_apply_right {α β γ} (b : β × γ) : +theorem sumProdDistrib_symm_apply_right (b : β × γ) : (sumProdDistrib α β γ).symm (inr b) = (inr b.1, b.2) := rfl #align equiv.sum_prod_distrib_symm_apply_right Equiv.sumProdDistrib_symm_apply_right /-- Type product is left distributive with respect to type sum up to an equivalence. -/ -def prodSumDistrib (α β γ : Sort _) : α × Sum β γ ≃ Sum (α × β) (α × γ) := +def prodSumDistrib (α β γ) : α × Sum β γ ≃ Sum (α × β) (α × γ) := calc α × Sum β γ ≃ Sum β γ × α := prodComm _ _ _ ≃ Sum (β × α) (γ × α) := sumProdDistrib _ _ _ _ ≃ Sum (α × β) (α × γ) := sumCongr (prodComm _ _) (prodComm _ _) - #align equiv.prod_sum_distrib Equiv.prodSumDistrib @[simp] -theorem prodSumDistrib_apply_left {α β γ} (a : α) (b : β) : +theorem prodSumDistrib_apply_left (a : α) (b : β) : prodSumDistrib α β γ (a, Sum.inl b) = Sum.inl (a, b) := rfl #align equiv.prod_sum_distrib_apply_left Equiv.prodSumDistrib_apply_left @[simp] -theorem prodSumDistrib_apply_right {α β γ} (a : α) (c : γ) : +theorem prodSumDistrib_apply_right (a : α) (c : γ) : prodSumDistrib α β γ (a, Sum.inr c) = Sum.inr (a, c) := rfl #align equiv.prod_sum_distrib_apply_right Equiv.prodSumDistrib_apply_right @[simp] -theorem prodSumDistrib_symm_apply_left {α β γ} (a : α × β) : +theorem prodSumDistrib_symm_apply_left (a : α × β) : (prodSumDistrib α β γ).symm (inl a) = (a.1, inl a.2) := rfl #align equiv.prod_sum_distrib_symm_apply_left Equiv.prodSumDistrib_symm_apply_left @[simp] -theorem prodSumDistrib_symm_apply_right {α β γ} (a : α × γ) : +theorem prodSumDistrib_symm_apply_right (a : α × γ) : (prodSumDistrib α β γ).symm (inr a) = (a.1, inr a.2) := rfl #align equiv.prod_sum_distrib_symm_apply_right Equiv.prodSumDistrib_symm_apply_right /-- An indexed sum of disjoint sums of types is equivalent to the sum of the indexed sums. -/ @[simps] -def sigmaSumDistrib {ι : Type _} (α β : ι → Type _) : - (Σi, Sum (α i) (β i)) ≃ Sum (Σi, α i) (Σi, β i) := +def sigmaSumDistrib (α β : ι → Type _) : + (Σ i, Sum (α i) (β i)) ≃ Sum (Σ i, α i) (Σ i, β i) := ⟨fun p => p.2.map (Sigma.mk p.1) (Sigma.mk p.1), Sum.elim (Sigma.map id fun _ => Sum.inl) (Sigma.map id fun _ => Sum.inr), fun p => by rcases p with ⟨i, a | b⟩ <;> rfl, fun p => by rcases p with (⟨i, a⟩ | ⟨i, b⟩) <;> rfl⟩ #align equiv.sigma_sum_distrib Equiv.sigmaSumDistrib -/-- The product of an indexed sum of types (formally, a `sigma`-type `Σ i, α i`) by a type `β` is +/-- The product of an indexed sum of types (formally, a `Sigma`-type `Σ i, α i`) by a type `β` is equivalent to the sum of products `Σ i, (α i × β)`. -/ -def sigmaProdDistrib {ι : Type _} (α : ι → Type _) (β : Type _) : (Σi, α i) × β ≃ Σi, α i × β := +def sigmaProdDistrib (α : ι → Type _) (β : Type _) : (Σ i, α i) × β ≃ Σ i, α i × β := ⟨fun p => ⟨p.1.1, (p.1.2, p.2)⟩, fun p => (⟨p.1, p.2.1⟩, p.2.2), fun p => by rcases p with ⟨⟨_, _⟩, _⟩ rfl, fun p => by @@ -962,7 +950,7 @@ def sigmaProdDistrib {ι : Type _} (α : ι → Type _) (β : Type _) : (Σi, α #align equiv.sigma_prod_distrib Equiv.sigmaProdDistrib /-- An equivalence that separates out the 0th fiber of `(Σ (n : ℕ), f n)`. -/ -def sigmaNatSucc (f : ℕ → Type u) : (Σn, f n) ≃ Sum (f 0) (Σn, f (n + 1)) := +def sigmaNatSucc (f : ℕ → Type u) : (Σ n, f n) ≃ Sum (f 0) (Σ n, f (n + 1)) := ⟨fun x => @Sigma.casesOn ℕ f (fun _ => Sum (f 0) (Σn, f (n + 1))) x fun n => @Nat.casesOn (fun i => f i → Sum (f 0) (Σn : ℕ, f (n + 1))) n (fun x : f 0 => Sum.inl x) @@ -971,18 +959,18 @@ def sigmaNatSucc (f : ℕ → Type u) : (Σn, f n) ≃ Sum (f 0) (Σn, f (n + 1) rintro (x | ⟨n, x⟩) <;> rfl⟩ #align equiv.sigma_nat_succ Equiv.sigmaNatSucc -/-- The product `bool × α` is equivalent to `α ⊕ α`. -/ +/-- The product `Bool × α` is equivalent to `α ⊕ α`. -/ @[simps] -def boolProdEquivSum (α : Type u) : Bool × α ≃ Sum α α where +def boolProdEquivSum (α) : Bool × α ≃ Sum α α where toFun p := cond p.1 (inr p.2) (inl p.2) invFun := Sum.elim (Prod.mk false) (Prod.mk true) left_inv := by rintro ⟨_ | _, _⟩ <;> rfl right_inv := by rintro (_ | _) <;> rfl #align equiv.bool_prod_equiv_sum Equiv.boolProdEquivSum -/-- The function type `bool → α` is equivalent to `α × α`. -/ +/-- The function type `Bool → α` is equivalent to `α × α`. -/ @[simps] -def boolArrowEquivProd (α : Type u) : (Bool → α) ≃ α × α where +def boolArrowEquivProd (α) : (Bool → α) ≃ α × α where toFun f := (f true, f false) invFun p b := cond b p.1 p.2 left_inv _ := funext <| Bool.forall_bool.2 ⟨rfl, rfl⟩ @@ -995,18 +983,18 @@ section open Sum Nat -/-- The set of natural numbers is equivalent to `ℕ ⊕ punit`. -/ -def natEquivNatSumPunit : ℕ ≃ Sum ℕ PUnit.{u + 1} where +/-- The set of natural numbers is equivalent to `ℕ ⊕ PUnit`. -/ +def natEquivNatSumPUnit : ℕ ≃ Sum ℕ PUnit where toFun n := Nat.casesOn n (inr PUnit.unit) inl invFun := Sum.elim Nat.succ fun _ => 0 left_inv n := by cases n <;> rfl right_inv := by rintro (_ | _ | _) <;> rfl -#align equiv.nat_equiv_nat_sum_punit Equiv.natEquivNatSumPunit +#align equiv.nat_equiv_nat_sum_punit Equiv.natEquivNatSumPUnit -/-- `ℕ ⊕ punit` is equivalent to `ℕ`. -/ -def natSumPunitEquivNat : Sum ℕ PUnit.{u + 1} ≃ ℕ := - natEquivNatSumPunit.symm -#align equiv.nat_sum_punit_equiv_nat Equiv.natSumPunitEquivNat +/-- `ℕ ⊕ Punit` is equivalent to `ℕ`. -/ +def natSumPUnitEquivNat : Sum ℕ PUnit ≃ ℕ := + natEquivNatSumPUnit.symm +#align equiv.nat_sum_punit_equiv_nat Equiv.natSumPUnitEquivNat /-- The type of integer numbers is equivalent to `ℕ ⊕ ℕ`. -/ def intEquivNatSumNat : ℤ ≃ Sum ℕ ℕ where @@ -1018,15 +1006,15 @@ def intEquivNatSumNat : ℤ ≃ Sum ℕ ℕ where end -/-- An equivalence between `α` and `β` generates an equivalence between `list α` and `list β`. -/ -def listEquivOfEquiv {α β : Type _} (e : α ≃ β) : List α ≃ List β where +/-- An equivalence between `α` and `β` generates an equivalence between `List α` and `List β`. -/ +def listEquivOfEquiv (e : α ≃ β) : List α ≃ List β where toFun := List.map e invFun := List.map e.symm left_inv l := by rw [List.map_map, e.symm_comp_self, List.map_id] right_inv l := by rw [List.map_map, e.self_comp_symm, List.map_id] #align equiv.list_equiv_of_equiv Equiv.listEquivOfEquiv -/-- If `α` is equivalent to `β`, then `unique α` is equivalent to `unique β`. -/ +/-- If `α` is equivalent to `β`, then `Unique α` is equivalent to `Unique β`. -/ def uniqueCongr (e : α ≃ β) : Unique α ≃ Unique β where toFun h := @Equiv.unique _ _ h e.symm invFun h := @Equiv.unique _ _ h e @@ -1034,7 +1022,7 @@ def uniqueCongr (e : α ≃ β) : Unique α ≃ Unique β where right_inv _ := Subsingleton.elim _ _ #align equiv.unique_congr Equiv.uniqueCongr -/-- If `α` is equivalent to `β`, then `is_empty α` is equivalent to `is_empty β`. -/ +/-- If `α` is equivalent to `β`, then `IsEmpty α` is equivalent to `IsEmpty β`. -/ theorem isEmpty_congr (e : α ≃ β) : IsEmpty α ↔ IsEmpty β := ⟨fun h => @Function.isEmpty _ _ h e.symm, fun h => @Function.isEmpty _ _ h e⟩ #align equiv.is_empty_congr Equiv.isEmpty_congr @@ -1049,7 +1037,7 @@ open Subtype /-- If `α` is equivalent to `β` and the predicates `p : α → Prop` and `q : β → Prop` are equivalent at corresponding points, then `{a // p a}` is equivalent to `{b // q b}`. -For the statement where `α = β`, that is, `e : perm α`, see `perm.subtype_perm`. -/ +For the statement where `α = β`, that is, `e : perm α`, see `Perm.subtypePerm`. -/ def subtypeEquiv {p : α → Prop} {q : β → Prop} (e : α ≃ β) (h : ∀ a, p a ↔ q (e a)) : { a : α // p a } ≃ { b : β // q b } where toFun a := ⟨e a, (h _).mp a.property⟩ @@ -1110,14 +1098,14 @@ def subtypeEquivOfSubtype' {p : α → Prop} (e : α ≃ β) : #align equiv.subtype_equiv_of_subtype' Equiv.subtypeEquivOfSubtype' /-- If two predicates are equal, then the corresponding subtypes are equivalent. -/ -def subtypeEquivProp {α : Type _} {p q : α → Prop} (h : p = q) : Subtype p ≃ Subtype q := +def subtypeEquivProp {p q : α → Prop} (h : p = q) : Subtype p ≃ Subtype q := subtypeEquiv (Equiv.refl α) fun _ => h ▸ Iff.rfl #align equiv.subtype_equiv_prop Equiv.subtypeEquivProp /-- A subtype of a subtype is equivalent to the subtype of elements satisfying both predicates. This version allows the “inner” predicate to depend on `h : p a`. -/ @[simps] -def subtypeSubtypeEquivSubtypeExists {α : Type u} (p : α → Prop) (q : Subtype p → Prop) : +def subtypeSubtypeEquivSubtypeExists (p : α → Prop) (q : Subtype p → Prop) : Subtype q ≃ { a : α // ∃ h : p a, q ⟨a, h⟩ } := ⟨fun a => ⟨a.1, a.1.2, by @@ -1128,7 +1116,7 @@ def subtypeSubtypeEquivSubtypeExists {α : Type u} (p : α → Prop) (q : Subtyp /-- A subtype of a subtype is equivalent to the subtype of elements satisfying both predicates. -/ @[simps] -def subtypeSubtypeEquivSubtypeInter {α : Type u} (p q : α → Prop) : +def subtypeSubtypeEquivSubtypeInter (p q : α → Prop) : { x : Subtype p // q x.1 } ≃ Subtype fun x => p x ∧ q x := (subtypeSubtypeEquivSubtypeExists p _).trans <| subtypeEquivRight fun x => @exists_prop (q x) (p x) @@ -1137,20 +1125,20 @@ def subtypeSubtypeEquivSubtypeInter {α : Type u} (p q : α → Prop) : /-- If the outer subtype has more restrictive predicate than the inner one, then we can drop the latter. -/ @[simps] -def subtypeSubtypeEquivSubtype {α : Type u} {p q : α → Prop} (h : ∀ {x}, q x → p x) : +def subtypeSubtypeEquivSubtype {p q : α → Prop} (h : ∀ {x}, q x → p x) : { x : Subtype p // q x.1 } ≃ Subtype q := (subtypeSubtypeEquivSubtypeInter p _).trans <| subtypeEquivRight fun _ => and_iff_right_of_imp h #align equiv.subtype_subtype_equiv_subtype Equiv.subtypeSubtypeEquivSubtype /-- If a proposition holds for all elements, then the subtype is equivalent to the original type. -/ -@[simps apply symmApply] -def subtypeUnivEquiv {α : Type u} {p : α → Prop} (h : ∀ x, p x) : Subtype p ≃ α := +@[simps apply symm_apply] +def subtypeUnivEquiv {p : α → Prop} (h : ∀ x, p x) : Subtype p ≃ α := ⟨fun x => x, fun x => ⟨x, h x⟩, fun _ => Subtype.eq rfl, fun _ => rfl⟩ #align equiv.subtype_univ_equiv Equiv.subtypeUnivEquiv /-- A subtype of a sigma-type is a sigma-type over a subtype. -/ -def subtypeSigmaEquiv {α : Type u} (p : α → Type v) (q : α → Prop) : { y : Sigma p // q y.1 } ≃ Σx : +def subtypeSigmaEquiv (p : α → Type v) (q : α → Prop) : { y : Sigma p // q y.1 } ≃ Σ x : Subtype q, p x.1 := ⟨fun x => ⟨⟨x.1.1, x.2⟩, x.1.2⟩, fun x => ⟨⟨x.1.1, x.2⟩, x.1.2⟩, fun _ => rfl, fun _ => rfl⟩ @@ -1158,25 +1146,24 @@ def subtypeSigmaEquiv {α : Type u} (p : α → Type v) (q : α → Prop) : { y /-- A sigma type over a subtype is equivalent to the sigma set over the original type, if the fiber is empty outside of the subset -/ -def sigmaSubtypeEquivOfSubset {α : Type u} (p : α → Type v) (q : α → Prop) (h : ∀ x, p x → q x) : - (Σx : Subtype q, p x) ≃ Σx : α, p x := +def sigmaSubtypeEquivOfSubset (p : α → Type v) (q : α → Prop) (h : ∀ x, p x → q x) : + (Σ x : Subtype q, p x) ≃ Σ x : α, p x := (subtypeSigmaEquiv p q).symm.trans <| subtypeUnivEquiv fun x => h x.1 x.2 #align equiv.sigma_subtype_equiv_of_subset Equiv.sigmaSubtypeEquivOfSubset /-- If a predicate `p : β → Prop` is true on the range of a map `f : α → β`, then `Σ y : {y // p y}, {x // f x = y}` is equivalent to `α`. -/ -def sigmaSubtypeFiberEquiv {α : Type u} {β : Type v} (f : α → β) (p : β → Prop) (h : ∀ x, p (f x)) : - (Σy : Subtype p, { x : α // f x = y }) ≃ α := +def sigmaSubtypeFiberEquiv {α β : Type _} (f : α → β) (p : β → Prop) (h : ∀ x, p (f x)) : + (Σ y : Subtype p, { x : α // f x = y }) ≃ α := calc _ ≃ Σy : β, { x : α // f x = y } := sigmaSubtypeEquivOfSubset _ p fun _ ⟨x, h'⟩ => h' ▸ h x _ ≃ α := sigmaFiberEquiv f - #align equiv.sigma_subtype_fiber_equiv Equiv.sigmaSubtypeFiberEquiv /-- If for each `x` we have `p x ↔ q (f x)`, then `Σ y : {y // q y}, f ⁻¹' {y}` is equivalent to `{x // p x}`. -/ -def sigmaSubtypeFiberEquivSubtype {α : Type u} {β : Type v} (f : α → β) {p : α → Prop} - {q : β → Prop} (h : ∀ x, p x ↔ q (f x)) : (Σy : Subtype q, { x : α // f x = y }) ≃ Subtype p := +def sigmaSubtypeFiberEquivSubtype {α β : Type _} (f : α → β) {p : α → Prop} {q : β → Prop} + (h : ∀ x, p x ↔ q (f x)) : (Σ y : Subtype q, { x : α // f x = y }) ≃ Subtype p := calc (Σy : Subtype q, { x : α // f x = y }) ≃ Σy : Subtype q, { x : Subtype p // Subtype.mk (f x) ((h x).1 x.2) = y } := by { @@ -1188,31 +1175,28 @@ def sigmaSubtypeFiberEquivSubtype {α : Type u} {β : Type v} (f : α → β) {p exact ⟨fun ⟨hp, h'⟩ => congr_arg Subtype.val h', fun h' => ⟨(h x).2 (h'.symm ▸ y.2), Subtype.eq h'⟩⟩ } _ ≃ Subtype p := sigmaFiberEquiv fun x : Subtype p => (⟨f x, (h x).1 x.property⟩ : Subtype q) - #align equiv.sigma_subtype_fiber_equiv_subtype Equiv.sigmaSubtypeFiberEquivSubtype -/-- A sigma type over an `option` is equivalent to the sigma set over the original type, +/-- A sigma type over an `Option` is equivalent to the sigma set over the original type, if the fiber is empty at none. -/ -def sigmaOptionEquivOfSome {α : Type u} (p : Option α → Type v) (h : p none → False) : - (Σx : Option α, p x) ≃ Σx : α, p (some x) := +def sigmaOptionEquivOfSome (p : Option α → Type v) (h : p none → False) : + (Σ x : Option α, p x) ≃ Σ x : α, p (some x) := haveI h' : ∀ x, p x → x.isSome := by intro x cases x · intro n exfalso exact h n - · intro _ exact rfl - (sigmaSubtypeEquivOfSubset _ _ h').symm.trans (sigmaCongrLeft' (optionIsSomeEquiv α)) #align equiv.sigma_option_equiv_of_some Equiv.sigmaOptionEquivOfSome -- ericr: this definition doesn't seem nice indentation-wise; is this just the Lean4 style? -/-- The `pi`-type `Π i, π i` is equivalent to the type of sections `f : ι → Σ i, π i` of the -`sigma` type such that for all `i` we have `(f i).fst = i`. -/ -def piEquivSubtypeSigma (ι : Type _) (π : ι → Type _) : - (∀ i, π i) ≃ { f : ι → Σi, π i // ∀ i, (f i).1 = i } := +/-- The `Pi`-type `∀ i, π i` is equivalent to the type of sections `f : ι → Σ i, π i` of the +`Sigma` type such that for all `i` we have `(f i).fst = i`. -/ +def piEquivSubtypeSigma (ι) (π : ι → Type _) : + (∀ i, π i) ≃ { f : ι → Σ i, π i // ∀ i, (f i).1 = i } := ⟨fun f => ⟨fun i => ⟨i, f i⟩, fun i => rfl⟩, fun f i => by rw [← f.2 i]; exact (f.1 i).2, fun f => funext fun i => rfl, @@ -1221,9 +1205,9 @@ def piEquivSubtypeSigma (ι : Type _) (π : ι → Type _) : Sigma.eq (hf i).symm <| eq_of_heq <| rec_heq_of_heq _ <| by simp⟩ #align equiv.pi_equiv_subtype_sigma Equiv.piEquivSubtypeSigma -/-- The set of functions `f : Π a, β a` such that for all `a` we have `p a (f a)` is equivalent -to the set of functions `Π a, {b : β a // p a b}`. -/ -def subtypePiEquivPi {α : Sort u} {β : α → Sort v} {p : ∀ a, β a → Prop} : +/-- The type of functions `f : ∀ a, β a` such that for all `a` we have `p a (f a)` is equivalent +to the type of functions `∀ a, {b : β a // p a b}`. -/ +def subtypePiEquivPi {β : α → Sort v} {p : ∀ a, β a → Prop} : { f : ∀ a, β a // ∀ a, p a (f a) } ≃ ∀ a, { b : β a // p a b } := ⟨fun f a => ⟨f.1 a, f.2 a⟩, fun f => ⟨fun a => (f a).1, fun a => (f a).2⟩, by rintro ⟨f, h⟩ @@ -1235,14 +1219,14 @@ def subtypePiEquivPi {α : Sort u} {β : α → Sort v} {p : ∀ a, β a → Pro /-- A subtype of a product defined by componentwise conditions is equivalent to a product of subtypes. -/ -def subtypeProdEquivProd {α : Type u} {β : Type v} {p : α → Prop} {q : β → Prop} : +def subtypeProdEquivProd {p : α → Prop} {q : β → Prop} : { c : α × β // p c.1 ∧ q c.2 } ≃ { a // p a } × { b // q b } := ⟨fun x => ⟨⟨x.1.1, x.2.1⟩, ⟨x.1.2, x.2.2⟩⟩, fun x => ⟨⟨x.1.1, x.2.1⟩, ⟨x.1.2, x.2.2⟩⟩, fun ⟨⟨_, _⟩, ⟨_, _⟩⟩ => rfl, fun ⟨⟨_, _⟩, ⟨_, _⟩⟩ => rfl⟩ #align equiv.subtype_prod_equiv_prod Equiv.subtypeProdEquivProd -/-- A subtype of a `prod` is equivalent to a sigma type whose fibers are subtypes. -/ -def subtypeProdEquivSigmaSubtype {α β : Type _} (p : α → β → Prop) : +/-- A subtype of a `Prod` is equivalent to a sigma type whose fibers are subtypes. -/ +def subtypeProdEquivSigmaSubtype (p : α → β → Prop) : { x : α × β // p x.1 x.2 } ≃ Σa, { b : β // p a b } where toFun x := ⟨x.1.1, x.1.2, x.property⟩ invFun x := ⟨⟨x.1, x.2⟩, x.2.property⟩ @@ -1250,7 +1234,7 @@ def subtypeProdEquivSigmaSubtype {α β : Type _} (p : α → β → Prop) : right_inv := fun ⟨a, b, pab⟩ => rfl #align equiv.subtype_prod_equiv_sigma_subtype Equiv.subtypeProdEquivSigmaSubtype -/-- The type `Π (i : α), β i` can be split as a product by separating the indices in `α` +/-- The type `∀ (i : α), β i` can be split as a product by separating the indices in `α` depending on whether they satisfy a predicate `p` or not. -/ @[simps] def piEquivPiSubtypeProd {α : Type _} (p : α → Prop) (β : α → Type _) [DecidablePred p] : @@ -1267,7 +1251,6 @@ def piEquivPiSubtypeProd {α : Type _} (p : α → Prop) (β : α → Type _) [D ext x by_cases h:p x <;> · simp only [h, dif_neg, dif_pos, not_false_iff] - #align equiv.pi_equiv_pi_subtype_prod Equiv.piEquivPiSubtypeProd /-- A product of types can be split as the binary product of one of the types and the product @@ -1286,7 +1269,6 @@ def piSplitAt {α : Type _} [DecidableEq α] (i : α) (β : α → Type _) : split_ifs with h · subst h; rfl · rfl - #align equiv.pi_split_at Equiv.piSplitAt /-- A product of copies of a type can be split as the binary product of one copy and the product @@ -1301,7 +1283,7 @@ end section subtypeEquivCodomain -variable {X : Type _} {Y : Type _} [DecidableEq X] {x : X} +variable [DecidableEq X] {x : X} /-- The type of all functions `X → Y` with prescribed values for all `x' ≠ x` is equivalent to the codomain `Y`. -/ @@ -1386,12 +1368,12 @@ section variable {α' β' : Type _} (e : Perm α') {p : β' → Prop} [DecidablePred p] (f : α' ≃ Subtype p) -/-- Extend the domain of `e : equiv.perm α` to one that is over `β` via `f : α → subtype p`, +/-- Extend the domain of `e : Equiv.Perm α` to one that is over `β` via `f : α → Subtype p`, where `p : β → Prop`, permuting only the `b : β` that satisfy `p b`. This can be used to extend the domain across a function `f : α → β`, -keeping everything outside of `set.range f` fixed. For this use-case `equiv` given by `f` can -be constructed by `equiv.of_left_inverse'` or `equiv.of_left_inverse` when there is a known -inverse, or `equiv.of_injective` in the general case.`. +keeping everything outside of `Set.range f` fixed. For this use-case `Equiv` given by `f` can +be constructed by `Equiv.of_left_inverse'` or `Equiv.of_left_inverse` when there is a known +inverse, or `Equiv.of_injective` in the general case.`. -/ def Perm.extendDomain : Perm β' := (permCongr f e).subtypeCongr (Equiv.refl _) @@ -1470,7 +1452,7 @@ section Swap variable [DecidableEq α] -/-- A helper function for `equiv.swap`. -/ +/-- A helper function for `Equiv.swap`. -/ def swapCore (a b r : α) : α := if r = a then b else if r = b then a else r #align equiv.swap_core Equiv.swapCore @@ -1663,7 +1645,7 @@ end Equiv namespace Function.Involutive -/-- Convert an involutive function `f` to a permutation with `to_fun = inv_fun = f`. -/ +/-- Convert an involutive function `f` to a permutation with `toFun = invFun = f`. -/ def toPerm (f : α → α) (h : Involutive f) : Equiv.Perm α := ⟨f, f, h.leftInverse, h.rightInverse⟩ #align function.involutive.to_perm Function.Involutive.toPerm @@ -1688,7 +1670,7 @@ theorem PLift.eq_up_iff_down_eq {x : PLift α} {y : α} : x = PLift.up y ↔ x.d Equiv.plift.eq_symm_apply #align plift.eq_up_iff_down_eq PLift.eq_up_iff_down_eq -theorem Function.Injective.map_swap {α β : Type _} [DecidableEq α] [DecidableEq β] {f : α → β} +theorem Function.Injective.map_swap [DecidableEq α] [DecidableEq β] {f : α → β} (hf : Function.Injective f) (x y z : α) : f (Equiv.swap x y z) = Equiv.swap (f x) (f y) (f z) := by conv_rhs => rw [Equiv.swap_apply_def] @@ -1807,7 +1789,7 @@ end section BinaryOp -variable {α₁ β₁ : Type _} (e : α₁ ≃ β₁) (f : α₁ → α₁ → α₁) +variable (e : α₁ ≃ β₁) (f : α₁ → α₁ → α₁) theorem semiconj_conj (f : α₁ → α₁) : Semiconj e f (e.conj f) := fun x => by simp #align equiv.semiconj_conj Equiv.semiconj_conj @@ -1859,7 +1841,7 @@ theorem Function.Injective.swap_comp #align function.injective.swap_comp Function.Injective.swap_comp /-- If `α` is a subsingleton, then it is equivalent to `α × α`. -/ -def subsingletonProdSelfEquiv {α : Type _} [Subsingleton α] : α × α ≃ α where +def subsingletonProdSelfEquiv [Subsingleton α] : α × α ≃ α where toFun p := p.1 invFun a := (a, a) left_inv _ := Subsingleton.elim _ _ @@ -1876,13 +1858,13 @@ def equivOfSubsingletonOfSubsingleton [Subsingleton α] [Subsingleton β] (f : right_inv _ := Subsingleton.elim _ _ #align equiv_of_subsingleton_of_subsingleton equivOfSubsingletonOfSubsingleton -/-- A nonempty subsingleton type is (noncomputably) equivalent to `punit`. -/ -noncomputable def Equiv.punitOfNonemptyOfSubsingleton - {α : Sort _} [h : Nonempty α] [Subsingleton α] : α ≃ PUnit.{v} := +/-- A nonempty subsingleton type is (noncomputably) equivalent to `PUnit`. -/ +noncomputable def Equiv.punitOfNonemptyOfSubsingleton [h : Nonempty α] [Subsingleton α] : + α ≃ PUnit := equivOfSubsingletonOfSubsingleton (fun _ => PUnit.unit) fun _ => h.some #align equiv.punit_of_nonempty_of_subsingleton Equiv.punitOfNonemptyOfSubsingleton -/-- `unique (unique α)` is equivalent to `unique α`. -/ +/-- `Unique (Unique α)` is equivalent to `Unique α`. -/ def uniqueUniqueEquiv : Unique (Unique α) ≃ Unique α := equivOfSubsingletonOfSubsingleton (fun h => h.default) fun h => { default := h, uniq := fun _ => Subsingleton.elim _ _ } @@ -1890,13 +1872,13 @@ def uniqueUniqueEquiv : Unique (Unique α) ≃ Unique α := namespace Function -theorem update_comp_equiv {α β α' : Sort _} [DecidableEq α'] [DecidableEq α] (f : α → β) +theorem update_comp_equiv [DecidableEq α'] [DecidableEq α] (f : α → β) (g : α' ≃ α) (a : α) (v : β) : update f a v ∘ g = update (f ∘ g) (g.symm a) v := by rw [← update_comp_eq_of_injective _ g.injective, g.apply_symm_apply] #align function.update_comp_equiv Function.update_comp_equiv -theorem update_apply_equiv_apply {α β α' : Sort _} [DecidableEq α'] [DecidableEq α] (f : α → β) +theorem update_apply_equiv_apply [DecidableEq α'] [DecidableEq α] (f : α → β) (g : α' ≃ α) (a : α) (v : β) (a' : α') : update f a v (g a') = update (f ∘ g) (g.symm a) v a' := congr_fun (update_comp_equiv f g a v) a' #align function.update_apply_equiv_apply Function.update_apply_equiv_apply @@ -1919,7 +1901,7 @@ theorem piCongrLeft'_update [DecidableEq α] [DecidableEq β] (P : α → Sort _ #align function.Pi_congr_left'_update Function.piCongrLeft'_update -theorem piCongrLeft''_symm_update [DecidableEq α] [DecidableEq β] (P : α → Sort _) (e : α ≃ β) +theorem piCongrLeft'_symm_update [DecidableEq α] [DecidableEq β] (P : α → Sort _) (e : α ≃ β) (f : ∀ b, P (e.symm b)) (b : β) (x : P (e.symm b)) : (e.piCongrLeft' P).symm (update f b x) = update ((e.piCongrLeft' P).symm f) (e.symm b) x := by simp [(e.piCongrLeft' P).symm_apply_eq, piCongrLeft'_update] From 70f7b0e238aec016a9c0ae0a582da50b1f5d9356 Mon Sep 17 00:00:00 2001 From: Arien Malec Date: Tue, 22 Nov 2022 14:03:28 -0800 Subject: [PATCH 073/127] addresed a theorem failure --- Mathlib/Logic/Equiv/Option.lean | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/Mathlib/Logic/Equiv/Option.lean b/Mathlib/Logic/Equiv/Option.lean index def27994411e3..a042396ff2b91 100644 --- a/Mathlib/Logic/Equiv/Option.lean +++ b/Mathlib/Logic/Equiv/Option.lean @@ -143,8 +143,7 @@ theorem some_remove_none_iff {x : α} : some (removeNone e x) = e none ↔ e.sym have h1 := congr_arg e.symm h rw [symm_apply_apply] at h1 simp only [false_iff_iff, apply_eq_iff_eq] - simp [h1] - + simp [h1, apply_eq_iff_eq] #align equiv.some_remove_none_iff Equiv.some_remove_none_iff From 525bbad2bf4d8d81b84df479e54813e956cc1089 Mon Sep 17 00:00:00 2001 From: Arien Malec Date: Tue, 22 Nov 2022 14:14:27 -0800 Subject: [PATCH 074/127] naming fix & style fixes --- Mathlib/Logic/Equiv/Option.lean | 59 ++++++++++++++++++++++----------- 1 file changed, 40 insertions(+), 19 deletions(-) diff --git a/Mathlib/Logic/Equiv/Option.lean b/Mathlib/Logic/Equiv/Option.lean index a042396ff2b91..66f9d00b4b5b9 100644 --- a/Mathlib/Logic/Equiv/Option.lean +++ b/Mathlib/Logic/Equiv/Option.lean @@ -160,13 +160,16 @@ theorem option_congr_injective : Function.Injective (optionCongr : α ≃ β → /-- Equivalences between `Option α` and `β` that send `none` to `x` are equivalent to equivalences between `α` and `{y : β // y ≠ x}`. -/ -def option_subtype [DecidableEq β] (x : β) : { e : Option α ≃ β // e none = x } ≃ (α ≃ { y : β // y ≠ x }) where +def option_subtype [DecidableEq β] (x : β) : + { e : Option α ≃ β // e none = x } ≃ (α ≃ { y : β // y ≠ x }) where toFun e := - { toFun := fun a => ⟨e a, ((EquivLike.injective _).ne_iff' e.property).2 (some_ne_none _)⟩, + { toFun := + fun a => ⟨e a, ((EquivLike.injective _).ne_iff' e.property).2 (some_ne_none _)⟩, invFun := fun b => - get - (ne_none_iff_is_some.1 - (((EquivLike.injective _).ne_iff' ((apply_eq_iff_eq_symm_apply _).1 e.property).symm).2 b.property)), + get _ + (ne_none_iff_isSome.1 + (((EquivLike.injective _).ne_iff' + ((apply_eq_iff_eq_symm_apply _).1 e.property).symm).2 b.property)), left_inv := fun a => by rw [← some_inj, some_get, ← coe_def] exact symm_apply_apply (e : Option α ≃ β) a, @@ -175,12 +178,14 @@ def option_subtype [DecidableEq β] (x : β) : { e : Option α ≃ β // e none simp exact apply_symm_apply _ _ } invFun e := - ⟨{ toFun := fun a => casesOn' a x (coe ∘ e), invFun := fun b => if h : b = x then none else e.symm ⟨b, h⟩, + ⟨{ toFun := fun a => casesOn' a x (coe ∘ e), + invFun := fun b => if h : b = x then none else e.symm ⟨b, h⟩, left_inv := fun a => by cases a · simp - simp only [cases_on'_some, Function.comp_apply, Subtype.coe_eta, symm_apply_apply, dite_eq_ite] + simp only [cases_on'_some, Function.comp_apply, Subtype.coe_eta, + symm_apply_apply, dite_eq_ite] exact if_neg (e a).property, right_inv := fun b => by by_cases h : b = x <;> simp [h] }, rfl⟩ @@ -197,46 +202,62 @@ def option_subtype [DecidableEq β] (x : β) : { e : Option α ≃ β // e none #align equiv.option_subtype Equiv.option_subtype @[simp] -theorem option_subtype_apply_apply [DecidableEq β] (x : β) (e : { e : Option α ≃ β // e none = x }) (a : α) (h) : - option_subtype x e a = ⟨(e : Option α ≃ β) a, h⟩ := +theorem option_subtype_apply_apply + [DecidableEq β] (x : β) + (e : { e : Option α ≃ β // e none = x }) + (a : α) + (h) : option_subtype x e a = ⟨(e : Option α ≃ β) a, h⟩ := rfl #align equiv.option_subtype_apply_apply Equiv.option_subtype_apply_apply @[simp] -theorem coe_option_subtype_apply_apply [DecidableEq β] (x : β) (e : { e : Option α ≃ β // e none = x }) (a : α) : - ↑(option_subtype x e a) = (e : Option α ≃ β) a := +theorem coe_option_subtype_apply_apply + [DecidableEq β] (x : β) + (e : { e : Option α ≃ β // e none = x }) + (a : α) : ↑(option_subtype x e a) = (e : Option α ≃ β) a := rfl #align equiv.coe_option_subtype_apply_apply Equiv.coe_option_subtype_apply_apply @[simp] -theorem option_subtype_apply_symm_apply [DecidableEq β] (x : β) (e : { e : Option α ≃ β // e none = x }) +theorem option_subtype_apply_symm_apply + [DecidableEq β] (x : β) + (e : { e : Option α ≃ β // e none = x }) (b : { y : β // y ≠ x }) : ↑((option_subtype x e).symm b) = (e : Option α ≃ β).symm b := by dsimp only [option_subtype] simp #align equiv.option_subtype_apply_symm_apply Equiv.option_subtype_apply_symm_apply @[simp] -theorem option_subtype_symm_apply_apply_coe [DecidableEq β] (x : β) (e : α ≃ { y : β // y ≠ x }) (a : α) : - (option_subtype x).symm e a = e a := +theorem option_subtype_symm_apply_apply_coe + [DecidableEq β] + (x : β) + (e : α ≃ { y : β // y ≠ x }) + (a : α) : (option_subtype x).symm e a = e a := rfl #align equiv.option_subtype_symm_apply_apply_coe Equiv.option_subtype_symm_apply_apply_coe @[simp] -theorem option_subtype_symm_apply_apply_some [DecidableEq β] (x : β) (e : α ≃ { y : β // y ≠ x }) (a : α) : - (option_subtype x).symm e (some a) = e a := +theorem option_subtype_symm_apply_apply_some + [DecidableEq β] + (x : β) + (e : α ≃ { y : β // y ≠ x }) + (a : α) : (option_subtype x).symm e (some a) = e a := rfl #align equiv.option_subtype_symm_apply_apply_some Equiv.option_subtype_symm_apply_apply_some @[simp] -theorem option_subtype_symm_apply_apply_none [DecidableEq β] (x : β) (e : α ≃ { y : β // y ≠ x }) : - (option_subtype x).symm e none = x := +theorem option_subtype_symm_apply_apply_none + [DecidableEq β] + (x : β) + (e : α ≃ { y : β // y ≠ x }) : (option_subtype x).symm e none = x := rfl #align equiv.option_subtype_symm_apply_apply_none Equiv.option_subtype_symm_apply_apply_none @[simp] theorem option_subtype_symm_apply_symm_apply [DecidableEq β] (x : β) (e : α ≃ { y : β // y ≠ x }) (b : { y : β // y ≠ x }) : ((option_subtype x).symm e : Option α ≃ β).symm b = e.symm b := by - simp only [option_subtype, coe_fn_symm_mk, Subtype.coe_mk, Subtype.coe_eta, dite_eq_ite, ite_eq_right_iff] + simp only [option_subtype, coe_fn_symm_mk, Subtype.coe_mk, + Subtype.coe_eta, dite_eq_ite, ite_eq_right_iff] exact fun h => False.elim (b.property h) #align equiv.option_subtype_symm_apply_symm_apply Equiv.option_subtype_symm_apply_symm_apply From 615cc29826b5968ccd90fdc123c4a556673fd66e Mon Sep 17 00:00:00 2001 From: Arien Malec Date: Tue, 22 Nov 2022 14:18:36 -0800 Subject: [PATCH 075/127] doc fix for linting --- Mathlib/Logic/Equiv/Option.lean | 3 +++ 1 file changed, 3 insertions(+) diff --git a/Mathlib/Logic/Equiv/Option.lean b/Mathlib/Logic/Equiv/Option.lean index 66f9d00b4b5b9..852025c1089d1 100644 --- a/Mathlib/Logic/Equiv/Option.lean +++ b/Mathlib/Logic/Equiv/Option.lean @@ -66,6 +66,9 @@ section RemoveNone variable (e : Option α ≃ Option β) +/-- If we have a value on one side of an `Equiv` of `Option` + we also have a value on the other side of the equivalence +-/ def remove_none_aux (x : α) : β := if h : (e (some x)).isSome then Option.get _ h else From 83e63fb15fb43eb9328d26ad5838eb78f17f424b Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?David=20W=C3=A4rn?= Date: Tue, 22 Nov 2022 23:43:47 +0100 Subject: [PATCH 076/127] remove bad (?) simp lemmas --- Mathlib/Logic/Equiv/Basic.lean | 40 +++++++++++----------------------- 1 file changed, 13 insertions(+), 27 deletions(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 266707e85fc51..98114de2cb2c4 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -290,7 +290,6 @@ theorem sumCongr_apply (ea : Equiv.Perm α) (eb : Equiv.Perm β) (x : Sum α β) -- porting note: it seems the general theorem about `Equiv` is now applied, so there's no need -- to have this version also have `@[simp]`. Similarly for below. -@[simp] theorem sumCongr_trans (e : Equiv.Perm α) (f : Equiv.Perm β) (g : Equiv.Perm α) (h : Equiv.Perm β) : (sumCongr e f).trans (sumCongr g h) = sumCongr (e.trans g) (f.trans h) := Equiv.sumCongr_trans e f g h @@ -404,11 +403,7 @@ theorem optionEquivSumPUnit_none : optionEquivSumPUnit α none = Sum.inr PUnit.u theorem optionEquivSumPUnit_some (a) : optionEquivSumPUnit α (some a) = Sum.inl a := rfl #align equiv.option_equiv_sum_punit_some Equiv.optionEquivSumPUnit_some - -@[simp] -theorem optionEquivSumPUnit_coe (a : α) : optionEquivSumPUnit α a = Sum.inl a := - rfl -#align equiv.option_equiv_sum_punit_coe Equiv.optionEquivSumPUnit_coe +#align equiv.option_equiv_sum_punit_coe Equiv.optionEquivSumPUnit_some @[simp] theorem optionEquivSumPUnit_symm_inl (a) : (optionEquivSumPUnit α).symm (Sum.inl a) = a := @@ -495,13 +490,14 @@ theorem sumCompl_apply_inr (p : α → Prop) [DecidablePred p] (x : { a // ¬p a rfl #align equiv.sum_compl_apply_inr Equiv.sumCompl_apply_inr -@[simp] +-- porting note: this had its `@[simp]` attribute removed, because the simpNF linter complained and +-- it doesn't seem like a good simp lemma? similarly for `sumCompl_apply_symm_of_neg` and +-- `Perm.subtypeCongr.{left,right}_apply` theorem sumCompl_apply_symm_of_pos (p : α → Prop) [DecidablePred p] (a : α) (h : p a) : (sumCompl p).symm a = Sum.inl ⟨a, h⟩ := dif_pos h #align equiv.sum_compl_apply_symm_of_pos Equiv.sumCompl_apply_symm_of_pos -@[simp] theorem sumCompl_apply_symm_of_neg (p : α → Prop) [DecidablePred p] (a : α) (h : ¬p a) : (sumCompl p).symm a = Sum.inr ⟨a, h⟩ := dif_neg h @@ -528,10 +524,10 @@ def Perm.subtypeCongr : Equiv.Perm ε := theorem Perm.subtypeCongr.apply (a : ε) : ep.subtypeCongr en a = if h : p a then (ep ⟨a, h⟩ : ε) else en ⟨a, h⟩ := by - by_cases h : p a <;> simp [Perm.subtypeCongr, h] + by_cases h : p a <;> + simp [Perm.subtypeCongr, h, sumCompl_apply_symm_of_pos, sumCompl_apply_symm_of_neg] #align equiv.perm.subtype_congr.apply Equiv.Perm.subtypeCongr.apply -@[simp] theorem Perm.subtypeCongr.left_apply {a : ε} (h : p a) : ep.subtypeCongr en a = ep ⟨a, h⟩ := by simp [Perm.subtypeCongr.apply, h] #align equiv.perm.subtype_congr.left_apply Equiv.Perm.subtypeCongr.left_apply @@ -541,7 +537,6 @@ theorem Perm.subtypeCongr.left_apply_subtype (a : { a // p a }) : ep.subtypeCong Perm.subtypeCongr.left_apply ep en a.property #align equiv.perm.subtype_congr.left_apply_subtype Equiv.Perm.subtypeCongr.left_apply_subtype -@[simp] theorem Perm.subtypeCongr.right_apply {a : ε} (h : ¬p a) : ep.subtypeCongr en a = en ⟨a, h⟩ := by simp [Perm.subtypeCongr.apply, h] #align equiv.perm.subtype_congr.right_apply Equiv.Perm.subtypeCongr.right_apply @@ -555,19 +550,15 @@ theorem Perm.subtypeCongr.right_apply_subtype (a : { a // ¬p a }) : ep.subtypeC theorem Perm.subtypeCongr.refl : Perm.subtypeCongr (Equiv.refl { a // p a }) (Equiv.refl { a // ¬p a }) = Equiv.refl ε := by ext x - by_cases h:p x <;> simp [h] + by_cases h:p x <;> simp [h, subtypeCongr.left_apply, subtypeCongr.right_apply] #align equiv.perm.subtype_congr.refl Equiv.Perm.subtypeCongr.refl @[simp] theorem Perm.subtypeCongr.symm : (ep.subtypeCongr en).symm = Perm.subtypeCongr ep.symm en.symm := by ext x by_cases h:p x - · have : p (ep.symm ⟨x, h⟩) := Subtype.property _ - simp [Perm.subtypeCongr.apply, h, symm_apply_eq, this] - - · have : ¬p (en.symm ⟨x, h⟩) := Subtype.property (en.symm _) - simp [Perm.subtypeCongr.apply, h, symm_apply_eq, this] - + · simp [Perm.subtypeCongr.apply, h, symm_apply_eq] + · simp [Perm.subtypeCongr.apply, h, symm_apply_eq] #align equiv.perm.subtype_congr.symm Equiv.Perm.subtypeCongr.symm @[simp] @@ -576,12 +567,8 @@ theorem Perm.subtypeCongr.trans : = Perm.subtypeCongr (ep.trans ep') (en.trans en') := by ext x by_cases h:p x - · have : p (ep ⟨x, h⟩) := Subtype.property _ - simp [Perm.subtypeCongr.apply, h, this] - - · have : ¬p (en ⟨x, h⟩) := Subtype.property (en _) - simp [Perm.subtypeCongr.apply, h, symm_apply_eq, this] - + · simp [Perm.subtypeCongr.apply, h] + · simp [Perm.subtypeCongr.apply, h, symm_apply_eq] #align equiv.perm.subtype_congr.trans Equiv.Perm.subtypeCongr.trans end sumCompl @@ -1386,11 +1373,11 @@ theorem Perm.extendDomain_apply_image (a : α') : e.extendDomain f (f a) = f (e theorem Perm.extendDomain_apply_subtype {b : β'} (h : p b) : e.extendDomain f b = f (e (f.symm ⟨b, h⟩)) := by - simp [Perm.extendDomain, h] + simp [Perm.extendDomain, h, subtypeCongr.left_apply] #align equiv.perm.extend_domain_apply_subtype Equiv.Perm.extendDomain_apply_subtype theorem Perm.extendDomain_apply_not_subtype {b : β'} (h : ¬p b) : e.extendDomain f b = b := by - simp [Perm.extendDomain, h] + simp [Perm.extendDomain, h, subtypeCongr.right_apply] #align equiv.perm.extend_domain_apply_not_subtype Equiv.Perm.extendDomain_apply_not_subtype @[simp] @@ -1908,4 +1895,3 @@ theorem piCongrLeft'_symm_update [DecidableEq α] [DecidableEq β] (P : α → S #align function.Pi_congr_left'_symm_update Function.piCongrLeft'_symm_update end Function -#lint From 1c482e33637d36087a8a40c7c9e7f476a29825fa Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 23 Nov 2022 09:45:15 +1100 Subject: [PATCH 077/127] fix one lint error --- Mathlib/Logic/Equiv/Basic.lean | 5 ++--- 1 file changed, 2 insertions(+), 3 deletions(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 266707e85fc51..28998a024405b 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -290,7 +290,6 @@ theorem sumCongr_apply (ea : Equiv.Perm α) (eb : Equiv.Perm β) (x : Sum α β) -- porting note: it seems the general theorem about `Equiv` is now applied, so there's no need -- to have this version also have `@[simp]`. Similarly for below. -@[simp] theorem sumCongr_trans (e : Equiv.Perm α) (f : Equiv.Perm β) (g : Equiv.Perm α) (h : Equiv.Perm β) : (sumCongr e f).trans (sumCongr g h) = sumCongr (e.trans g) (f.trans h) := Equiv.sumCongr_trans e f g h @@ -502,8 +501,8 @@ theorem sumCompl_apply_symm_of_pos (p : α → Prop) [DecidablePred p] (a : α) #align equiv.sum_compl_apply_symm_of_pos Equiv.sumCompl_apply_symm_of_pos @[simp] -theorem sumCompl_apply_symm_of_neg (p : α → Prop) [DecidablePred p] (a : α) - (h : ¬p a) : (sumCompl p).symm a = Sum.inr ⟨a, h⟩ := +theorem sumCompl_apply_symm_of_neg (p : α → Prop) [DecidablePred p] (a : α) (h : ¬p a) : + (sumCompl p).symm a = Sum.inr ⟨a, h⟩ := dif_neg h #align equiv.sum_compl_apply_symm_of_neg Equiv.sumCompl_apply_symm_of_neg From a01d23e9f6fa1c16f1beaa085f4068943f426538 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 23 Nov 2022 09:45:30 +1100 Subject: [PATCH 078/127] remove #lint --- Mathlib/Logic/Equiv/Basic.lean | 1 - 1 file changed, 1 deletion(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 28998a024405b..6b28c4c192b7c 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -1907,4 +1907,3 @@ theorem piCongrLeft'_symm_update [DecidableEq α] [DecidableEq β] (P : α → S #align function.Pi_congr_left'_symm_update Function.piCongrLeft'_symm_update end Function -#lint From 097865a8196990429f359990c8854ea614e837aa Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 23 Nov 2022 10:28:20 +1100 Subject: [PATCH 079/127] don't making the coercion twice --- Mathlib/Logic/Equiv/Defs.lean | 12 +++++++++--- 1 file changed, 9 insertions(+), 3 deletions(-) diff --git a/Mathlib/Logic/Equiv/Defs.lean b/Mathlib/Logic/Equiv/Defs.lean index 3ba9bbccf93a1..42b5c4a4fb979 100644 --- a/Mathlib/Logic/Equiv/Defs.lean +++ b/Mathlib/Logic/Equiv/Defs.lean @@ -92,8 +92,6 @@ instance : EquivLike (α ≃ β) α β where right_inv := right_inv coe_injective' e₁ e₂ h₁ h₂ := by cases e₁; cases e₂; congr -instance : CoeFun (α ≃ β) fun _ => α → β := ⟨toFun⟩ - /-- The map `(r ≃ s) → (r → s)` is injective. -/ theorem coe_fn_injective : @Function.Injective (α ≃ β) (α → β) (fun e => e) := FunLike.coe_injective' @@ -132,10 +130,18 @@ instance inhabited' : Inhabited (α ≃ α) := ⟨Equiv.refl α⟩ protected def symm (e : α ≃ β) : β ≃ α := ⟨e.invFun, e.toFun, e.right_inv, e.left_inv⟩ /-- See Note [custom simps projection] -/ +def Simps.apply (e : α ≃ β) : α → β := e def Simps.symm_apply (e : α ≃ β) : β → α := e.symm initialize_simps_projections Equiv (toFun → apply, invFun → symm_apply) +-- Porting note: +-- Added these lemmas as restatements of `left_inv` and `right_inv`, +-- which use the coercions. +-- We might even consider switching the names, and having these as a public API. +theorem left_inv' (e : α ≃ β) : Function.LeftInverse e.symm e := e.left_inv +theorem right_inv' (e : α ≃ β) : Function.RightInverse e.symm e := e.right_inv + /-- Composition of equivalences `e₁ : α ≃ β` and `e₂ : β ≃ γ`. -/ -- Porting note: `trans` attribute rejects this lemma because of implicit arguments. -- @[trans] @@ -662,7 +668,7 @@ end Perm @[simps apply] def sigmaCongrLeft {β : α₂ → Sort _} (e : α₁ ≃ α₂) : (Σ a : α₁, β (e a)) ≃ Σ a : α₂, β a where toFun a := ⟨e a.1, a.2⟩ - invFun a := ⟨e.symm a.1, (e.right_inv a.1).symm ▸ a.2⟩ + invFun a := ⟨e.symm a.1, (e.right_inv' a.1).symm ▸ a.2⟩ -- porting note: this was a pretty gnarly match already, and it got worse after porting left_inv := fun ⟨a, b⟩ => match (motive := ∀ a' (h : a' = a), Sigma.mk _ (congr_arg e h.symm ▸ b) = ⟨a, b⟩) From 84cfc0e354524ab67fa0ad97ffb9c4563c6fcda5 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 23 Nov 2022 10:29:10 +1100 Subject: [PATCH 080/127] merge --- Mathlib/Logic/Equiv/Basic.lean | 40 +++++++++++++++++++++++----------- 1 file changed, 27 insertions(+), 13 deletions(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 138243bed3939..f862dd267fc80 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -290,6 +290,7 @@ theorem sumCongr_apply (ea : Equiv.Perm α) (eb : Equiv.Perm β) (x : Sum α β) -- porting note: it seems the general theorem about `Equiv` is now applied, so there's no need -- to have this version also have `@[simp]`. Similarly for below. +@[simp] theorem sumCongr_trans (e : Equiv.Perm α) (f : Equiv.Perm β) (g : Equiv.Perm α) (h : Equiv.Perm β) : (sumCongr e f).trans (sumCongr g h) = sumCongr (e.trans g) (f.trans h) := Equiv.sumCongr_trans e f g h @@ -403,7 +404,11 @@ theorem optionEquivSumPUnit_none : optionEquivSumPUnit α none = Sum.inr PUnit.u theorem optionEquivSumPUnit_some (a) : optionEquivSumPUnit α (some a) = Sum.inl a := rfl #align equiv.option_equiv_sum_punit_some Equiv.optionEquivSumPUnit_some -#align equiv.option_equiv_sum_punit_coe Equiv.optionEquivSumPUnit_some + +@[simp] +theorem optionEquivSumPUnit_coe (a : α) : optionEquivSumPUnit α a = Sum.inl a := + rfl +#align equiv.option_equiv_sum_punit_coe Equiv.optionEquivSumPUnit_coe @[simp] theorem optionEquivSumPUnit_symm_inl (a) : (optionEquivSumPUnit α).symm (Sum.inl a) = a := @@ -490,14 +495,13 @@ theorem sumCompl_apply_inr (p : α → Prop) [DecidablePred p] (x : { a // ¬p a rfl #align equiv.sum_compl_apply_inr Equiv.sumCompl_apply_inr --- porting note: this had its `@[simp]` attribute removed, because the simpNF linter complained and --- it doesn't seem like a good simp lemma? similarly for `sumCompl_apply_symm_of_neg` and --- `Perm.subtypeCongr.{left,right}_apply` +@[simp] theorem sumCompl_apply_symm_of_pos (p : α → Prop) [DecidablePred p] (a : α) (h : p a) : (sumCompl p).symm a = Sum.inl ⟨a, h⟩ := dif_pos h #align equiv.sum_compl_apply_symm_of_pos Equiv.sumCompl_apply_symm_of_pos +@[simp] theorem sumCompl_apply_symm_of_neg (p : α → Prop) [DecidablePred p] (a : α) (h : ¬p a) : (sumCompl p).symm a = Sum.inr ⟨a, h⟩ := dif_neg h @@ -524,10 +528,10 @@ def Perm.subtypeCongr : Equiv.Perm ε := theorem Perm.subtypeCongr.apply (a : ε) : ep.subtypeCongr en a = if h : p a then (ep ⟨a, h⟩ : ε) else en ⟨a, h⟩ := by - by_cases h : p a <;> - simp [Perm.subtypeCongr, h, sumCompl_apply_symm_of_pos, sumCompl_apply_symm_of_neg] + by_cases h : p a <;> simp [Perm.subtypeCongr, h] #align equiv.perm.subtype_congr.apply Equiv.Perm.subtypeCongr.apply +@[simp] theorem Perm.subtypeCongr.left_apply {a : ε} (h : p a) : ep.subtypeCongr en a = ep ⟨a, h⟩ := by simp [Perm.subtypeCongr.apply, h] #align equiv.perm.subtype_congr.left_apply Equiv.Perm.subtypeCongr.left_apply @@ -537,6 +541,7 @@ theorem Perm.subtypeCongr.left_apply_subtype (a : { a // p a }) : ep.subtypeCong Perm.subtypeCongr.left_apply ep en a.property #align equiv.perm.subtype_congr.left_apply_subtype Equiv.Perm.subtypeCongr.left_apply_subtype +@[simp] theorem Perm.subtypeCongr.right_apply {a : ε} (h : ¬p a) : ep.subtypeCongr en a = en ⟨a, h⟩ := by simp [Perm.subtypeCongr.apply, h] #align equiv.perm.subtype_congr.right_apply Equiv.Perm.subtypeCongr.right_apply @@ -550,15 +555,19 @@ theorem Perm.subtypeCongr.right_apply_subtype (a : { a // ¬p a }) : ep.subtypeC theorem Perm.subtypeCongr.refl : Perm.subtypeCongr (Equiv.refl { a // p a }) (Equiv.refl { a // ¬p a }) = Equiv.refl ε := by ext x - by_cases h:p x <;> simp [h, subtypeCongr.left_apply, subtypeCongr.right_apply] + by_cases h:p x <;> simp [h] #align equiv.perm.subtype_congr.refl Equiv.Perm.subtypeCongr.refl @[simp] theorem Perm.subtypeCongr.symm : (ep.subtypeCongr en).symm = Perm.subtypeCongr ep.symm en.symm := by ext x by_cases h:p x - · simp [Perm.subtypeCongr.apply, h, symm_apply_eq] - · simp [Perm.subtypeCongr.apply, h, symm_apply_eq] + · have : p (ep.symm ⟨x, h⟩) := Subtype.property _ + simp [Perm.subtypeCongr.apply, h, symm_apply_eq, this] + + · have : ¬p (en.symm ⟨x, h⟩) := Subtype.property (en.symm _) + simp [Perm.subtypeCongr.apply, h, symm_apply_eq, this] + #align equiv.perm.subtype_congr.symm Equiv.Perm.subtypeCongr.symm @[simp] @@ -567,8 +576,12 @@ theorem Perm.subtypeCongr.trans : = Perm.subtypeCongr (ep.trans ep') (en.trans en') := by ext x by_cases h:p x - · simp [Perm.subtypeCongr.apply, h] - · simp [Perm.subtypeCongr.apply, h, symm_apply_eq] + · have : p (ep ⟨x, h⟩) := Subtype.property _ + simp [Perm.subtypeCongr.apply, h, this] + + · have : ¬p (en ⟨x, h⟩) := Subtype.property (en _) + simp [Perm.subtypeCongr.apply, h, symm_apply_eq, this] + #align equiv.perm.subtype_congr.trans Equiv.Perm.subtypeCongr.trans end sumCompl @@ -1373,11 +1386,11 @@ theorem Perm.extendDomain_apply_image (a : α') : e.extendDomain f (f a) = f (e theorem Perm.extendDomain_apply_subtype {b : β'} (h : p b) : e.extendDomain f b = f (e (f.symm ⟨b, h⟩)) := by - simp [Perm.extendDomain, h, subtypeCongr.left_apply] + simp [Perm.extendDomain, h] #align equiv.perm.extend_domain_apply_subtype Equiv.Perm.extendDomain_apply_subtype theorem Perm.extendDomain_apply_not_subtype {b : β'} (h : ¬p b) : e.extendDomain f b = b := by - simp [Perm.extendDomain, h, subtypeCongr.right_apply] + simp [Perm.extendDomain, h] #align equiv.perm.extend_domain_apply_not_subtype Equiv.Perm.extendDomain_apply_not_subtype @[simp] @@ -1895,3 +1908,4 @@ theorem piCongrLeft'_symm_update [DecidableEq α] [DecidableEq β] (P : α → S #align function.Pi_congr_left'_symm_update Function.piCongrLeft'_symm_update end Function +#lint From 9e43f15d4e66f499fed9f62e71b24204bebfbaff Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 23 Nov 2022 10:36:42 +1100 Subject: [PATCH 081/127] almost there --- Mathlib/Logic/Equiv/Basic.lean | 15 +++------------ Mathlib/Logic/Equiv/Defs.lean | 3 +++ 2 files changed, 6 insertions(+), 12 deletions(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index f862dd267fc80..446abd085035b 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -813,7 +813,7 @@ theorem eq_of_prodExtendRight_ne {e : Perm β₁} {a a' : α₁} {b : β₁} @[simp] theorem fst_prodExtendRight (ab : α₁ × β₁) : (prodExtendRight a e ab).fst = ab.fst := by rw [prodExtendRight] - dsimp only + dsimp split_ifs with h · rw [h] · rfl @@ -1803,20 +1803,11 @@ instance [IsAssociative α₁ f] : IsAssociative β₁ (e.arrowCongr (e.arrowCon instance [IsIdempotent α₁ f] : IsIdempotent β₁ (e.arrowCongr (e.arrowCongr e) f) := (e.semiconj₂_conj f).isIdempotent_right e.surjective --- porting note: the coe changes make `EmbeddingLike.apply_eq_iff_eq` almost impossible to use instance [IsLeftCancel α₁ f] : IsLeftCancel β₁ (e.arrowCongr (e.arrowCongr e) f) := - ⟨e.surjective.forall₃.2 fun x y z => by - simp only [arrowCongr, coe_fn_symm_mk, comp_apply, symm_apply_apply] - intro h - rw [@IsLeftCancel.left_cancel _ f _ x y z _] - exact e.injective h⟩ + ⟨e.surjective.forall₃.2 fun x y z => by simpa using @IsLeftCancel.left_cancel _ f _ x y z⟩ instance [IsRightCancel α₁ f] : IsRightCancel β₁ (e.arrowCongr (e.arrowCongr e) f) := - ⟨e.surjective.forall₃.2 fun x y z => by - simp only [arrowCongr, coe_fn_symm_mk, comp_apply, symm_apply_apply] - intro h - rw [@IsRightCancel.right_cancel _ f _ x y z _] - exact e.injective h⟩ + ⟨e.surjective.forall₃.2 fun x y z => by simpa using @IsRightCancel.right_cancel _ f _ x y z⟩ end BinaryOp diff --git a/Mathlib/Logic/Equiv/Defs.lean b/Mathlib/Logic/Equiv/Defs.lean index 42b5c4a4fb979..dc618c4429213 100644 --- a/Mathlib/Logic/Equiv/Defs.lean +++ b/Mathlib/Logic/Equiv/Defs.lean @@ -92,6 +92,9 @@ instance : EquivLike (α ≃ β) α β where right_inv := right_inv coe_injective' e₁ e₂ h₁ h₂ := by cases e₁; cases e₂; congr +@[simp] theorem coe_fn_mk (f : α → β) (g l r) : (Equiv.mk f g l r : α → β) = f := +rfl + /-- The map `(r ≃ s) → (r → s)` is injective. -/ theorem coe_fn_injective : @Function.Injective (α ≃ β) (α → β) (fun e => e) := FunLike.coe_injective' From b9678126cf8a9c7e72420d93a531883c29e1a37a Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 23 Nov 2022 10:37:33 +1100 Subject: [PATCH 082/127] and lints! --- Mathlib/Logic/Equiv/Basic.lean | 2 -- 1 file changed, 2 deletions(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 446abd085035b..0279328ef5a17 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -290,7 +290,6 @@ theorem sumCongr_apply (ea : Equiv.Perm α) (eb : Equiv.Perm β) (x : Sum α β) -- porting note: it seems the general theorem about `Equiv` is now applied, so there's no need -- to have this version also have `@[simp]`. Similarly for below. -@[simp] theorem sumCongr_trans (e : Equiv.Perm α) (f : Equiv.Perm β) (g : Equiv.Perm α) (h : Equiv.Perm β) : (sumCongr e f).trans (sumCongr g h) = sumCongr (e.trans g) (f.trans h) := Equiv.sumCongr_trans e f g h @@ -1899,4 +1898,3 @@ theorem piCongrLeft'_symm_update [DecidableEq α] [DecidableEq β] (P : α → S #align function.Pi_congr_left'_symm_update Function.piCongrLeft'_symm_update end Function -#lint From b6a40982c9e0fc49d684cd0b0cf181d3b52fbbf1 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 23 Nov 2022 10:40:38 +1100 Subject: [PATCH 083/127] bump --- lean_packages/manifest.json | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/lean_packages/manifest.json b/lean_packages/manifest.json index 369179063781c..9a302046d3bf1 100644 --- a/lean_packages/manifest.json +++ b/lean_packages/manifest.json @@ -1,7 +1,7 @@ {"version": 2, "packages": [{"url": "https://github.com/mhuisi/lean4-cli", - "rev": "b88b0ab3ad5f1352d107d3e8740ff15dc1fbb5b1", + "rev": "5a858c32963b6b19be0d477a30a1f4b6c120be7e", "name": "Cli", "inputRev": "nightly"}, {"url": "https://github.com/hargonix/LeanInk", @@ -13,7 +13,7 @@ "name": "aesop", "inputRev": "master"}, {"url": "https://github.com/leanprover/doc-gen4", - "rev": "664a86e08bba7f4415fb0fdaef04c4508af702d1", + "rev": "69d48b174dbe404f37fe7e572de12f0c21992729", "name": "doc-gen4", "inputRev": "main"}, {"url": "https://github.com/xubaiw/Unicode.lean", @@ -21,7 +21,7 @@ "name": "Unicode", "inputRev": "main"}, {"url": "https://github.com/leanprover/std4", - "rev": "eaac7d455caa50a9c493f9c0c1fd13aaadfe3b03", + "rev": "68581a4685ce2ae8906f237651b3381d1aaa8b41", "name": "std", "inputRev": "main"}, {"url": "https://github.com/leanprover/lake", From a3a8b25c476399b474549f62643f8fa82c88f92f Mon Sep 17 00:00:00 2001 From: Moritz Doll Date: Wed, 23 Nov 2022 11:14:03 +1100 Subject: [PATCH 084/127] fixed coe --- Mathlib/Logic/Equiv/Option.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Logic/Equiv/Option.lean b/Mathlib/Logic/Equiv/Option.lean index 852025c1089d1..c0b57321f8dc9 100644 --- a/Mathlib/Logic/Equiv/Option.lean +++ b/Mathlib/Logic/Equiv/Option.lean @@ -181,7 +181,7 @@ def option_subtype [DecidableEq β] (x : β) : simp exact apply_symm_apply _ _ } invFun e := - ⟨{ toFun := fun a => casesOn' a x (coe ∘ e), + ⟨{ toFun := fun a => casesOn' a x (Subtype.val ∘ toFun ∘ e), invFun := fun b => if h : b = x then none else e.symm ⟨b, h⟩, left_inv := fun a => by cases a From 053a10e37e5234dafbfceb575582f4a322170814 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 23 Nov 2022 13:13:07 +1100 Subject: [PATCH 085/127] fix casing issue --- Mathlib/Order/Compare.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/Mathlib/Order/Compare.lean b/Mathlib/Order/Compare.lean index b7a8b301885f1..6eef634908075 100644 --- a/Mathlib/Order/Compare.lean +++ b/Mathlib/Order/Compare.lean @@ -53,13 +53,13 @@ def Compares [LT α] : Ordering → α → α → Prop | gt, a, b => a > b @[simp] -lemma Compares_lt [LT α] (a b : α) : Compares lt a b = (a < b) := rfl +lemma compares_lt [LT α] (a b : α) : Compares lt a b = (a < b) := rfl @[simp] -lemma Compares_eq [LT α] (a b : α) : Compares eq a b = (a = b) := rfl +lemma compares_eq [LT α] (a b : α) : Compares eq a b = (a = b) := rfl @[simp] -lemma Compares_gt [LT α] (a b : α) : Compares gt a b = (a > b) := rfl +lemma compares_gt [LT α] (a b : α) : Compares gt a b = (a > b) := rfl theorem compares_swap [LT α] {a b : α} {o : Ordering} : o.swap.Compares a b ↔ o.Compares b a := by cases o From 41e4084c25f015e9effa737019fe0a1dade20dce Mon Sep 17 00:00:00 2001 From: Arien Malec Date: Tue, 22 Nov 2022 18:45:25 -0800 Subject: [PATCH 086/127] fixed naming --- Mathlib/Logic/Equiv/Option.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Logic/Equiv/Option.lean b/Mathlib/Logic/Equiv/Option.lean index c0b57321f8dc9..192b5488d69b1 100644 --- a/Mathlib/Logic/Equiv/Option.lean +++ b/Mathlib/Logic/Equiv/Option.lean @@ -187,7 +187,7 @@ def option_subtype [DecidableEq β] (x : β) : cases a · simp - simp only [cases_on'_some, Function.comp_apply, Subtype.coe_eta, + simp only [casesOn'_some, Function.comp_apply, Subtype.coe_eta, symm_apply_apply, dite_eq_ite] exact if_neg (e a).property, right_inv := fun b => by by_cases h : b = x <;> simp [h] }, From f4479e4b8c2cb5fbb0f22fe029d3b43cb719cb3c Mon Sep 17 00:00:00 2001 From: Arien Malec Date: Tue, 22 Nov 2022 18:46:30 -0800 Subject: [PATCH 087/127] style fix of rfl proofs --- Mathlib/Logic/Equiv/Option.lean | 6 ++---- 1 file changed, 2 insertions(+), 4 deletions(-) diff --git a/Mathlib/Logic/Equiv/Option.lean b/Mathlib/Logic/Equiv/Option.lean index 192b5488d69b1..4a32e15cd022e 100644 --- a/Mathlib/Logic/Equiv/Option.lean +++ b/Mathlib/Logic/Equiv/Option.lean @@ -209,16 +209,14 @@ theorem option_subtype_apply_apply [DecidableEq β] (x : β) (e : { e : Option α ≃ β // e none = x }) (a : α) - (h) : option_subtype x e a = ⟨(e : Option α ≃ β) a, h⟩ := - rfl + (h) : option_subtype x e a = ⟨(e : Option α ≃ β) a, h⟩ := rfl #align equiv.option_subtype_apply_apply Equiv.option_subtype_apply_apply @[simp] theorem coe_option_subtype_apply_apply [DecidableEq β] (x : β) (e : { e : Option α ≃ β // e none = x }) - (a : α) : ↑(option_subtype x e a) = (e : Option α ≃ β) a := - rfl + (a : α) : ↑(option_subtype x e a) = (e : Option α ≃ β) a := rfl #align equiv.coe_option_subtype_apply_apply Equiv.coe_option_subtype_apply_apply @[simp] From 0ce71aeba210b227155e494b21d14882e36e7a79 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 23 Nov 2022 13:57:54 +1100 Subject: [PATCH 088/127] explicit type ascriptions --- Mathlib/Order/Monotone.lean | 40 ++++++++++++++++++++++--------------- 1 file changed, 24 insertions(+), 16 deletions(-) diff --git a/Mathlib/Order/Monotone.lean b/Mathlib/Order/Monotone.lean index 04e98aad651d1..e88e186537cad 100644 --- a/Mathlib/Order/Monotone.lean +++ b/Mathlib/Order/Monotone.lean @@ -130,13 +130,17 @@ theorem antitone_comp_ofDual_iff : Antitone (f ∘ ofDual) ↔ Monotone f := forall_swap #align antitone_comp_of_dual_iff antitone_comp_ofDual_iff +-- Porting note: +-- Here (and below) without the type ascription, Lean is seeing through the +-- defeq `βᵒᵈ = β` and picking up the wrong `Preorder` instance. +-- https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/logic.2Eequiv.2Ebasic.20mathlib4.23631/near/311744939 @[simp] -theorem monotone_toDual_comp_iff : Monotone (toDual ∘ f) ↔ Antitone f := +theorem monotone_toDual_comp_iff : Monotone (toDual ∘ f : α → βᵒᵈ) ↔ Antitone f := Iff.rfl #align monotone_to_dual_comp_iff monotone_toDual_comp_iff @[simp] -theorem antitone_toDual_comp_iff : Antitone (toDual ∘ f) ↔ Monotone f := +theorem antitone_toDual_comp_iff : Antitone (toDual ∘ f : α → βᵒᵈ) ↔ Monotone f := Iff.rfl #align antitone_to_dual_comp_iff antitone_toDual_comp_iff @@ -151,12 +155,12 @@ theorem antitoneOn_comp_ofDual_iff : AntitoneOn (f ∘ ofDual) s ↔ MonotoneOn #align antitone_on_comp_of_dual_iff antitoneOn_comp_ofDual_iff @[simp] -theorem monotoneOn_toDual_comp_iff : MonotoneOn (toDual ∘ f) s ↔ AntitoneOn f s := +theorem monotoneOn_toDual_comp_iff : MonotoneOn (toDual ∘ f : α → βᵒᵈ) s ↔ AntitoneOn f s := Iff.rfl #align monotone_on_to_dual_comp_iff monotoneOn_toDual_comp_iff @[simp] -theorem antitoneOn_toDual_comp_iff : AntitoneOn (toDual ∘ f) s ↔ MonotoneOn f s := +theorem antitoneOn_toDual_comp_iff : AntitoneOn (toDual ∘ f : α → βᵒᵈ) s ↔ MonotoneOn f s := Iff.rfl #align antitone_on_to_dual_comp_iff antitoneOn_toDual_comp_iff @@ -171,12 +175,12 @@ theorem strictAnti_comp_ofDual_iff : StrictAnti (f ∘ ofDual) ↔ StrictMono f #align strict_anti_comp_of_dual_iff strictAnti_comp_ofDual_iff @[simp] -theorem strictMono_toDual_comp_iff : StrictMono (toDual ∘ f) ↔ StrictAnti f := +theorem strictMono_toDual_comp_iff : StrictMono (toDual ∘ f : α → βᵒᵈ) ↔ StrictAnti f := Iff.rfl #align strict_mono_to_dual_comp_iff strictMono_toDual_comp_iff @[simp] -theorem strictAnti_toDual_comp_iff : StrictAnti (toDual ∘ f) ↔ StrictMono f := +theorem strictAnti_toDual_comp_iff : StrictAnti (toDual ∘ f : α → βᵒᵈ) ↔ StrictMono f := Iff.rfl #align strict_anti_to_dual_comp_iff strictAnti_toDual_comp_iff @@ -191,39 +195,43 @@ theorem strictAntiOn_comp_ofDual_iff : StrictAntiOn (f ∘ ofDual) s ↔ StrictM #align strict_anti_on_comp_of_dual_iff strictAntiOn_comp_ofDual_iff @[simp] -theorem strictMonoOn_toDual_comp_iff : StrictMonoOn (toDual ∘ f) s ↔ StrictAntiOn f s := +theorem strictMonoOn_toDual_comp_iff : StrictMonoOn (toDual ∘ f : α → βᵒᵈ) s ↔ StrictAntiOn f s := Iff.rfl #align strict_mono_on_to_dual_comp_iff strictMonoOn_toDual_comp_iff @[simp] -theorem strictAntiOn_toDual_comp_iff : StrictAntiOn (toDual ∘ f) s ↔ StrictMonoOn f s := +theorem strictAntiOn_toDual_comp_iff : StrictAntiOn (toDual ∘ f : α → βᵒᵈ) s ↔ StrictMonoOn f s := Iff.rfl #align strict_anti_on_to_dual_comp_iff strictAntiOn_toDual_comp_iff -protected theorem Monotone.dual (hf : Monotone f) : Monotone (toDual ∘ f ∘ ofDual) := +protected theorem Monotone.dual (hf : Monotone f) : Monotone (toDual ∘ f ∘ ofDual : αᵒᵈ → βᵒᵈ) := swap hf -protected theorem Antitone.dual (hf : Antitone f) : Antitone (toDual ∘ f ∘ ofDual) := +protected theorem Antitone.dual (hf : Antitone f) : Antitone (toDual ∘ f ∘ ofDual : αᵒᵈ → βᵒᵈ) := swap hf -protected theorem MonotoneOn.dual (hf : MonotoneOn f s) : MonotoneOn (toDual ∘ f ∘ ofDual) s := +protected theorem MonotoneOn.dual (hf : MonotoneOn f s) : + MonotoneOn (toDual ∘ f ∘ ofDual : αᵒᵈ → βᵒᵈ) s := swap₂ hf -protected theorem AntitoneOn.dual (hf : AntitoneOn f s) : AntitoneOn (toDual ∘ f ∘ ofDual) s := +protected theorem AntitoneOn.dual (hf : AntitoneOn f s) : + AntitoneOn (toDual ∘ f ∘ ofDual : αᵒᵈ → βᵒᵈ) s := swap₂ hf -protected theorem StrictMono.dual (hf : StrictMono f) : StrictMono (toDual ∘ f ∘ ofDual) := +protected theorem StrictMono.dual (hf : StrictMono f) : + StrictMono (toDual ∘ f ∘ ofDual : αᵒᵈ → βᵒᵈ) := swap hf -protected theorem StrictAnti.dual (hf : StrictAnti f) : StrictAnti (toDual ∘ f ∘ ofDual) := +protected theorem StrictAnti.dual (hf : StrictAnti f) : + StrictAnti (toDual ∘ f ∘ ofDual : αᵒᵈ → βᵒᵈ) := swap hf protected theorem StrictMonoOn.dual (hf : StrictMonoOn f s) : - StrictMonoOn (toDual ∘ f ∘ ofDual) s := + StrictMonoOn (toDual ∘ f ∘ ofDual : αᵒᵈ → βᵒᵈ) s := swap₂ hf protected theorem StrictAntiOn.dual (hf : StrictAntiOn f s) : - StrictAntiOn (toDual ∘ f ∘ ofDual) s := + StrictAntiOn (toDual ∘ f ∘ ofDual : αᵒᵈ → βᵒᵈ) s := swap₂ hf alias antitone_comp_ofDual_iff ↔ _ Monotone.dual_left From c9ba6e7f40352fdf26841395a360346c626d274b Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 23 Nov 2022 14:14:15 +1100 Subject: [PATCH 089/127] remove unneeded simps --- Mathlib/Order/Synonym.lean | 12 ++++++++---- 1 file changed, 8 insertions(+), 4 deletions(-) diff --git a/Mathlib/Order/Synonym.lean b/Mathlib/Order/Synonym.lean index 057bb94e21cd6..38463abbab7fd 100644 --- a/Mathlib/Order/Synonym.lean +++ b/Mathlib/Order/Synonym.lean @@ -70,12 +70,14 @@ theorem ofDual_toDual (a : α) : ofDual (toDual a) = a := rfl #align order_dual.of_dual_to_dual OrderDual.ofDual_toDual -@[simp] +-- Porting note: +-- removed @[simp] since this already follows by `simp only [EmbeddingLike.apply_eq_iff_eq]` theorem toDual_inj {a b : α} : toDual a = toDual b ↔ a = b := Iff.rfl #align order_dual.to_dual_inj OrderDual.toDual_inj -@[simp] +-- Porting note: +-- removed @[simp] since this already follows by `simp only [EmbeddingLike.apply_eq_iff_eq]` theorem ofDual_inj {a b : αᵒᵈ} : ofDual a = ofDual b ↔ a = b := Iff.rfl #align order_dual.of_dual_inj OrderDual.ofDual_inj @@ -178,12 +180,14 @@ theorem ofLex_toLex (a : α) : ofLex (toLex a) = a := rfl #align of_lex_to_lex ofLex_toLex -@[simp] +-- Porting note: +-- removed @[simp] since this already follows by `simp only [EmbeddingLike.apply_eq_iff_eq]` theorem toLex_inj {a b : α} : toLex a = toLex b ↔ a = b := Iff.rfl #align to_lex_inj toLex_inj -@[simp] +-- Porting note: +-- removed @[simp] since this already follows by `simp only [EmbeddingLike.apply_eq_iff_eq]` theorem ofLex_inj {a b : Lex α} : ofLex a = ofLex b ↔ a = b := Iff.rfl #align of_lex_inj ofLex_inj From 4a7b7fca4601e84a80a9fe8849cca4ebc0015d27 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 23 Nov 2022 14:36:22 +1100 Subject: [PATCH 090/127] lint --- Mathlib/Logic/Equiv/Defs.lean | 1 + 1 file changed, 1 insertion(+) diff --git a/Mathlib/Logic/Equiv/Defs.lean b/Mathlib/Logic/Equiv/Defs.lean index dc618c4429213..7f8972299a5c1 100644 --- a/Mathlib/Logic/Equiv/Defs.lean +++ b/Mathlib/Logic/Equiv/Defs.lean @@ -134,6 +134,7 @@ protected def symm (e : α ≃ β) : β ≃ α := ⟨e.invFun, e.toFun, e.right_ /-- See Note [custom simps projection] -/ def Simps.apply (e : α ≃ β) : α → β := e +/-- See Note [custom simps projection] -/ def Simps.symm_apply (e : α ≃ β) : β → α := e.symm initialize_simps_projections Equiv (toFun → apply, invFun → symm_apply) From c4824fa8eeedb45c62d53b8272f64a74a4739102 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Wed, 23 Nov 2022 14:39:01 +1100 Subject: [PATCH 091/127] add a linink to a Lean 4 issue --- Mathlib/Algebra/Group/Defs.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Algebra/Group/Defs.lean b/Mathlib/Algebra/Group/Defs.lean index ae0d21a06e1cc..18b1fc529511d 100644 --- a/Mathlib/Algebra/Group/Defs.lean +++ b/Mathlib/Algebra/Group/Defs.lean @@ -254,7 +254,7 @@ attribute [to_additive] MulOneClass @[ext, to_additive] theorem MulOneClass.ext {M : Type u} : ∀ ⦃m₁ m₂ : MulOneClass M⦄, m₁.mul = m₂.mul → m₁ = m₂ := by rintro @⟨⟨one₁⟩, ⟨mul₁⟩, one_mul₁, mul_one₁⟩ @⟨⟨one₂⟩, ⟨mul₂⟩, one_mul₂, mul_one₂⟩ ⟨rfl⟩ - -- FIXME: + -- FIXME (See https://github.com/leanprover/lean4/issues/1711) -- congr suffices one₁ = one₂ by cases this; rfl exact (one_mul₂ one₁).symm.trans (mul_one₁ one₂) From c2f70bde0d168b87b991fa54197c403affc334a1 Mon Sep 17 00:00:00 2001 From: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> Date: Wed, 23 Nov 2022 16:19:48 +0100 Subject: [PATCH 092/127] Update Mathlib/Logic/Equiv/Basic.lean Co-authored-by: Kevin Buzzard --- Mathlib/Logic/Equiv/Basic.lean | 1 - 1 file changed, 1 deletion(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 0279328ef5a17..bfc470e01bddc 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -638,7 +638,6 @@ def piComm (φ : α → β → Sort _) : (∀ a b, φ a b) ≃ ∀ b a, φ a b : ⟨swap, swap, fun _ => rfl, fun _ => rfl⟩ #align equiv.Pi_comm Equiv.piComm --- up to here **TODO** remove this @[simp] theorem piComm_symm {φ : α → β → Sort _} : (piComm φ).symm = (piComm <| swap φ) := rfl From 7ce749c34a55c0a7b9eb3d93d8247d09b6d5d2ca Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Wed, 23 Nov 2022 16:04:52 +0000 Subject: [PATCH 093/127] fix a bunch of names, long line --- Mathlib/Logic/Equiv/Option.lean | 85 +++++++++++++++++---------------- 1 file changed, 43 insertions(+), 42 deletions(-) diff --git a/Mathlib/Logic/Equiv/Option.lean b/Mathlib/Logic/Equiv/Option.lean index 4a32e15cd022e..d6288fa9d6712 100644 --- a/Mathlib/Logic/Equiv/Option.lean +++ b/Mathlib/Logic/Equiv/Option.lean @@ -13,9 +13,9 @@ import Mathlib.Logic.Equiv.Defs We define -* `Equiv.option_congr`: the `Option α ≃ Option β` constructed from `e : α ≃ β` by sending `none` to +* `Equiv.optionCongr`: the `Option α ≃ Option β` constructed from `e : α ≃ β` by sending `none` to `none`, and applying a `e` elsewhere. -* `Equiv.remove_none`: the `α ≃ β` constructed from `Option α ≃ Option β` by removing `none` from +* `Equiv.removeNone`: the `α ≃ β` constructed from `Option α ≃ Option β` by removing `none` from both sides. -/ @@ -38,27 +38,27 @@ def optionCongr (e : α ≃ β) : Option α ≃ Option β where #align equiv.option_congr Equiv.optionCongr @[simp] -theorem option_congr_refl : optionCongr (Equiv.refl α) = Equiv.refl _ := +theorem optionCongr_refl : optionCongr (Equiv.refl α) = Equiv.refl _ := ext <| congr_fun Option.map_id -#align equiv.option_congr_refl Equiv.option_congr_refl +#align equiv.option_congr_refl Equiv.optionCongr_refl @[simp] -theorem option_congr_symm (e : α ≃ β) : (optionCongr e).symm = optionCongr e.symm := +theorem optionCongr_symm (e : α ≃ β) : (optionCongr e).symm = optionCongr e.symm := rfl -#align equiv.option_congr_symm Equiv.option_congr_symm +#align equiv.option_congr_symm Equiv.optionCongr_symm @[simp] -theorem option_congr_trans (e₁ : α ≃ β) (e₂ : β ≃ γ) : +theorem optionCongr_trans (e₁ : α ≃ β) (e₂ : β ≃ γ) : (optionCongr e₁).trans (optionCongr e₂) = optionCongr (e₁.trans e₂) := ext <| Option.map_map _ _ -#align equiv.option_congr_trans Equiv.option_congr_trans +#align equiv.option_congr_trans Equiv.optionCongr_trans /-- When `α` and `β` are in the same universe, this is the same as the result of `EquivFunctor.map_equiv`. -/ -theorem option_congr_eq_equiv_function_map_equiv {α β : Type _} (e : α ≃ β) : +theorem optionCongr_eq_equivFunction_mapEquiv {α β : Type _} (e : α ≃ β) : optionCongr e = EquivFunctor.mapEquiv Option e := rfl -#align equiv.option_congr_eq_equiv_function_map_equiv Equiv.option_congr_eq_equiv_function_map_equiv +#align equiv.option_congr_eq_equiv_function_map_equiv Equiv.optionCongr_eq_equivFunction_mapEquiv end OptionCongr @@ -69,7 +69,7 @@ variable (e : Option α ≃ Option β) /-- If we have a value on one side of an `Equiv` of `Option` we also have a value on the other side of the equivalence -/ -def remove_none_aux (x : α) : β := +def removeNone_aux (x : α) : β := if h : (e (some x)).isSome then Option.get _ h else Option.get _ <| @@ -79,44 +79,44 @@ def remove_none_aux (x : α) : β := rw [Option.not_isSome_iff_eq_none, ← hn] at h exact Option.some_ne_none _ (e.injective h) -#align equiv.remove_none_aux Equiv.remove_none_aux +#align equiv.remove_none_aux Equiv.removeNone_aux -theorem remove_none_aux_some {x : α} (h : ∃ x', e (some x) = some x') : - some (remove_none_aux e x) = e (some x) := - by simp [remove_none_aux, Option.isSome_iff_exists.mpr h] -#align equiv.remove_none_aux_some Equiv.remove_none_aux_some +theorem removeNone_aux_some {x : α} (h : ∃ x', e (some x) = some x') : + some (removeNone_aux e x) = e (some x) := + by simp [removeNone_aux, Option.isSome_iff_exists.mpr h] +#align equiv.remove_none_aux_some Equiv.removeNone_aux_some -theorem remove_none_aux_none {x : α} (h : e (some x) = none) : - some (remove_none_aux e x) = e none := by - simp [remove_none_aux, Option.not_isSome_iff_eq_none.mpr h] -#align equiv.remove_none_aux_none Equiv.remove_none_aux_none +theorem removeNone_aux_none {x : α} (h : e (some x) = none) : + some (removeNone_aux e x) = e none := by + simp [removeNone_aux, Option.not_isSome_iff_eq_none.mpr h] +#align equiv.remove_none_aux_none Equiv.removeNone_aux_none -theorem remove_none_aux_inv (x : α) : remove_none_aux e.symm (remove_none_aux e x) = x := +theorem removeNone_aux_inv (x : α) : removeNone_aux e.symm (removeNone_aux e x) = x := Option.some_injective _ (by - cases h1 : e.symm (some (remove_none_aux e x)) <;> cases h2 : e (some x) - · rw [remove_none_aux_none _ h1] + cases h1 : e.symm (some (removeNone_aux e x)) <;> cases h2 : e (some x) + · rw [removeNone_aux_none _ h1] exact (e.eq_symm_apply.mpr h2).symm - · rw [remove_none_aux_some _ ⟨_, h2⟩] at h1 + · rw [removeNone_aux_some _ ⟨_, h2⟩] at h1 simp at h1 - · rw [remove_none_aux_none _ h2] at h1 + · rw [removeNone_aux_none _ h2] at h1 simp at h1 - · rw [remove_none_aux_some _ ⟨_, h1⟩] - rw [remove_none_aux_some _ ⟨_, h2⟩] + · rw [removeNone_aux_some _ ⟨_, h1⟩] + rw [removeNone_aux_some _ ⟨_, h2⟩] simp ) -#align equiv.remove_none_aux_inv Equiv.remove_none_aux_inv +#align equiv.remove_none_aux_inv Equiv.removeNone_aux_inv /-- Given an equivalence between two `Option` types, eliminate `none` from that equivalence by mapping `e.symm none` to `e none`. -/ def removeNone : α ≃ β where - toFun := remove_none_aux e - invFun := remove_none_aux e.symm - left_inv := remove_none_aux_inv e - right_inv := remove_none_aux_inv e.symm + toFun := removeNone_aux e + invFun := removeNone_aux e.symm + left_inv := removeNone_aux_inv e + right_inv := removeNone_aux_inv e.symm #align equiv.remove_none Equiv.removeNone @[simp] @@ -124,12 +124,13 @@ theorem remove_none_symm : (removeNone e).symm = removeNone e.symm := rfl #align equiv.remove_none_symm Equiv.remove_none_symm -theorem remove_none_some {x : α} (h : ∃ x', e (some x) = some x') : some (removeNone e x) = e (some x) := - remove_none_aux_some e h +theorem remove_none_some {x : α} (h : ∃ x', e (some x) = some x') : + some (removeNone e x) = e (some x) := + removeNone_aux_some e h #align equiv.remove_none_some Equiv.remove_none_some theorem remove_none_none {x : α} (h : e (some x) = none) : some (removeNone e x) = e none := - remove_none_aux_none e h + removeNone_aux_none e h #align equiv.remove_none_none Equiv.remove_none_none @[simp] @@ -137,7 +138,7 @@ theorem option_symm_apply_none_iff : e.symm none = none ↔ e none = none := ⟨fun h => by simpa using (congr_arg e h).symm, fun h => by simpa using (congr_arg e.symm h).symm⟩ #align equiv.option_symm_apply_none_iff Equiv.option_symm_apply_none_iff -theorem some_remove_none_iff {x : α} : some (removeNone e x) = e none ↔ e.symm none = some x := by +theorem some_removeNone_iff {x : α} : some (removeNone e x) = e none ↔ e.symm none = some x := by cases' h : e (some x) with a · rw [remove_none_none _ h] simpa using (congr_arg e.symm h).symm @@ -148,18 +149,18 @@ theorem some_remove_none_iff {x : α} : some (removeNone e x) = e none ↔ e.sym simp only [false_iff_iff, apply_eq_iff_eq] simp [h1, apply_eq_iff_eq] -#align equiv.some_remove_none_iff Equiv.some_remove_none_iff +#align equiv.some_remove_none_iff Equiv.some_removeNone_iff @[simp] -theorem remove_none_option_congr (e : α ≃ β) : removeNone e.optionCongr = e := +theorem removeNone_option_congr (e : α ≃ β) : removeNone e.optionCongr = e := Equiv.ext fun x => Option.some_injective _ <| remove_none_some _ ⟨e x, by simp [EquivFunctor.map]⟩ -#align equiv.remove_none_option_congr Equiv.remove_none_option_congr +#align equiv.remove_none_option_congr Equiv.removeNone_option_congr end RemoveNone -theorem option_congr_injective : Function.Injective (optionCongr : α ≃ β → Option α ≃ Option β) := - Function.LeftInverse.injective remove_none_option_congr -#align equiv.option_congr_injective Equiv.option_congr_injective +theorem optionCongr_injective : Function.Injective (optionCongr : α ≃ β → Option α ≃ Option β) := + Function.LeftInverse.injective removeNone_option_congr +#align equiv.option_congr_injective Equiv.optionCongr_injective /-- Equivalences between `Option α` and `β` that send `none` to `x` are equivalent to equivalences between `α` and `{y : β // y ≠ x}`. -/ From 89327d4e5782fb88de5129bb2e4b79b0e8c200cc Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Wed, 23 Nov 2022 16:42:29 +0000 Subject: [PATCH 094/127] Update Mathlib/Logic/Equiv/Basic.lean --- Mathlib/Logic/Equiv/Basic.lean | 1 - 1 file changed, 1 deletion(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index bfc470e01bddc..99edb8f9ec07d 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -5,7 +5,6 @@ Authors: Leonardo de Moura, Mario Carneiro Ported by: Kevin Buzzard, Ruben Vorster, Scott Morrison, Eric Rodriguez -/ import Mathlib.Data.Bool.Basic -import Mathlib.Data.Prod.Basic import Mathlib.Data.Sigma.Basic import Mathlib.Data.Subtype import Mathlib.Data.Sum.Basic From 6cb4aba39e150e76b5bcaf1d494fbd91be1c12d6 Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Wed, 23 Nov 2022 16:42:43 +0000 Subject: [PATCH 095/127] Update Mathlib/Logic/Equiv/Basic.lean --- Mathlib/Logic/Equiv/Basic.lean | 1 - 1 file changed, 1 deletion(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 99edb8f9ec07d..593d300b027f5 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -6,7 +6,6 @@ Ported by: Kevin Buzzard, Ruben Vorster, Scott Morrison, Eric Rodriguez -/ import Mathlib.Data.Bool.Basic import Mathlib.Data.Sigma.Basic -import Mathlib.Data.Subtype import Mathlib.Data.Sum.Basic import Mathlib.Init.Data.Sigma.Basic import Mathlib.Logic.Equiv.Defs From aa15dfed4d71d225453d9318e5e3eeb1b5b6beac Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Wed, 23 Nov 2022 16:43:02 +0000 Subject: [PATCH 096/127] Update Mathlib/Logic/Equiv/Basic.lean --- Mathlib/Logic/Equiv/Basic.lean | 1 - 1 file changed, 1 deletion(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 593d300b027f5..69b3e1dcaa9c5 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -10,7 +10,6 @@ import Mathlib.Data.Sum.Basic import Mathlib.Init.Data.Sigma.Basic import Mathlib.Logic.Equiv.Defs import Mathlib.Logic.Function.Conjugate -import Mathlib.Tactic.SplitIfs import Mathlib.Tactic.Convert import Mathlib.Tactic.Contrapose import Mathlib.Tactic.GeneralizeProofs From 5e8fbbd891de72ddd33d5f2af00abcae796cff46 Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Wed, 23 Nov 2022 16:43:21 +0000 Subject: [PATCH 097/127] Update Mathlib/Logic/Equiv/Basic.lean --- Mathlib/Logic/Equiv/Basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 69b3e1dcaa9c5..d9d18ba80539c 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -127,7 +127,7 @@ def prodAssoc (α β γ) : (α × β) × γ ≃ α × β × γ := fun ⟨_, ⟨_, _⟩⟩ => rfl⟩ #align equiv.prod_assoc Equiv.prodAssoc -/-- Functions on `α × β` are equivalent to functions `α → β → γ`. -/ +/-- `γ`-valued functions on `α × β` are equivalent to functions `α → β → γ`. -/ @[simps (config := { fullyApplied := false })] def curry (α β γ) : (α × β → γ) ≃ (α → β → γ) where toFun := Function.curry From 3adebd04c4b7d1c6addfca7a2121f954ee232277 Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Wed, 23 Nov 2022 16:43:57 +0000 Subject: [PATCH 098/127] Update Mathlib/Logic/Equiv/Basic.lean --- Mathlib/Logic/Equiv/Basic.lean | 2 -- 1 file changed, 2 deletions(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index d9d18ba80539c..d685ce0b29d7b 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -510,8 +510,6 @@ def subtypeCongr {p q : α → Prop} [DecidablePred p] [DecidablePred q] (sumCompl p).symm.trans ((sumCongr e f).trans (sumCompl q)) #align equiv.subtype_congr Equiv.subtypeCongr -open Equiv - variable {p : ε → Prop} [DecidablePred p] variable (ep ep' : Perm { a // p a }) (en en' : Perm { a // ¬p a }) From 0c46476ec0b4edade3a26314a28fe3c34f3a67d2 Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Wed, 23 Nov 2022 16:44:15 +0000 Subject: [PATCH 099/127] Update Mathlib/Logic/Equiv/Basic.lean --- Mathlib/Logic/Equiv/Basic.lean | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index d685ce0b29d7b..ecd30c1fb4ed5 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -1628,8 +1628,7 @@ def setValue (f : α ≃ β) (a : α) (b : β) : α ≃ β := @[simp] theorem setValue_eq (f : α ≃ β) (a : α) (b : β) : setValue f a b a = b := by - dsimp [setValue] - simp [swap_apply_left] + simp [setValue, swap_apply_left] #align equiv.set_value_eq Equiv.setValue_eq end Swap From f7b42273e2dcc28e81db7cd2bfa06559d8052ffd Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Wed, 23 Nov 2022 16:44:30 +0000 Subject: [PATCH 100/127] Update Mathlib/Logic/Equiv/Basic.lean --- Mathlib/Logic/Equiv/Basic.lean | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index ecd30c1fb4ed5..6c0d38ef849e7 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -1614,7 +1614,8 @@ theorem sumCongr_refl_swap {α β : Sort _} [DecidableEq α] [DecidableEq β] (i cases x · simp [Sum.map, swap_apply_of_ne_of_ne] - · simp [Sum.map, swap_apply_def] + · simp only [Equiv.sumCongr_apply, Sum.map, coe_refl, comp.right_id, Sum.elim_inr, comp_apply, + swap_apply_def, Sum.inr.injEq] split_ifs <;> rfl #align equiv.perm.sumCongr_refl_swap Equiv.Perm.sumCongr_refl_swap From f9423ba973e268e0b150aa83e0584b418007709a Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Wed, 23 Nov 2022 16:44:41 +0000 Subject: [PATCH 101/127] Update Mathlib/Logic/Equiv/Basic.lean --- Mathlib/Logic/Equiv/Basic.lean | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 6c0d38ef849e7..468d16f38e87a 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -1600,7 +1600,8 @@ theorem sumCongr_swap_refl {α β : Sort _} [DecidableEq α] [DecidableEq β] (i Equiv.Perm.sumCongr (Equiv.swap i j) (Equiv.refl β) = Equiv.swap (Sum.inl i) (Sum.inl j) := by ext x cases x - · simp [Sum.map, swap_apply_def] + · simp only [Equiv.sumCongr_apply, Sum.map, coe_refl, comp.right_id, Sum.elim_inl, comp_apply, + swap_apply_def, Sum.inl.injEq] split_ifs <;> rfl · simp [Sum.map, swap_apply_of_ne_of_ne] From 4639b4ef6dab5ae2ba6dde787f6e80ce81e3797a Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Wed, 23 Nov 2022 16:44:53 +0000 Subject: [PATCH 102/127] Update Mathlib/Logic/Equiv/Basic.lean --- Mathlib/Logic/Equiv/Basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 468d16f38e87a..e4a988eff3089 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -1587,7 +1587,7 @@ theorem swap_apply_ne_self_iff {a b x : α} : swap a b x ≠ x ↔ a ≠ b ∧ ( by_cases hax:x = a · simp [hax, eq_comm] - by_cases hbx:x = b + by_cases hbx : x = b · simp [hbx] simp [hab, hax, hbx, swap_apply_of_ne_of_ne] From a820ebd03b0b890cb33abc90ccc653eace9514af Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Wed, 23 Nov 2022 16:45:12 +0000 Subject: [PATCH 103/127] Update Mathlib/Logic/Equiv/Basic.lean --- Mathlib/Logic/Equiv/Basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index e4a988eff3089..40d19063decb0 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -1584,7 +1584,7 @@ theorem swap_apply_ne_self_iff {a b x : α} : swap a b x ≠ x ↔ a ≠ b ∧ ( by_cases hab:a = b · simp [hab] - by_cases hax:x = a + by_cases hax : x = a · simp [hax, eq_comm] by_cases hbx : x = b From 976292e4ebfd73c1fcc3d74b79e862119d8f0ed3 Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Wed, 23 Nov 2022 16:45:46 +0000 Subject: [PATCH 104/127] Update Mathlib/Logic/Equiv/Basic.lean --- Mathlib/Logic/Equiv/Basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 40d19063decb0..30b91c3c06607 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -594,7 +594,7 @@ def subtypePreimage : { x : α → β // x ∘ Subtype.val = x₀ } ≃ ({ a // left_inv := fun ⟨x, hx⟩ => Subtype.val_injective <| funext fun a => by - dsimp + dsimp only split_ifs { rw [← hx]; rfl } { rfl } From 1fc9f8124c663499d01d34d9f6099339a8e0b428 Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Wed, 23 Nov 2022 16:46:04 +0000 Subject: [PATCH 105/127] Update Mathlib/Logic/Equiv/Basic.lean --- Mathlib/Logic/Equiv/Basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 30b91c3c06607..fa89a54e5beca 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -601,7 +601,7 @@ def subtypePreimage : { x : α → β // x ∘ Subtype.val = x₀ } ≃ ({ a // right_inv x := funext fun ⟨a, h⟩ => show dite (p a) _ _ = _ by - dsimp + dsimp only rw [dif_neg h] #align equiv.subtype_preimage Equiv.subtypePreimage From dc4a4cdb544ebf53779a8baddfd3722ddbc87385 Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Wed, 23 Nov 2022 16:46:23 +0000 Subject: [PATCH 106/127] Update Mathlib/Logic/Equiv/Basic.lean --- Mathlib/Logic/Equiv/Basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index fa89a54e5beca..7dd390bfd9522 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -981,7 +981,7 @@ def natEquivNatSumPUnit : ℕ ≃ Sum ℕ PUnit where toFun n := Nat.casesOn n (inr PUnit.unit) inl invFun := Sum.elim Nat.succ fun _ => 0 left_inv n := by cases n <;> rfl - right_inv := by rintro (_ | _ | _) <;> rfl + right_inv := by rintro (_ | _) <;> rfl #align equiv.nat_equiv_nat_sum_punit Equiv.natEquivNatSumPUnit /-- `ℕ ⊕ Punit` is equivalent to `ℕ`. -/ From 095f8773fc7b8c43ab6ce11a2010d22d3e4e309a Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Wed, 23 Nov 2022 16:47:14 +0000 Subject: [PATCH 107/127] Update Mathlib/Logic/Equiv/Basic.lean --- Mathlib/Logic/Equiv/Basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 7dd390bfd9522..6119e9a708874 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -1549,7 +1549,7 @@ theorem symm_trans_swap_trans [DecidableEq β] (a b : α) (e : α ≃ β) : have : ∀ a, e.symm x = a ↔ x = e a := fun a => by rw [@eq_comm _ (e.symm x)] constructor <;> intros <;> simp_all - simp [swap_apply_def, this] + simp [trans_apply, swap_apply_def, this] split_ifs <;> simp #align equiv.symm_trans_swap_trans Equiv.symm_trans_swap_trans From a085bcbde1fc1668815f43c0c408774881f565ff Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Wed, 23 Nov 2022 16:47:40 +0000 Subject: [PATCH 108/127] Update Mathlib/Logic/Equiv/Basic.lean --- Mathlib/Logic/Equiv/Basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 6119e9a708874..194f1b44299c7 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -1567,7 +1567,7 @@ theorem swap_apply_self (i j a : α) : swap i j (swap i j a) = a := by /-- A function is invariant to a swap if it is equal at both elements -/ theorem apply_swap_eq_self {v : α → β} {i j : α} (hv : v i = v j) (k : α) : v (swap i j k) = v k := by - by_cases hi:k = i + by_cases hi : k = i · rw [hi, swap_apply_left, hv] by_cases hj:k = j From 57363d6acd8d8bdd1e460329fabd1ff25556f0f7 Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Wed, 23 Nov 2022 16:48:04 +0000 Subject: [PATCH 109/127] Update Mathlib/Logic/Equiv/Basic.lean --- Mathlib/Logic/Equiv/Basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 194f1b44299c7..ba6f3abe73aac 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -1570,7 +1570,7 @@ theorem apply_swap_eq_self {v : α → β} {i j : α} (hv : v i = v j) (k : α) by_cases hi : k = i · rw [hi, swap_apply_left, hv] - by_cases hj:k = j + by_cases hj : k = j · rw [hj, swap_apply_right, hv] rw [swap_apply_of_ne_of_ne hi hj] From 6e84e1419bd08d2bc6a3001e338f99ce0c1cdbc4 Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Wed, 23 Nov 2022 16:48:23 +0000 Subject: [PATCH 110/127] Update Mathlib/Logic/Equiv/Basic.lean --- Mathlib/Logic/Equiv/Basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index ba6f3abe73aac..9276a350349e7 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -1581,7 +1581,7 @@ theorem swap_apply_eq_iff {x y z w : α} : swap x y z = w ↔ z = swap x y w := #align equiv.swap_apply_eq_iff Equiv.swap_apply_eq_iff theorem swap_apply_ne_self_iff {a b x : α} : swap a b x ≠ x ↔ a ≠ b ∧ (x = a ∨ x = b) := by - by_cases hab:a = b + by_cases hab : a = b · simp [hab] by_cases hax : x = a From 6e8e74eaf92c89a5b05d9fa8542bdaa80c9a36a1 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Thu, 24 Nov 2022 03:51:06 +1100 Subject: [PATCH 111/127] Update Mathlib/Logic/Equiv/Basic.lean Co-authored-by: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> --- Mathlib/Logic/Equiv/Basic.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 9276a350349e7..bdaf0247eaa1e 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -387,8 +387,8 @@ theorem emptySum_apply_inr [IsEmpty α] (b : β) : emptySum α β (Sum.inr b) = /-- `Option α` is equivalent to `α ⊕ punit` -/ def optionEquivSumPUnit (α) : Option α ≃ Sum α PUnit := ⟨fun o => o.elim (inr PUnit.unit) inl, fun s => s.elim some fun _ => none, - fun o => by cases o <;> rfl, fun s => by - rcases s with (_ | ⟨⟨⟩⟩) <;> rfl⟩ + fun o => by cases o <;> rfl, + fun s => by rcases s with (_ | ⟨⟨⟩⟩) <;> rfl⟩ #align equiv.option_equiv_sum_punit Equiv.optionEquivSumPUnit @[simp] From 5e06896167343334bf646eff9d171ab02da1bcfa Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Thu, 24 Nov 2022 03:51:14 +1100 Subject: [PATCH 112/127] Update Mathlib/Logic/Equiv/Basic.lean Co-authored-by: Ruben Van de Velde <65514131+Ruben-VandeVelde@users.noreply.github.com> --- Mathlib/Logic/Equiv/Basic.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index bdaf0247eaa1e..d25228a67cffc 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -596,8 +596,8 @@ def subtypePreimage : { x : α → β // x ∘ Subtype.val = x₀ } ≃ ({ a // funext fun a => by dsimp only split_ifs - { rw [← hx]; rfl } - { rfl } + · rw [← hx]; rfl + · rfl right_inv x := funext fun ⟨a, h⟩ => show dite (p a) _ _ = _ by From d3d27f1ce8be832151781daa9400a8a9529b62d5 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Thu, 24 Nov 2022 03:53:07 +1100 Subject: [PATCH 113/127] syntax --- Mathlib/Logic/Equiv/Basic.lean | 22 ++++++++++++---------- 1 file changed, 12 insertions(+), 10 deletions(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 9276a350349e7..22da0b4e04ab1 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -361,11 +361,14 @@ theorem sumAssoc_symm_apply_inr_inr {α β γ} (c) : (sumAssoc α β γ).symm (i /-- Sum with `IsEmpty` is equivalent to the original type. -/ @[simps symm_apply] -def sumEmpty (α β) [IsEmpty β] : Sum α β ≃ α := - ⟨Sum.elim id isEmptyElim, inl, fun s => by +def sumEmpty (α β) [IsEmpty β] : Sum α β ≃ α where + toFun := Sum.elim id isEmptyElim + invFun := inl + left_inv s := by rcases s with (_ | x) - rfl - exact isEmptyElim x, fun a => rfl⟩ + · rfl + · exact isEmptyElim x + right_inv _ := rfl #align equiv.sum_empty Equiv.sumEmpty @[simp] @@ -818,12 +821,11 @@ section /-- The type of functions to a product `α × β` is equivalent to the type of pairs of functions `γ → α` and `γ → β`. -/ -def arrowProdEquivProdArrow (α β γ : Type _) : (γ → α × β) ≃ (γ → α) × (γ → β) := - ⟨fun f => (fun c => (f c).1, fun c => (f c).2), fun p c => (p.1 c, p.2 c), - fun f => funext fun c => Prod.mk.eta, - fun p => by - cases p - rfl⟩ +def arrowProdEquivProdArrow (α β γ : Type _) : (γ → α × β) ≃ (γ → α) × (γ → β) where + toFun := fun f => (fun c => (f c).1, fun c => (f c).2) + invFun := fun p c => (p.1 c, p.2 c) + left_inv := fun f => funext fun c => Prod.mk.eta + right_inv := fun p => by cases p; rfl #align equiv.arrow_prod_equiv_prod_arrow Equiv.arrowProdEquivProdArrow open Sum From 3fe9f4c7bef958e11a3028434197b86d2280ea62 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Thu, 24 Nov 2022 03:56:47 +1100 Subject: [PATCH 114/127] use where more often --- Mathlib/Logic/Equiv/Basic.lean | 32 ++++++++++++++++++-------------- 1 file changed, 18 insertions(+), 14 deletions(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 22da0b4e04ab1..dd35515ccf72f 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -1187,37 +1187,41 @@ def sigmaOptionEquivOfSome (p : Option α → Type v) (h : p none → False) : (sigmaSubtypeEquivOfSubset _ _ h').symm.trans (sigmaCongrLeft' (optionIsSomeEquiv α)) #align equiv.sigma_option_equiv_of_some Equiv.sigmaOptionEquivOfSome --- ericr: this definition doesn't seem nice indentation-wise; is this just the Lean4 style? /-- The `Pi`-type `∀ i, π i` is equivalent to the type of sections `f : ι → Σ i, π i` of the `Sigma` type such that for all `i` we have `(f i).fst = i`. -/ def piEquivSubtypeSigma (ι) (π : ι → Type _) : - (∀ i, π i) ≃ { f : ι → Σ i, π i // ∀ i, (f i).1 = i } := - ⟨fun f => ⟨fun i => ⟨i, f i⟩, fun i => rfl⟩, - fun f i => by rw [← f.2 i]; exact (f.1 i).2, - fun f => funext fun i => rfl, - fun ⟨f, hf⟩ => + (∀ i, π i) ≃ { f : ι → Σ i, π i // ∀ i, (f i).1 = i } where + toFun := fun f => ⟨fun i => ⟨i, f i⟩, fun i => rfl⟩ + invFun := fun f i => by rw [← f.2 i]; exact (f.1 i).2 + left_inv := fun f => funext fun i => rfl + right_inv := fun ⟨f, hf⟩ => Subtype.eq <| funext fun i => - Sigma.eq (hf i).symm <| eq_of_heq <| rec_heq_of_heq _ <| by simp⟩ + Sigma.eq (hf i).symm <| eq_of_heq <| rec_heq_of_heq _ <| by simp #align equiv.pi_equiv_subtype_sigma Equiv.piEquivSubtypeSigma /-- The type of functions `f : ∀ a, β a` such that for all `a` we have `p a (f a)` is equivalent to the type of functions `∀ a, {b : β a // p a b}`. -/ def subtypePiEquivPi {β : α → Sort v} {p : ∀ a, β a → Prop} : - { f : ∀ a, β a // ∀ a, p a (f a) } ≃ ∀ a, { b : β a // p a b } := - ⟨fun f a => ⟨f.1 a, f.2 a⟩, fun f => ⟨fun a => (f a).1, fun a => (f a).2⟩, by + { f : ∀ a, β a // ∀ a, p a (f a) } ≃ ∀ a, { b : β a // p a b } where + toFun := fun f a => ⟨f.1 a, f.2 a⟩ + invFun := fun f => ⟨fun a => (f a).1, fun a => (f a).2⟩ + left_inv := by rintro ⟨f, h⟩ - rfl, by + rfl + right_inv := by rintro f funext a - exact Subtype.ext_val rfl⟩ + exact Subtype.ext_val rfl #align equiv.subtype_pi_equiv_pi Equiv.subtypePiEquivPi /-- A subtype of a product defined by componentwise conditions is equivalent to a product of subtypes. -/ def subtypeProdEquivProd {p : α → Prop} {q : β → Prop} : - { c : α × β // p c.1 ∧ q c.2 } ≃ { a // p a } × { b // q b } := - ⟨fun x => ⟨⟨x.1.1, x.2.1⟩, ⟨x.1.2, x.2.2⟩⟩, fun x => ⟨⟨x.1.1, x.2.1⟩, ⟨x.1.2, x.2.2⟩⟩, - fun ⟨⟨_, _⟩, ⟨_, _⟩⟩ => rfl, fun ⟨⟨_, _⟩, ⟨_, _⟩⟩ => rfl⟩ + { c : α × β // p c.1 ∧ q c.2 } ≃ { a // p a } × { b // q b } where + toFun := fun x => ⟨⟨x.1.1, x.2.1⟩, ⟨x.1.2, x.2.2⟩⟩ + invFun := fun x => ⟨⟨x.1.1, x.2.1⟩, ⟨x.1.2, x.2.2⟩⟩ + left_inv := fun ⟨⟨_, _⟩, ⟨_, _⟩⟩ => rfl + right_inv := fun ⟨⟨_, _⟩, ⟨_, _⟩⟩ => rfl #align equiv.subtype_prod_equiv_prod Equiv.subtypeProdEquivProd /-- A subtype of a `Prod` is equivalent to a sigma type whose fibers are subtypes. -/ From 7b61a9a7c9591d833ad51c65dd334af32b493fe5 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Thu, 24 Nov 2022 04:14:13 +1100 Subject: [PATCH 115/127] comment about mwe --- Mathlib/Logic/Equiv/Basic.lean | 2 ++ 1 file changed, 2 insertions(+) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index d84dd6d9892c2..434b8bfa8bd58 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -1425,7 +1425,9 @@ def subtypeQuotientEquivQuotientSubtype (p₁ : α → Prop) [s₁ : Setoid α] invFun a := Quotient.liftOn a (fun a => (⟨⟦a.1⟧, (hp₂ _).1 a.2⟩ : { x // p₂ x })) fun a b hab => Subtype.ext_val (Quotient.sound ((h _ _).1 hab)) + -- Porting note: -- for some reason, doing this as a `fun ⟨a, ha⟩ => ` block breaks things. + -- There is a stand-alone version of this at https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/logic.2Eequiv.2Ebasic.20mathlib4.23631/near/311869098 left_inv t := by rcases t with ⟨a, ha⟩; exact Quotient.inductionOn a (fun b hb => rfl) ha right_inv a := Quotient.inductionOn a fun ⟨a, ha⟩ => rfl #align equiv.subtype_quotient_equiv_quotient_subtype Equiv.subtypeQuotientEquivQuotientSubtype From 277e12ab84755bbb7ed4cdbc2c7a69f7400039d9 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Thu, 24 Nov 2022 04:22:52 +1100 Subject: [PATCH 116/127] don't touch manifest --- lean_packages/manifest.json | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/lean_packages/manifest.json b/lean_packages/manifest.json index 9a302046d3bf1..369179063781c 100644 --- a/lean_packages/manifest.json +++ b/lean_packages/manifest.json @@ -1,7 +1,7 @@ {"version": 2, "packages": [{"url": "https://github.com/mhuisi/lean4-cli", - "rev": "5a858c32963b6b19be0d477a30a1f4b6c120be7e", + "rev": "b88b0ab3ad5f1352d107d3e8740ff15dc1fbb5b1", "name": "Cli", "inputRev": "nightly"}, {"url": "https://github.com/hargonix/LeanInk", @@ -13,7 +13,7 @@ "name": "aesop", "inputRev": "master"}, {"url": "https://github.com/leanprover/doc-gen4", - "rev": "69d48b174dbe404f37fe7e572de12f0c21992729", + "rev": "664a86e08bba7f4415fb0fdaef04c4508af702d1", "name": "doc-gen4", "inputRev": "main"}, {"url": "https://github.com/xubaiw/Unicode.lean", @@ -21,7 +21,7 @@ "name": "Unicode", "inputRev": "main"}, {"url": "https://github.com/leanprover/std4", - "rev": "68581a4685ce2ae8906f237651b3381d1aaa8b41", + "rev": "eaac7d455caa50a9c493f9c0c1fd13aaadfe3b03", "name": "std", "inputRev": "main"}, {"url": "https://github.com/leanprover/lake", From d9109cd4cb5f2747d09b7ded3b5330e6d34c0e68 Mon Sep 17 00:00:00 2001 From: Kevin Buzzard Date: Wed, 23 Nov 2022 18:17:49 +0000 Subject: [PATCH 117/127] compiling --- Mathlib/Logic/Equiv/Option.lean | 40 ++++++++++++++++----------------- 1 file changed, 19 insertions(+), 21 deletions(-) diff --git a/Mathlib/Logic/Equiv/Option.lean b/Mathlib/Logic/Equiv/Option.lean index d6288fa9d6712..7c48a57df6895 100644 --- a/Mathlib/Logic/Equiv/Option.lean +++ b/Mathlib/Logic/Equiv/Option.lean @@ -162,43 +162,44 @@ theorem optionCongr_injective : Function.Injective (optionCongr : α ≃ β → Function.LeftInverse.injective removeNone_option_congr #align equiv.option_congr_injective Equiv.optionCongr_injective +set_option maxHeartbeats 600000 -- next def times out with default heartbeats (200000) + /-- Equivalences between `Option α` and `β` that send `none` to `x` are equivalent to equivalences between `α` and `{y : β // y ≠ x}`. -/ def option_subtype [DecidableEq β] (x : β) : { e : Option α ≃ β // e none = x } ≃ (α ≃ { y : β // y ≠ x }) where toFun e := { toFun := - fun a => ⟨e a, ((EquivLike.injective _).ne_iff' e.property).2 (some_ne_none _)⟩, + fun a => ⟨(e : Option α ≃ β) a, ((EquivLike.injective _).ne_iff' e.property).2 (some_ne_none _)⟩, invFun := fun b => get _ (ne_none_iff_isSome.1 (((EquivLike.injective _).ne_iff' ((apply_eq_iff_eq_symm_apply _).1 e.property).symm).2 b.property)), left_inv := fun a => by - rw [← some_inj, some_get, ← coe_def] + rw [← some_inj, some_get] exact symm_apply_apply (e : Option α ≃ β) a, right_inv := fun b => by ext - simp - exact apply_symm_apply _ _ } + simp } invFun e := - ⟨{ toFun := fun a => casesOn' a x (Subtype.val ∘ toFun ∘ e), - invFun := fun b => if h : b = x then none else e.symm ⟨b, h⟩, + ⟨{ toFun := fun a => casesOn' a x (Subtype.val ∘ e), + invFun := fun b => if h : b = x then none else e.symm ⟨b, h⟩, left_inv := fun a => by - cases a - · simp - - simp only [casesOn'_some, Function.comp_apply, Subtype.coe_eta, - symm_apply_apply, dite_eq_ite] - exact if_neg (e a).property, - right_inv := fun b => by by_cases h : b = x <;> simp [h] }, + cases a with + | none => simp + | some a => + simp only [casesOn'_some, Function.comp_apply, Subtype.coe_eta, + symm_apply_apply, dite_eq_ite] + exact if_neg (e a).property, + right_inv := fun b => by by_cases h : b = x <;> simp [h] }, rfl⟩ left_inv e := by ext a cases a · simpa using e.property.symm - · simpa + · simp right_inv e := by ext a @@ -230,11 +231,8 @@ theorem option_subtype_apply_symm_apply #align equiv.option_subtype_apply_symm_apply Equiv.option_subtype_apply_symm_apply @[simp] -theorem option_subtype_symm_apply_apply_coe - [DecidableEq β] - (x : β) - (e : α ≃ { y : β // y ≠ x }) - (a : α) : (option_subtype x).symm e a = e a := +theorem option_subtype_symm_apply_apply_coe [DecidableEq β] (x : β) (e : α ≃ { y : β // y ≠ x }) + (a : α) : ((option_subtype x).symm e : Option α ≃ β) a = e a := rfl #align equiv.option_subtype_symm_apply_apply_coe Equiv.option_subtype_symm_apply_apply_coe @@ -243,7 +241,7 @@ theorem option_subtype_symm_apply_apply_some [DecidableEq β] (x : β) (e : α ≃ { y : β // y ≠ x }) - (a : α) : (option_subtype x).symm e (some a) = e a := + (a : α) : ((option_subtype x).symm e : Option α ≃ β) (some a) = e a := rfl #align equiv.option_subtype_symm_apply_apply_some Equiv.option_subtype_symm_apply_apply_some @@ -251,7 +249,7 @@ theorem option_subtype_symm_apply_apply_some theorem option_subtype_symm_apply_apply_none [DecidableEq β] (x : β) - (e : α ≃ { y : β // y ≠ x }) : (option_subtype x).symm e none = x := + (e : α ≃ { y : β // y ≠ x }) : ((option_subtype x).symm e : Option α ≃ β) none = x := rfl #align equiv.option_subtype_symm_apply_apply_none Equiv.option_subtype_symm_apply_apply_none From 5e7eb5746a06c562b91cfd87d7eb06b2cd845a82 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Thu, 24 Nov 2022 06:15:12 +1100 Subject: [PATCH 118/127] cleanup --- Mathlib/Logic/Equiv/Basic.lean | 5 +---- 1 file changed, 1 insertion(+), 4 deletions(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index 434b8bfa8bd58..e0a718c495a43 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -1425,10 +1425,7 @@ def subtypeQuotientEquivQuotientSubtype (p₁ : α → Prop) [s₁ : Setoid α] invFun a := Quotient.liftOn a (fun a => (⟨⟦a.1⟧, (hp₂ _).1 a.2⟩ : { x // p₂ x })) fun a b hab => Subtype.ext_val (Quotient.sound ((h _ _).1 hab)) - -- Porting note: - -- for some reason, doing this as a `fun ⟨a, ha⟩ => ` block breaks things. - -- There is a stand-alone version of this at https://leanprover.zulipchat.com/#narrow/stream/287929-mathlib4/topic/logic.2Eequiv.2Ebasic.20mathlib4.23631/near/311869098 - left_inv t := by rcases t with ⟨a, ha⟩; exact Quotient.inductionOn a (fun b hb => rfl) ha + left_inv t := by exact fun ⟨a, ha⟩ => Quotient.inductionOn a (fun b hb => rfl) ha right_inv a := Quotient.inductionOn a fun ⟨a, ha⟩ => rfl #align equiv.subtype_quotient_equiv_quotient_subtype Equiv.subtypeQuotientEquivQuotientSubtype From 913e4699dbda3d4acb217d251eb7d88664af93c4 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Thu, 24 Nov 2022 06:25:29 +1100 Subject: [PATCH 119/127] oops --- Mathlib/Logic/Equiv/Basic.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/Mathlib/Logic/Equiv/Basic.lean b/Mathlib/Logic/Equiv/Basic.lean index e0a718c495a43..15d66351e5273 100644 --- a/Mathlib/Logic/Equiv/Basic.lean +++ b/Mathlib/Logic/Equiv/Basic.lean @@ -1425,7 +1425,7 @@ def subtypeQuotientEquivQuotientSubtype (p₁ : α → Prop) [s₁ : Setoid α] invFun a := Quotient.liftOn a (fun a => (⟨⟦a.1⟧, (hp₂ _).1 a.2⟩ : { x // p₂ x })) fun a b hab => Subtype.ext_val (Quotient.sound ((h _ _).1 hab)) - left_inv t := by exact fun ⟨a, ha⟩ => Quotient.inductionOn a (fun b hb => rfl) ha + left_inv := by exact fun ⟨a, ha⟩ => Quotient.inductionOn a (fun b hb => rfl) ha right_inv a := Quotient.inductionOn a fun ⟨a, ha⟩ => rfl #align equiv.subtype_quotient_equiv_quotient_subtype Equiv.subtypeQuotientEquivQuotientSubtype From 8bc0328718a3da7455921310cfd5f279e50cee99 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Thu, 24 Nov 2022 06:37:47 +1100 Subject: [PATCH 120/127] long line --- Mathlib/Logic/Equiv/Option.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Logic/Equiv/Option.lean b/Mathlib/Logic/Equiv/Option.lean index 7c48a57df6895..6732b1186ac7a 100644 --- a/Mathlib/Logic/Equiv/Option.lean +++ b/Mathlib/Logic/Equiv/Option.lean @@ -169,8 +169,8 @@ equivalences between `α` and `{y : β // y ≠ x}`. -/ def option_subtype [DecidableEq β] (x : β) : { e : Option α ≃ β // e none = x } ≃ (α ≃ { y : β // y ≠ x }) where toFun e := - { toFun := - fun a => ⟨(e : Option α ≃ β) a, ((EquivLike.injective _).ne_iff' e.property).2 (some_ne_none _)⟩, + { toFun := fun a => + ⟨(e : Option α ≃ β) a, ((EquivLike.injective _).ne_iff' e.property).2 (some_ne_none _)⟩, invFun := fun b => get _ (ne_none_iff_isSome.1 From 563a6bcd21b2beddc8ed56249588245837005700 Mon Sep 17 00:00:00 2001 From: Arien Malec Date: Fri, 25 Nov 2022 09:51:20 -0800 Subject: [PATCH 121/127] naming fix & minor docu fix --- Mathlib/Logic/Equiv/Option.lean | 20 ++++++++++---------- lake-packages/Qq | 1 + lake-packages/aesop | 1 + lake-packages/std | 1 + 4 files changed, 13 insertions(+), 10 deletions(-) create mode 160000 lake-packages/Qq create mode 160000 lake-packages/aesop create mode 160000 lake-packages/std diff --git a/Mathlib/Logic/Equiv/Option.lean b/Mathlib/Logic/Equiv/Option.lean index 6732b1186ac7a..6ebce3b81004f 100644 --- a/Mathlib/Logic/Equiv/Option.lean +++ b/Mathlib/Logic/Equiv/Option.lean @@ -14,7 +14,7 @@ import Mathlib.Logic.Equiv.Defs We define * `Equiv.optionCongr`: the `Option α ≃ Option β` constructed from `e : α ≃ β` by sending `none` to - `none`, and applying a `e` elsewhere. + `none`, and applying `e` elsewhere. * `Equiv.removeNone`: the `α ≃ β` constructed from `Option α ≃ Option β` by removing `none` from both sides. -/ @@ -120,18 +120,18 @@ def removeNone : α ≃ β where #align equiv.remove_none Equiv.removeNone @[simp] -theorem remove_none_symm : (removeNone e).symm = removeNone e.symm := +theorem removeNone_symm : (removeNone e).symm = removeNone e.symm := rfl -#align equiv.remove_none_symm Equiv.remove_none_symm +#align equiv.remove_none_symm Equiv.removeNone_symm -theorem remove_none_some {x : α} (h : ∃ x', e (some x) = some x') : +theorem removeNone_some {x : α} (h : ∃ x', e (some x) = some x') : some (removeNone e x) = e (some x) := removeNone_aux_some e h -#align equiv.remove_none_some Equiv.remove_none_some +#align equiv.remove_none_some Equiv.removeNone_some -theorem remove_none_none {x : α} (h : e (some x) = none) : some (removeNone e x) = e none := +theorem removeNone_none {x : α} (h : e (some x) = none) : some (removeNone e x) = e none := removeNone_aux_none e h -#align equiv.remove_none_none Equiv.remove_none_none +#align equiv.remove_none_none Equiv.removeNone_none @[simp] theorem option_symm_apply_none_iff : e.symm none = none ↔ e none = none := @@ -140,10 +140,10 @@ theorem option_symm_apply_none_iff : e.symm none = none ↔ e none = none := theorem some_removeNone_iff {x : α} : some (removeNone e x) = e none ↔ e.symm none = some x := by cases' h : e (some x) with a - · rw [remove_none_none _ h] + · rw [removeNone_none _ h] simpa using (congr_arg e.symm h).symm - · rw [remove_none_some _ ⟨a, h⟩] + · rw [removeNone_some _ ⟨a, h⟩] have h1 := congr_arg e.symm h rw [symm_apply_apply] at h1 simp only [false_iff_iff, apply_eq_iff_eq] @@ -153,7 +153,7 @@ theorem some_removeNone_iff {x : α} : some (removeNone e x) = e none ↔ e.symm @[simp] theorem removeNone_option_congr (e : α ≃ β) : removeNone e.optionCongr = e := - Equiv.ext fun x => Option.some_injective _ <| remove_none_some _ ⟨e x, by simp [EquivFunctor.map]⟩ + Equiv.ext fun x => Option.some_injective _ <| removeNone_some _ ⟨e x, by simp [EquivFunctor.map]⟩ #align equiv.remove_none_option_congr Equiv.removeNone_option_congr end RemoveNone diff --git a/lake-packages/Qq b/lake-packages/Qq new file mode 160000 index 0000000000000..7ac99aa3fac48 --- /dev/null +++ b/lake-packages/Qq @@ -0,0 +1 @@ +Subproject commit 7ac99aa3fac487bec1d5860e751b99c7418298cf diff --git a/lake-packages/aesop b/lake-packages/aesop new file mode 160000 index 0000000000000..2e457511b82e4 --- /dev/null +++ b/lake-packages/aesop @@ -0,0 +1 @@ +Subproject commit 2e457511b82e4cedfdce86a81045711f747f32b5 diff --git a/lake-packages/std b/lake-packages/std new file mode 160000 index 0000000000000..c7a08a351d39c --- /dev/null +++ b/lake-packages/std @@ -0,0 +1 @@ +Subproject commit c7a08a351d39cdcb014c0bdbbc41550b436d54f5 From c01906947f7fbf855b88ac6f520d07ca60e11e81 Mon Sep 17 00:00:00 2001 From: Arien Malec Date: Fri, 25 Nov 2022 09:52:32 -0800 Subject: [PATCH 122/127] naming fix per review --- Mathlib/Logic/Equiv/Option.lean | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/Mathlib/Logic/Equiv/Option.lean b/Mathlib/Logic/Equiv/Option.lean index 6ebce3b81004f..269d11dca16a8 100644 --- a/Mathlib/Logic/Equiv/Option.lean +++ b/Mathlib/Logic/Equiv/Option.lean @@ -152,14 +152,14 @@ theorem some_removeNone_iff {x : α} : some (removeNone e x) = e none ↔ e.symm #align equiv.some_remove_none_iff Equiv.some_removeNone_iff @[simp] -theorem removeNone_option_congr (e : α ≃ β) : removeNone e.optionCongr = e := +theorem removeNone_optionCongr (e : α ≃ β) : removeNone e.optionCongr = e := Equiv.ext fun x => Option.some_injective _ <| removeNone_some _ ⟨e x, by simp [EquivFunctor.map]⟩ -#align equiv.remove_none_option_congr Equiv.removeNone_option_congr +#align equiv.remove_none_option_congr Equiv.removeNone_optionCongr end RemoveNone theorem optionCongr_injective : Function.Injective (optionCongr : α ≃ β → Option α ≃ Option β) := - Function.LeftInverse.injective removeNone_option_congr + Function.LeftInverse.injective removeNone_optionCongr #align equiv.option_congr_injective Equiv.optionCongr_injective set_option maxHeartbeats 600000 -- next def times out with default heartbeats (200000) From 1b7ea460f52e5dbb81ec9cf06bc8b57e36eb3c5b Mon Sep 17 00:00:00 2001 From: Arien Malec Date: Fri, 25 Nov 2022 09:59:15 -0800 Subject: [PATCH 123/127] fix naming issue --- Mathlib/Logic/Equiv/Option.lean | 50 ++++++++++++++++----------------- 1 file changed, 25 insertions(+), 25 deletions(-) diff --git a/Mathlib/Logic/Equiv/Option.lean b/Mathlib/Logic/Equiv/Option.lean index 269d11dca16a8..f3f04b864d603 100644 --- a/Mathlib/Logic/Equiv/Option.lean +++ b/Mathlib/Logic/Equiv/Option.lean @@ -166,7 +166,7 @@ set_option maxHeartbeats 600000 -- next def times out with default heartbeats (2 /-- Equivalences between `Option α` and `β` that send `none` to `x` are equivalent to equivalences between `α` and `{y : β // y ≠ x}`. -/ -def option_subtype [DecidableEq β] (x : β) : +def optionSubtype [DecidableEq β] (x : β) : { e : Option α ≃ β // e none = x } ≃ (α ≃ { y : β // y ≠ x }) where toFun e := { toFun := fun a => @@ -204,61 +204,61 @@ def option_subtype [DecidableEq β] (x : β) : right_inv e := by ext a rfl -#align equiv.option_subtype Equiv.option_subtype +#align equiv.option_subtype Equiv.optionSubtype @[simp] -theorem option_subtype_apply_apply +theorem optionSubtype_apply_apply [DecidableEq β] (x : β) (e : { e : Option α ≃ β // e none = x }) (a : α) - (h) : option_subtype x e a = ⟨(e : Option α ≃ β) a, h⟩ := rfl -#align equiv.option_subtype_apply_apply Equiv.option_subtype_apply_apply + (h) : optionSubtype x e a = ⟨(e : Option α ≃ β) a, h⟩ := rfl +#align equiv.option_subtype_apply_apply Equiv.optionSubtype_apply_apply @[simp] -theorem coe_option_subtype_apply_apply +theorem coe_optionSubtype_apply_apply [DecidableEq β] (x : β) (e : { e : Option α ≃ β // e none = x }) - (a : α) : ↑(option_subtype x e a) = (e : Option α ≃ β) a := rfl -#align equiv.coe_option_subtype_apply_apply Equiv.coe_option_subtype_apply_apply + (a : α) : ↑(optionSubtype x e a) = (e : Option α ≃ β) a := rfl +#align equiv.coe_option_subtype_apply_apply Equiv.coe_optionSubtype_apply_apply @[simp] -theorem option_subtype_apply_symm_apply +theorem optionSubtype_apply_symm_apply [DecidableEq β] (x : β) (e : { e : Option α ≃ β // e none = x }) - (b : { y : β // y ≠ x }) : ↑((option_subtype x e).symm b) = (e : Option α ≃ β).symm b := by - dsimp only [option_subtype] + (b : { y : β // y ≠ x }) : ↑((optionSubtype x e).symm b) = (e : Option α ≃ β).symm b := by + dsimp only [optionSubtype] simp -#align equiv.option_subtype_apply_symm_apply Equiv.option_subtype_apply_symm_apply +#align equiv.option_subtype_apply_symm_apply Equiv.optionSubtype_apply_symm_apply @[simp] -theorem option_subtype_symm_apply_apply_coe [DecidableEq β] (x : β) (e : α ≃ { y : β // y ≠ x }) - (a : α) : ((option_subtype x).symm e : Option α ≃ β) a = e a := +theorem optionSubtype_symm_apply_apply_coe [DecidableEq β] (x : β) (e : α ≃ { y : β // y ≠ x }) + (a : α) : ((optionSubtype x).symm e : Option α ≃ β) a = e a := rfl -#align equiv.option_subtype_symm_apply_apply_coe Equiv.option_subtype_symm_apply_apply_coe +#align equiv.option_subtype_symm_apply_apply_coe Equiv.optionSubtype_symm_apply_apply_coe @[simp] -theorem option_subtype_symm_apply_apply_some +theorem optionSubtype_symm_apply_apply_some [DecidableEq β] (x : β) (e : α ≃ { y : β // y ≠ x }) - (a : α) : ((option_subtype x).symm e : Option α ≃ β) (some a) = e a := + (a : α) : ((optionSubtype x).symm e : Option α ≃ β) (some a) = e a := rfl -#align equiv.option_subtype_symm_apply_apply_some Equiv.option_subtype_symm_apply_apply_some +#align equiv.option_subtype_symm_apply_apply_some Equiv.optionSubtype_symm_apply_apply_some @[simp] -theorem option_subtype_symm_apply_apply_none +theorem optionSubtype_symm_apply_apply_none [DecidableEq β] (x : β) - (e : α ≃ { y : β // y ≠ x }) : ((option_subtype x).symm e : Option α ≃ β) none = x := + (e : α ≃ { y : β // y ≠ x }) : ((optionSubtype x).symm e : Option α ≃ β) none = x := rfl -#align equiv.option_subtype_symm_apply_apply_none Equiv.option_subtype_symm_apply_apply_none +#align equiv.option_subtype_symm_apply_apply_none Equiv.optionSubtype_symm_apply_apply_none @[simp] -theorem option_subtype_symm_apply_symm_apply [DecidableEq β] (x : β) (e : α ≃ { y : β // y ≠ x }) - (b : { y : β // y ≠ x }) : ((option_subtype x).symm e : Option α ≃ β).symm b = e.symm b := by - simp only [option_subtype, coe_fn_symm_mk, Subtype.coe_mk, +theorem optionSubtype_symm_apply_symm_apply [DecidableEq β] (x : β) (e : α ≃ { y : β // y ≠ x }) + (b : { y : β // y ≠ x }) : ((optionSubtype x).symm e : Option α ≃ β).symm b = e.symm b := by + simp only [optionSubtype, coe_fn_symm_mk, Subtype.coe_mk, Subtype.coe_eta, dite_eq_ite, ite_eq_right_iff] exact fun h => False.elim (b.property h) -#align equiv.option_subtype_symm_apply_symm_apply Equiv.option_subtype_symm_apply_symm_apply +#align equiv.option_subtype_symm_apply_symm_apply Equiv.optionSubtype_symm_apply_symm_apply end Equiv From 205213bb0120c96a7ea36689a27c9fad72270056 Mon Sep 17 00:00:00 2001 From: Arien Malec Date: Fri, 25 Nov 2022 10:00:06 -0800 Subject: [PATCH 124/127] fix indentation issue --- Mathlib/Logic/Equiv/Option.lean | 14 +++++++------- 1 file changed, 7 insertions(+), 7 deletions(-) diff --git a/Mathlib/Logic/Equiv/Option.lean b/Mathlib/Logic/Equiv/Option.lean index f3f04b864d603..64d316215b332 100644 --- a/Mathlib/Logic/Equiv/Option.lean +++ b/Mathlib/Logic/Equiv/Option.lean @@ -208,17 +208,17 @@ def optionSubtype [DecidableEq β] (x : β) : @[simp] theorem optionSubtype_apply_apply - [DecidableEq β] (x : β) - (e : { e : Option α ≃ β // e none = x }) - (a : α) - (h) : optionSubtype x e a = ⟨(e : Option α ≃ β) a, h⟩ := rfl + [DecidableEq β] (x : β) + (e : { e : Option α ≃ β // e none = x }) + (a : α) + (h) : optionSubtype x e a = ⟨(e : Option α ≃ β) a, h⟩ := rfl #align equiv.option_subtype_apply_apply Equiv.optionSubtype_apply_apply @[simp] theorem coe_optionSubtype_apply_apply - [DecidableEq β] (x : β) - (e : { e : Option α ≃ β // e none = x }) - (a : α) : ↑(optionSubtype x e a) = (e : Option α ≃ β) a := rfl + [DecidableEq β] (x : β) + (e : { e : Option α ≃ β // e none = x }) + (a : α) : ↑(optionSubtype x e a) = (e : Option α ≃ β) a := rfl #align equiv.coe_option_subtype_apply_apply Equiv.coe_optionSubtype_apply_apply @[simp] From 1c2551978b824517daad95b0cef23cace9766d2b Mon Sep 17 00:00:00 2001 From: Arien Malec Date: Sun, 27 Nov 2022 20:11:29 -0800 Subject: [PATCH 125/127] fixed naming (+ remove lake_packages) --- Mathlib/Logic/Equiv/Option.lean | 4 ++-- lake-packages/Qq | 1 - lake-packages/aesop | 1 - lake-packages/std | 1 - 4 files changed, 2 insertions(+), 5 deletions(-) delete mode 160000 lake-packages/Qq delete mode 160000 lake-packages/aesop delete mode 160000 lake-packages/std diff --git a/Mathlib/Logic/Equiv/Option.lean b/Mathlib/Logic/Equiv/Option.lean index 64d316215b332..200ebdbdfc084 100644 --- a/Mathlib/Logic/Equiv/Option.lean +++ b/Mathlib/Logic/Equiv/Option.lean @@ -55,10 +55,10 @@ theorem optionCongr_trans (e₁ : α ≃ β) (e₂ : β ≃ γ) : /-- When `α` and `β` are in the same universe, this is the same as the result of `EquivFunctor.map_equiv`. -/ -theorem optionCongr_eq_equivFunction_mapEquiv {α β : Type _} (e : α ≃ β) : +theorem optionCongr_eq_equivFunctor_mapEquiv {α β : Type _} (e : α ≃ β) : optionCongr e = EquivFunctor.mapEquiv Option e := rfl -#align equiv.option_congr_eq_equiv_function_map_equiv Equiv.optionCongr_eq_equivFunction_mapEquiv +#align equiv.option_congr_eq_equiv_function_map_equiv Equiv.optionCongr_eq_equivFunctor_mapEquiv end OptionCongr diff --git a/lake-packages/Qq b/lake-packages/Qq deleted file mode 160000 index 7ac99aa3fac48..0000000000000 --- a/lake-packages/Qq +++ /dev/null @@ -1 +0,0 @@ -Subproject commit 7ac99aa3fac487bec1d5860e751b99c7418298cf diff --git a/lake-packages/aesop b/lake-packages/aesop deleted file mode 160000 index 2e457511b82e4..0000000000000 --- a/lake-packages/aesop +++ /dev/null @@ -1 +0,0 @@ -Subproject commit 2e457511b82e4cedfdce86a81045711f747f32b5 diff --git a/lake-packages/std b/lake-packages/std deleted file mode 160000 index c7a08a351d39c..0000000000000 --- a/lake-packages/std +++ /dev/null @@ -1 +0,0 @@ -Subproject commit c7a08a351d39cdcb014c0bdbbc41550b436d54f5 From 5e66077abbc407ab57700ff0c23e452839e8ba16 Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Tue, 29 Nov 2022 03:26:45 +1100 Subject: [PATCH 126/127] replace slow simp --- Mathlib/Logic/Equiv/Option.lean | 12 ++++++------ 1 file changed, 6 insertions(+), 6 deletions(-) diff --git a/Mathlib/Logic/Equiv/Option.lean b/Mathlib/Logic/Equiv/Option.lean index 200ebdbdfc084..27d70cb2159c9 100644 --- a/Mathlib/Logic/Equiv/Option.lean +++ b/Mathlib/Logic/Equiv/Option.lean @@ -162,8 +162,6 @@ theorem optionCongr_injective : Function.Injective (optionCongr : α ≃ β → Function.LeftInverse.injective removeNone_optionCongr #align equiv.option_congr_injective Equiv.optionCongr_injective -set_option maxHeartbeats 600000 -- next def times out with default heartbeats (200000) - /-- Equivalences between `Option α` and `β` that send `none` to `x` are equivalent to equivalences between `α` and `{y : β // y ≠ x}`. -/ def optionSubtype [DecidableEq β] (x : β) : @@ -192,15 +190,17 @@ def optionSubtype [DecidableEq β] (x : β) : simp only [casesOn'_some, Function.comp_apply, Subtype.coe_eta, symm_apply_apply, dite_eq_ite] exact if_neg (e a).property, - right_inv := fun b => by by_cases h : b = x <;> simp [h] }, + right_inv := fun b => by + by_cases h : b = x <;> simp [h] }, rfl⟩ left_inv e := by ext a cases a · simpa using e.property.symm - - · simp - + -- Porting note: this cases had been by `simpa`, + -- but `simp` here is mysteriously slow, even after squeezing. + -- `rfl` closes the goal quickly, so we use that. + · rfl right_inv e := by ext a rfl From 2527966ccb67fa783d97b3382ea4a98b584c94eb Mon Sep 17 00:00:00 2001 From: Scott Morrison Date: Tue, 29 Nov 2022 03:28:47 +1100 Subject: [PATCH 127/127] Apply suggestions from code review --- Mathlib/Logic/Equiv/Option.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/Mathlib/Logic/Equiv/Option.lean b/Mathlib/Logic/Equiv/Option.lean index 27d70cb2159c9..7f13a69f83603 100644 --- a/Mathlib/Logic/Equiv/Option.lean +++ b/Mathlib/Logic/Equiv/Option.lean @@ -28,7 +28,7 @@ variable {α β γ : Type _} section OptionCongr -/-- A universe-polymorphic version of `EquivFunctor.map_equiv Option e`. -/ +/-- A universe-polymorphic version of `EquivFunctor.mapEquiv Option e`. -/ @[simps apply] def optionCongr (e : α ≃ β) : Option α ≃ Option β where toFun := Option.map e @@ -54,7 +54,7 @@ theorem optionCongr_trans (e₁ : α ≃ β) (e₂ : β ≃ γ) : #align equiv.option_congr_trans Equiv.optionCongr_trans /-- When `α` and `β` are in the same universe, this is the same as the result of -`EquivFunctor.map_equiv`. -/ +`EquivFunctor.mapEquiv`. -/ theorem optionCongr_eq_equivFunctor_mapEquiv {α β : Type _} (e : α ≃ β) : optionCongr e = EquivFunctor.mapEquiv Option e := rfl