refactor(hott): replace 'assert'-expr with 'have'-expr
This commit is contained in:
parent
deb1b3dc79
commit
cc8d9bc7ff
33 changed files with 214 additions and 173 deletions
|
@ -69,12 +69,14 @@ namespace category
|
|||
@precategory.mk _ cone_hom
|
||||
abstract begin
|
||||
intro x y,
|
||||
assert H : cone_hom x y ≃ Σ(f : x ⟶ y), Πi, cone_to_nat y i ∘ f = cone_to_nat x i,
|
||||
{ fapply equiv.MK,
|
||||
have H : cone_hom x y ≃ Σ(f : x ⟶ y), Πi, cone_to_nat y i ∘ f = cone_to_nat x i,
|
||||
begin
|
||||
fapply equiv.MK,
|
||||
{ intro f, induction f, constructor, assumption},
|
||||
{ intro v, induction v, constructor, assumption},
|
||||
{ intro v, induction v, reflexivity},
|
||||
{ intro f, induction f, reflexivity}},
|
||||
{ intro f, induction f, reflexivity}
|
||||
end,
|
||||
apply is_trunc.is_trunc_equiv_closed_rev, exact H,
|
||||
fapply sigma.is_trunc_sigma, intros,
|
||||
apply is_trunc_succ, apply pi.is_trunc_pi, intros, esimp,
|
||||
|
|
|
@ -65,12 +65,14 @@ namespace category
|
|||
open equalizer_category_hom
|
||||
theorem is_set_equalizer_category_hom (b₁ b₂ : bool) : is_set (equalizer_category_hom b₁ b₂) :=
|
||||
begin
|
||||
assert H : Πb b', equalizer_category_hom b b' ≃ bool.rec (bool.rec empty bool) (λb, empty) b b',
|
||||
{ intro b b', fapply equiv.MK,
|
||||
have H : Πb b', equalizer_category_hom b b' ≃ bool.rec (bool.rec empty bool) (λb, empty) b b',
|
||||
begin
|
||||
intro b b', fapply equiv.MK,
|
||||
{ intro x, induction x, exact ff, exact tt},
|
||||
{ intro v, induction b: induction b': induction v, exact f1, exact f2},
|
||||
{ intro v, induction b: induction b': induction v: reflexivity},
|
||||
{ intro x, induction x: reflexivity}},
|
||||
{ intro x, induction x: reflexivity}
|
||||
end,
|
||||
apply is_trunc_equiv_closed_rev, apply H,
|
||||
induction b₁: induction b₂: exact _
|
||||
end
|
||||
|
@ -110,14 +112,16 @@ namespace category
|
|||
theorem is_set_pullback_category_hom (b₁ b₂ : pullback_category_ob)
|
||||
: is_set (pullback_category_hom b₁ b₂) :=
|
||||
begin
|
||||
assert H : Πb b', pullback_category_hom b b' ≃
|
||||
have H : Πb b', pullback_category_hom b b' ≃
|
||||
pullback_category_ob.rec (λb, empty) (λb, empty)
|
||||
(pullback_category_ob.rec unit unit empty) b' b,
|
||||
{ intro b b', fapply equiv.MK,
|
||||
begin
|
||||
intro b b', fapply equiv.MK,
|
||||
{ intro x, induction x: exact star},
|
||||
{ intro v, induction b: induction b': induction v, exact f1, exact f2},
|
||||
{ intro v, induction b: induction b': induction v: reflexivity},
|
||||
{ intro x, induction x: reflexivity}},
|
||||
{ intro x, induction x: reflexivity}
|
||||
end,
|
||||
apply is_trunc_equiv_closed_rev, apply H,
|
||||
induction b₁: induction b₂: exact _
|
||||
end
|
||||
|
|
|
@ -66,8 +66,8 @@ namespace category
|
|||
ua !equiv_equiv_iso
|
||||
|
||||
definition is_univalent_Set (A B : set) : is_equiv (iso_of_eq : A = B → A ≅ B) :=
|
||||
assert H₁ : is_equiv (@iso_of_equiv A B ∘ @equiv_of_eq A B ∘ subtype_eq_inv _ _ ∘
|
||||
@ap _ _ (to_fun (trunctype.sigma_char 0)) A B), from
|
||||
have H₁ : is_equiv (@iso_of_equiv A B ∘ @equiv_of_eq A B ∘ subtype_eq_inv _ _ ∘
|
||||
@ap _ _ (to_fun (trunctype.sigma_char 0)) A B), from
|
||||
@is_equiv_compose _ _ _ _ _
|
||||
(@is_equiv_compose _ _ _ _ _
|
||||
(@is_equiv_compose _ _ _ _ _
|
||||
|
|
|
@ -48,16 +48,19 @@ namespace category
|
|||
begin
|
||||
apply is_prop.mk,
|
||||
intro G G', cases G with G η ε H K, cases G' with G' η' ε' H' K',
|
||||
assert lem₁ : Π(p : G = G'), p ▸ η = η' → p ▸ ε = ε'
|
||||
have lem₁ : Π(p : G = G'), p ▸ η = η' → p ▸ ε = ε'
|
||||
→ is_left_adjoint.mk G η ε H K = is_left_adjoint.mk G' η' ε' H' K',
|
||||
{ intros p q r, induction p, induction q, induction r, esimp,
|
||||
apply apd011 (is_left_adjoint.mk G η ε) !is_prop.elim !is_prop.elim},
|
||||
assert lem₂ : Π (d : carrier D),
|
||||
begin
|
||||
intros p q r, induction p, induction q, induction r, esimp,
|
||||
apply apd011 (is_left_adjoint.mk G η ε) !is_prop.elim !is_prop.elim
|
||||
end,
|
||||
have lem₂ : Π (d : carrier D),
|
||||
(to_fun_hom G (natural_map ε' d) ∘
|
||||
natural_map η (to_fun_ob G' d)) ∘
|
||||
to_fun_hom G' (natural_map ε d) ∘
|
||||
natural_map η' (to_fun_ob G d) = id,
|
||||
{ intro d, esimp,
|
||||
begin
|
||||
intro d, esimp,
|
||||
rewrite [assoc],
|
||||
rewrite [-assoc (G (ε' d))],
|
||||
esimp, rewrite [nf_fn_eq_fn_nf_pt' G' ε η d],
|
||||
|
@ -70,13 +73,15 @@ namespace category
|
|||
rewrite [↑functor.compose, -respect_comp G],
|
||||
rewrite [H' (G d)],
|
||||
rewrite [respect_id,▸*,id_right],
|
||||
apply K},
|
||||
assert lem₃ : Π (d : carrier D),
|
||||
apply K
|
||||
end,
|
||||
have lem₃ : Π (d : carrier D),
|
||||
(to_fun_hom G' (natural_map ε d) ∘
|
||||
natural_map η' (to_fun_ob G d)) ∘
|
||||
to_fun_hom G (natural_map ε' d) ∘
|
||||
natural_map η (to_fun_ob G' d) = id,
|
||||
{ intro d, esimp,
|
||||
begin
|
||||
intro d, esimp,
|
||||
rewrite [assoc, -assoc (G' (ε d))],
|
||||
esimp, rewrite [nf_fn_eq_fn_nf_pt' G ε' η' d],
|
||||
esimp, rewrite [assoc], esimp, rewrite [-assoc],
|
||||
|
@ -88,7 +93,8 @@ namespace category
|
|||
rewrite [↑functor.compose, -respect_comp G'],
|
||||
rewrite [H (G' d)],
|
||||
rewrite [respect_id,▸*,id_right],
|
||||
apply K'},
|
||||
apply K'
|
||||
end,
|
||||
fapply lem₁,
|
||||
{ fapply functor.eq_of_pointwise_iso,
|
||||
{ fapply change_natural_map,
|
||||
|
|
|
@ -47,8 +47,9 @@ namespace category
|
|||
begin
|
||||
fconstructor,
|
||||
{ exact (to_fun_hom F)⁻¹ᶠ f},
|
||||
{ assert H : is_iso (F ((to_fun_hom F)⁻¹ᶠ f)),
|
||||
{ have H' : is_iso (to_hom f), from _, exact (right_inv (to_fun_hom F) (to_hom f))⁻¹ ▸ H'},
|
||||
{ have H : is_iso (F ((to_fun_hom F)⁻¹ᶠ f)), from
|
||||
have H' : is_iso (to_hom f), from _,
|
||||
(right_inv (to_fun_hom F) (to_hom f))⁻¹ ▸ H',
|
||||
exact reflect_is_iso F _},
|
||||
end
|
||||
|
||||
|
|
|
@ -231,13 +231,15 @@ namespace category
|
|||
theorem is_prop_is_equivalence [instance] {C : Category} {D : Precategory} (F : C ⇒ D)
|
||||
: is_prop (is_equivalence F) :=
|
||||
begin
|
||||
assert f : is_equivalence F ≃ Σ(H : is_left_adjoint F), is_iso (unit F) × is_iso (counit F),
|
||||
{ fapply equiv.MK,
|
||||
have f : is_equivalence F ≃ Σ(H : is_left_adjoint F), is_iso (unit F) × is_iso (counit F),
|
||||
begin
|
||||
fapply equiv.MK,
|
||||
{ intro H, induction H, fconstructor: constructor, repeat (esimp;assumption) },
|
||||
{ intro H, induction H with H1 H2, induction H1, induction H2, constructor,
|
||||
repeat (esimp at *;assumption)},
|
||||
{ intro H, induction H with H1 H2, induction H1, induction H2, reflexivity},
|
||||
{ intro H, induction H, reflexivity}},
|
||||
{ intro H, induction H, reflexivity}
|
||||
end,
|
||||
apply is_trunc_equiv_closed_rev, exact f,
|
||||
end
|
||||
|
||||
|
|
|
@ -18,10 +18,10 @@ namespace category
|
|||
definition functor_limit_object [constructor]
|
||||
[H : has_limits_of_shape D I] (F : I ⇒ D ^c C) : C ⇒ D :=
|
||||
begin
|
||||
assert lem : Π(c d : carrier C) (f : hom c d) ⦃i j : carrier I⦄ (k : i ⟶ j),
|
||||
have lem : Π(c d : carrier C) (f : hom c d) ⦃i j : carrier I⦄ (k : i ⟶ j),
|
||||
(constant2_functor F d) k ∘ to_fun_hom (F i) f ∘ limit_morphism (constant2_functor F c) i =
|
||||
to_fun_hom (F j) f ∘ limit_morphism (constant2_functor F c) j,
|
||||
{ intro c d f i j k, rewrite [-limit_commute _ k,▸*,+assoc,▸*,-naturality (F k) f]},
|
||||
begin intro c d f i j k, rewrite [-limit_commute _ k,▸*,+assoc,▸*,-naturality (F k) f] end,
|
||||
|
||||
fapply functor.mk,
|
||||
{ intro c, exact limit_object (constant2_functor F c)},
|
||||
|
|
|
@ -76,11 +76,13 @@ namespace category
|
|||
: preserves_existing_limits (yoneda_embedding C) :=
|
||||
begin
|
||||
intro I F H Gη, induction H with y H, induction Gη with G η, esimp at *,
|
||||
assert lem : Π (i : carrier I),
|
||||
have lem : Π (i : carrier I),
|
||||
nat_trans_hom_functor_left (natural_map (cone_to_nat y) i)
|
||||
∘n preserves_existing_limits_yoneda_embedding_lemma y η = natural_map η i,
|
||||
{ intro i, apply nat_trans_eq, intro c, apply eq_of_homotopy, intro x,
|
||||
esimp, refine !assoc ⬝ !id_right ⬝ !to_hom_limit_commute},
|
||||
begin
|
||||
intro i, apply nat_trans_eq, intro c, apply eq_of_homotopy, intro x,
|
||||
esimp, refine !assoc ⬝ !id_right ⬝ !to_hom_limit_commute
|
||||
end,
|
||||
fapply is_contr.mk,
|
||||
{ fapply cone_hom.mk,
|
||||
{ exact preserves_existing_limits_yoneda_embedding_lemma y η},
|
||||
|
@ -104,7 +106,7 @@ namespace category
|
|||
[H : is_left_adjoint F] : preserves_existing_limits F :=
|
||||
begin
|
||||
intro I G K dη, induction K with y K, induction dη with d η, esimp at *,
|
||||
-- assert lem : Π (i : carrier I),
|
||||
-- have lem : Π (i : carrier I),
|
||||
-- nat_trans_hom_functor_left (natural_map (cone_to_nat y) i)
|
||||
-- ∘n preserves_existing_limits_yoneda_embedding_lemma y η = natural_map η i,
|
||||
-- { intro i, apply nat_trans_eq, intro c, apply eq_of_homotopy, intro x,
|
||||
|
|
|
@ -57,8 +57,8 @@ namespace category
|
|||
: is_prop (has_terminal_object D) :=
|
||||
begin
|
||||
apply is_prop.mk, intro t₁ t₂, induction t₁ with d₁ H₁, induction t₂ with d₂ H₂,
|
||||
assert p : d₁ = d₂,
|
||||
{ apply eq_of_iso, apply terminal_iso_terminal},
|
||||
have p : d₁ = d₂,
|
||||
begin apply eq_of_iso, apply terminal_iso_terminal end,
|
||||
induction p, exact ap _ !is_prop.elim
|
||||
end
|
||||
|
||||
|
|
|
@ -173,11 +173,11 @@ namespace category
|
|||
esimp at *,
|
||||
revert q, eapply homotopy2.rec_on @p, esimp, clear p, intro p q, induction p,
|
||||
esimp at *,
|
||||
assert H : comp1 = comp2,
|
||||
{ apply eq_of_homotopy3, intros, apply eq_of_homotopy2, intros, apply q},
|
||||
have H : comp1 = comp2,
|
||||
begin apply eq_of_homotopy3, intros, apply eq_of_homotopy2, intros, apply q end,
|
||||
induction H,
|
||||
assert K : ID1 = ID2,
|
||||
{ apply eq_of_homotopy, intro a, exact !ir'⁻¹ ⬝ !il},
|
||||
have K : ID1 = ID2,
|
||||
begin apply eq_of_homotopy, intro a, exact !ir'⁻¹ ⬝ !il end,
|
||||
induction K,
|
||||
apply ap0111111 (precategory.mk' hom1 comp1 ID1): apply is_prop.elim
|
||||
end
|
||||
|
@ -197,19 +197,21 @@ namespace category
|
|||
/- if we need to prove properties about precategory_eq, it might be easier with the following proof:
|
||||
begin
|
||||
induction C with hom1 comp1 ID1, induction D with hom2 comp2 ID2, esimp at *,
|
||||
assert H : Σ(s : hom1 = hom2), (λa b, equiv_of_eq (apd100 s a b)) = p,
|
||||
{ fconstructor,
|
||||
have H : Σ(s : hom1 = hom2), (λa b, equiv_of_eq (apd100 s a b)) = p,
|
||||
begin
|
||||
fconstructor,
|
||||
{ apply eq_of_homotopy2, intros, apply ua, apply p},
|
||||
{ apply eq_of_homotopy2, intros, rewrite [to_right_inv !eq_equiv_homotopy2, equiv_of_eq_ua]}},
|
||||
{ apply eq_of_homotopy2, intros, rewrite [to_right_inv !eq_equiv_homotopy2, equiv_of_eq_ua]}
|
||||
end,
|
||||
induction H with H1 H2, induction H1, esimp at H2,
|
||||
assert K : (λa b, equiv.refl) = p,
|
||||
{ refine _ ⬝ H2, apply eq_of_homotopy2, intros, exact !equiv_of_eq_refl⁻¹},
|
||||
have K : (λa b, equiv.refl) = p,
|
||||
begin refine _ ⬝ H2, apply eq_of_homotopy2, intros, exact !equiv_of_eq_refl⁻¹ end,
|
||||
induction K, clear H2,
|
||||
esimp at *,
|
||||
assert H : comp1 = comp2,
|
||||
{ apply eq_of_homotopy3, intros, apply eq_of_homotopy2, intros, apply q},
|
||||
assert K : ID1 = ID2,
|
||||
{ apply eq_of_homotopy, intros, apply r},
|
||||
have H : comp1 = comp2,
|
||||
begin apply eq_of_homotopy3, intros, apply eq_of_homotopy2, intros, apply q end,
|
||||
have K : ID1 = ID2,
|
||||
begin apply eq_of_homotopy, intros, apply r end,
|
||||
induction H, induction K,
|
||||
apply ap0111111 (precategory.mk' hom1 comp1 ID1): apply is_prop.elim
|
||||
end
|
||||
|
|
|
@ -243,7 +243,7 @@ section field
|
|||
by rewrite [mul.comm, (!div_mul_cancel Hb)]
|
||||
|
||||
theorem one_div_add_one_div (Ha : a ≠ 0) (Hb : b ≠ 0) : 1 / a + 1 / b = (a + b) / (a * b) :=
|
||||
assert a * b ≠ 0, from (division_ring.mul_ne_zero Ha Hb),
|
||||
have a * b ≠ 0, from (division_ring.mul_ne_zero Ha Hb),
|
||||
by rewrite [add.comm, -(field.div_mul_left Ha this), -(field.div_mul_right Hb this), *division.def,
|
||||
-right_distrib]
|
||||
|
||||
|
|
|
@ -304,7 +304,7 @@ section group
|
|||
|
||||
lemma is_conj.symm (a b : A) : a ~ b → b ~ a :=
|
||||
assume Pab, obtain x (Pconj : x ∘c b = a), from Pab,
|
||||
assert Pxinv : x⁻¹ ∘c x ∘c b = x⁻¹ ∘c a, begin congruence, assumption end,
|
||||
have Pxinv : x⁻¹ ∘c x ∘c b = x⁻¹ ∘c a, begin congruence, assumption end,
|
||||
sigma.mk x⁻¹ (inverse (conj_inv_cancel x b ▸ Pxinv))
|
||||
|
||||
lemma is_conj.trans (a b c : A) : a ~ b → b ~ c → a ~ c :=
|
||||
|
|
|
@ -166,16 +166,16 @@ section linear_ordered_field
|
|||
|
||||
theorem div_lt_div_of_mul_sub_mul_div_neg (Hc : c ≠ 0) (Hd : d ≠ 0)
|
||||
(H : (a * d - b * c) / (c * d) < 0) : a / c < b / d :=
|
||||
assert H1 : (a * d - c * b) / (c * d) < 0, by rewrite [mul.comm c b]; exact H,
|
||||
assert H2 : a / c - b / d < 0, by rewrite [!div_sub_div Hc Hd]; exact H1,
|
||||
assert H3 : a / c - b / d + b / d < 0 + b / d, from add_lt_add_right H2 _,
|
||||
have H1 : (a * d - c * b) / (c * d) < 0, by rewrite [mul.comm c b]; exact H,
|
||||
have H2 : a / c - b / d < 0, by rewrite [!div_sub_div Hc Hd]; exact H1,
|
||||
have H3 : a / c - b / d + b / d < 0 + b / d, from add_lt_add_right H2 _,
|
||||
begin rewrite [zero_add at H3, sub_eq_add_neg at H3, neg_add_cancel_right at H3], exact H3 end
|
||||
|
||||
theorem div_le_div_of_mul_sub_mul_div_nonpos (Hc : c ≠ 0) (Hd : d ≠ 0)
|
||||
(H : (a * d - b * c) / (c * d) ≤ 0) : a / c ≤ b / d :=
|
||||
assert H1 : (a * d - c * b) / (c * d) ≤ 0, by rewrite [mul.comm c b]; exact H,
|
||||
assert H2 : a / c - b / d ≤ 0, by rewrite [!div_sub_div Hc Hd]; exact H1,
|
||||
assert H3 : a / c - b / d + b / d ≤ 0 + b / d, from add_le_add_right H2 _,
|
||||
have H1 : (a * d - c * b) / (c * d) ≤ 0, by rewrite [mul.comm c b]; exact H,
|
||||
have H2 : a / c - b / d ≤ 0, by rewrite [!div_sub_div Hc Hd]; exact H1,
|
||||
have H3 : a / c - b / d + b / d ≤ 0 + b / d, from add_le_add_right H2 _,
|
||||
begin rewrite [zero_add at H3, sub_eq_add_neg at H3, neg_add_cancel_right at H3], exact H3 end
|
||||
|
||||
theorem div_pos_of_pos_of_pos (Ha : 0 < a) (Hb : 0 < b) : 0 < a / b :=
|
||||
|
@ -342,12 +342,12 @@ section linear_ordered_field
|
|||
|
||||
theorem exists_add_lt_prod_pos_of_lt (H : b < a) : Σ c : A, b + c < a × c > 0 :=
|
||||
sigma.mk ((a - b) / (1 + 1))
|
||||
(pair (assert H2 : a + a > (b + b) + (a - b), from calc
|
||||
(pair (have H2 : a + a > (b + b) + (a - b), from calc
|
||||
a + a > b + a : add_lt_add_right H
|
||||
... = b + a + b - b : add_sub_cancel
|
||||
... = b + b + a - b : add.right_comm
|
||||
... = (b + b) + (a - b) : add_sub,
|
||||
assert H3 : (a + a) / 2 > ((b + b) + (a - b)) / 2,
|
||||
have H3 : (a + a) / 2 > ((b + b) + (a - b)) / 2,
|
||||
from div_lt_div_of_lt_of_pos H2 two_pos,
|
||||
by rewrite [one_add_one_eq_two, sub_eq_add_neg, add_self_div_two at H3, -div_add_div_same at H3, add_self_div_two at H3];
|
||||
exact H3)
|
||||
|
@ -414,7 +414,7 @@ section discrete_linear_ordered_field
|
|||
), le_of_one_le_div Hb H'
|
||||
|
||||
theorem le_of_one_div_le_one_div_of_neg (H : b < 0) (Hl : 1 / a ≤ 1 / b) : b ≤ a :=
|
||||
assert Ha : a ≠ 0, from ne_of_lt (neg_of_one_div_neg (calc
|
||||
have Ha : a ≠ 0, from ne_of_lt (neg_of_one_div_neg (calc
|
||||
1 / a ≤ 1 / b : Hl
|
||||
... < 0 : one_div_neg_of_neg H)),
|
||||
have H' : -b > 0, from neg_pos_of_neg H,
|
||||
|
@ -504,7 +504,7 @@ section discrete_linear_ordered_field
|
|||
by rewrite [abs_of_neg H', abs_of_neg (one_div_neg_of_neg H'),
|
||||
-(division_ring.one_div_neg_eq_neg_one_div (ne_of_lt H'))]
|
||||
else
|
||||
assert Heq : a = 0, from eq_of_le_of_ge (le_of_not_gt H) (le_of_not_gt H'),
|
||||
have Heq : a = 0, from eq_of_le_of_ge (le_of_not_gt H) (le_of_not_gt H'),
|
||||
by rewrite [Heq, div_zero, *abs_zero, div_zero])
|
||||
|
||||
theorem sign_eq_div_abs (a : A) : sign a = a / (abs a) :=
|
||||
|
|
|
@ -201,12 +201,12 @@ structure ordered_comm_group [class] (A : Type) extends add_comm_group A, order_
|
|||
|
||||
theorem ordered_comm_group.le_of_add_le_add_left [s : ordered_comm_group A] {a b c : A}
|
||||
(H : a + b ≤ a + c) : b ≤ c :=
|
||||
assert H' : -a + (a + b) ≤ -a + (a + c), from ordered_comm_group.add_le_add_left _ _ H _,
|
||||
have H' : -a + (a + b) ≤ -a + (a + c), from ordered_comm_group.add_le_add_left _ _ H _,
|
||||
by rewrite *neg_add_cancel_left at H'; exact H'
|
||||
|
||||
theorem ordered_comm_group.lt_of_add_lt_add_left [s : ordered_comm_group A] {a b c : A}
|
||||
(H : a + b < a + c) : b < c :=
|
||||
assert H' : -a + (a + b) < -a + (a + c), from ordered_comm_group.add_lt_add_left _ _ H _,
|
||||
have H' : -a + (a + b) < -a + (a + c), from ordered_comm_group.add_lt_add_left _ _ H _,
|
||||
by rewrite *neg_add_cancel_left at H'; exact H'
|
||||
|
||||
definition ordered_comm_group.to_ordered_cancel_comm_monoid [trans_instance]
|
||||
|
@ -347,7 +347,7 @@ section
|
|||
iff.mp !add_le_iff_le_sub_right
|
||||
|
||||
theorem le_add_iff_neg_add_le : a ≤ b + c ↔ -b + a ≤ c :=
|
||||
assert H: a ≤ b + c ↔ -b + a ≤ -b + (b + c), from iff.symm (!add_le_add_left_iff),
|
||||
have H: a ≤ b + c ↔ -b + a ≤ -b + (b + c), from iff.symm (!add_le_add_left_iff),
|
||||
by rewrite neg_add_cancel_left at H; exact H
|
||||
|
||||
theorem le_add_of_neg_add_le {a b c : A} : -b + a ≤ c → a ≤ b + c :=
|
||||
|
@ -366,7 +366,7 @@ section
|
|||
iff.mp !le_add_iff_sub_left_le
|
||||
|
||||
theorem le_add_iff_sub_right_le : a ≤ b + c ↔ a - c ≤ b :=
|
||||
assert H: a ≤ b + c ↔ a - c ≤ b + c - c, from iff.symm (!add_le_add_right_iff),
|
||||
have H: a ≤ b + c ↔ a - c ≤ b + c - c, from iff.symm (!add_le_add_right_iff),
|
||||
by rewrite [sub_eq_add_neg (b+c) c at H, add_neg_cancel_right at H]; exact H
|
||||
|
||||
theorem le_add_of_sub_right_le {a b c : A} : a - c ≤ b → a ≤ b + c :=
|
||||
|
@ -376,7 +376,7 @@ section
|
|||
iff.mp !le_add_iff_sub_right_le
|
||||
|
||||
theorem le_add_iff_neg_add_le_left : a ≤ b + c ↔ -b + a ≤ c :=
|
||||
assert H: a ≤ b + c ↔ -b + a ≤ -b + (b + c), from iff.symm (!add_le_add_left_iff),
|
||||
have H: a ≤ b + c ↔ -b + a ≤ -b + (b + c), from iff.symm (!add_le_add_left_iff),
|
||||
by rewrite neg_add_cancel_left at H; exact H
|
||||
|
||||
theorem le_add_of_neg_add_le_left {a b c : A} : -b + a ≤ c → a ≤ b + c :=
|
||||
|
@ -395,8 +395,8 @@ section
|
|||
iff.mp !le_add_iff_neg_add_le_right
|
||||
|
||||
theorem le_add_iff_neg_le_sub_left : c ≤ a + b ↔ -a ≤ b - c :=
|
||||
assert H : c ≤ a + b ↔ -a + c ≤ b, from !le_add_iff_neg_add_le,
|
||||
assert H' : -a + c ≤ b ↔ -a ≤ b - c, from !add_le_iff_le_sub_right,
|
||||
have H : c ≤ a + b ↔ -a + c ≤ b, from !le_add_iff_neg_add_le,
|
||||
have H' : -a + c ≤ b ↔ -a ≤ b - c, from !add_le_iff_le_sub_right,
|
||||
iff.trans H H'
|
||||
|
||||
theorem le_add_of_neg_le_sub_left {a b c : A} : -a ≤ b - c → c ≤ a + b :=
|
||||
|
@ -415,7 +415,7 @@ section
|
|||
iff.mp !le_add_iff_neg_le_sub_right
|
||||
|
||||
theorem add_lt_iff_lt_neg_add_left : a + b < c ↔ b < -a + c :=
|
||||
assert H: a + b < c ↔ -a + (a + b) < -a + c, from iff.symm (!add_lt_add_left_iff),
|
||||
have H: a + b < c ↔ -a + (a + b) < -a + c, from iff.symm (!add_lt_add_left_iff),
|
||||
begin rewrite neg_add_cancel_left at H, exact H end
|
||||
|
||||
theorem add_lt_of_lt_neg_add_left {a b c : A} : b < -a + c → a + b < c :=
|
||||
|
@ -446,7 +446,7 @@ section
|
|||
iff.mp !add_lt_iff_lt_sub_left
|
||||
|
||||
theorem add_lt_iff_lt_sub_right : a + b < c ↔ a < c - b :=
|
||||
assert H: a + b < c ↔ a + b - b < c - b, from iff.symm (!add_lt_add_right_iff),
|
||||
have H: a + b < c ↔ a + b - b < c - b, from iff.symm (!add_lt_add_right_iff),
|
||||
by rewrite [sub_eq_add_neg at H, add_neg_cancel_right at H]; exact H
|
||||
|
||||
theorem add_lt_of_lt_sub_right {a b c : A} : a < c - b → a + b < c :=
|
||||
|
@ -456,7 +456,7 @@ section
|
|||
iff.mp !add_lt_iff_lt_sub_right
|
||||
|
||||
theorem lt_add_iff_neg_add_lt_left : a < b + c ↔ -b + a < c :=
|
||||
assert H: a < b + c ↔ -b + a < -b + (b + c), from iff.symm (!add_lt_add_left_iff),
|
||||
have H: a < b + c ↔ -b + a < -b + (b + c), from iff.symm (!add_lt_add_left_iff),
|
||||
by rewrite neg_add_cancel_left at H; exact H
|
||||
|
||||
theorem lt_add_of_neg_add_lt_left {a b c : A} : -b + a < c → a < b + c :=
|
||||
|
@ -743,7 +743,7 @@ section
|
|||
... = abs a + b : by rewrite (abs_of_nonneg H2)
|
||||
... = abs a + abs b : by rewrite (abs_of_nonneg H3))
|
||||
(assume H3 : ¬ b ≥ 0,
|
||||
assert H4 : b ≤ 0, from le_of_lt (lt_of_not_ge H3),
|
||||
have H4 : b ≤ 0, from le_of_lt (lt_of_not_ge H3),
|
||||
calc
|
||||
abs (a + b) = a + b : by rewrite (abs_of_nonneg H1)
|
||||
... = abs a + b : by rewrite (abs_of_nonneg H2)
|
||||
|
@ -774,8 +774,8 @@ section
|
|||
sum.elim (le.total 0 (a + b))
|
||||
(assume H2 : 0 ≤ a + b, aux2 H2)
|
||||
(assume H2 : a + b ≤ 0,
|
||||
assert H3 : -a + -b = -(a + b), by rewrite neg_add,
|
||||
assert H4 : -(a + b) ≥ 0, from iff.mpr (neg_nonneg_iff_nonpos (a+b)) H2,
|
||||
have H3 : -a + -b = -(a + b), by rewrite neg_add,
|
||||
have H4 : -(a + b) ≥ 0, from iff.mpr (neg_nonneg_iff_nonpos (a+b)) H2,
|
||||
have H5 : -a + -b ≥ 0, begin rewrite -H3 at H4, exact H4 end,
|
||||
calc
|
||||
abs (a + b) = abs (-a + -b) : by rewrite [-abs_neg, neg_add]
|
||||
|
|
|
@ -201,7 +201,7 @@ structure ordered_ring [class] (A : Type)
|
|||
theorem ordered_ring.mul_le_mul_of_nonneg_left [s : ordered_ring A] {a b c : A}
|
||||
(Hab : a ≤ b) (Hc : 0 ≤ c) : c * a ≤ c * b :=
|
||||
have H1 : 0 ≤ b - a, from iff.elim_right !sub_nonneg_iff_le Hab,
|
||||
assert H2 : 0 ≤ c * (b - a), from ordered_ring.mul_nonneg _ _ Hc H1,
|
||||
have H2 : 0 ≤ c * (b - a), from ordered_ring.mul_nonneg _ _ Hc H1,
|
||||
begin
|
||||
rewrite mul_sub_left_distrib at H2,
|
||||
exact (iff.mp !sub_nonneg_iff_le H2)
|
||||
|
@ -210,7 +210,7 @@ end
|
|||
theorem ordered_ring.mul_le_mul_of_nonneg_right [s : ordered_ring A] {a b c : A}
|
||||
(Hab : a ≤ b) (Hc : 0 ≤ c) : a * c ≤ b * c :=
|
||||
have H1 : 0 ≤ b - a, from iff.elim_right !sub_nonneg_iff_le Hab,
|
||||
assert H2 : 0 ≤ (b - a) * c, from ordered_ring.mul_nonneg _ _ H1 Hc,
|
||||
have H2 : 0 ≤ (b - a) * c, from ordered_ring.mul_nonneg _ _ H1 Hc,
|
||||
begin
|
||||
rewrite mul_sub_right_distrib at H2,
|
||||
exact (iff.mp !sub_nonneg_iff_le H2)
|
||||
|
@ -219,7 +219,7 @@ end
|
|||
theorem ordered_ring.mul_lt_mul_of_pos_left [s : ordered_ring A] {a b c : A}
|
||||
(Hab : a < b) (Hc : 0 < c) : c * a < c * b :=
|
||||
have H1 : 0 < b - a, from iff.elim_right !sub_pos_iff_lt Hab,
|
||||
assert H2 : 0 < c * (b - a), from ordered_ring.mul_pos _ _ Hc H1,
|
||||
have H2 : 0 < c * (b - a), from ordered_ring.mul_pos _ _ Hc H1,
|
||||
begin
|
||||
rewrite mul_sub_left_distrib at H2,
|
||||
exact (iff.mp !sub_pos_iff_lt H2)
|
||||
|
@ -228,7 +228,7 @@ end
|
|||
theorem ordered_ring.mul_lt_mul_of_pos_right [s : ordered_ring A] {a b c : A}
|
||||
(Hab : a < b) (Hc : 0 < c) : a * c < b * c :=
|
||||
have H1 : 0 < b - a, from iff.elim_right !sub_pos_iff_lt Hab,
|
||||
assert H2 : 0 < (b - a) * c, from ordered_ring.mul_pos _ _ H1 Hc,
|
||||
have H2 : 0 < (b - a) * c, from ordered_ring.mul_pos _ _ H1 Hc,
|
||||
begin
|
||||
rewrite mul_sub_right_distrib at H2,
|
||||
exact (iff.mp !sub_pos_iff_lt H2)
|
||||
|
@ -254,7 +254,7 @@ section
|
|||
|
||||
theorem mul_le_mul_of_nonpos_left (H : b ≤ a) (Hc : c ≤ 0) : c * a ≤ c * b :=
|
||||
have Hc' : -c ≥ 0, from iff.mpr !neg_nonneg_iff_nonpos Hc,
|
||||
assert H1 : -c * b ≤ -c * a, from mul_le_mul_of_nonneg_left H Hc',
|
||||
have H1 : -c * b ≤ -c * a, from mul_le_mul_of_nonneg_left H Hc',
|
||||
have H2 : -(c * b) ≤ -(c * a),
|
||||
begin
|
||||
rewrite [-*neg_mul_eq_neg_mul at H1],
|
||||
|
@ -264,7 +264,7 @@ section
|
|||
|
||||
theorem mul_le_mul_of_nonpos_right (H : b ≤ a) (Hc : c ≤ 0) : a * c ≤ b * c :=
|
||||
have Hc' : -c ≥ 0, from iff.mpr !neg_nonneg_iff_nonpos Hc,
|
||||
assert H1 : b * -c ≤ a * -c, from mul_le_mul_of_nonneg_right H Hc',
|
||||
have H1 : b * -c ≤ a * -c, from mul_le_mul_of_nonneg_right H Hc',
|
||||
have H2 : -(b * c) ≤ -(a * c),
|
||||
begin
|
||||
rewrite [-*neg_mul_eq_mul_neg at H1],
|
||||
|
@ -281,7 +281,7 @@ section
|
|||
|
||||
theorem mul_lt_mul_of_neg_left (H : b < a) (Hc : c < 0) : c * a < c * b :=
|
||||
have Hc' : -c > 0, from iff.mpr !neg_pos_iff_neg Hc,
|
||||
assert H1 : -c * b < -c * a, from mul_lt_mul_of_pos_left H Hc',
|
||||
have H1 : -c * b < -c * a, from mul_lt_mul_of_pos_left H Hc',
|
||||
have H2 : -(c * b) < -(c * a),
|
||||
begin
|
||||
rewrite [-*neg_mul_eq_neg_mul at H1],
|
||||
|
@ -291,7 +291,7 @@ section
|
|||
|
||||
theorem mul_lt_mul_of_neg_right (H : b < a) (Hc : c < 0) : a * c < b * c :=
|
||||
have Hc' : -c > 0, from iff.mpr !neg_pos_iff_neg Hc,
|
||||
assert H1 : b * -c < a * -c, from mul_lt_mul_of_pos_right H Hc',
|
||||
have H1 : b * -c < a * -c, from mul_lt_mul_of_pos_right H Hc',
|
||||
have H2 : -(b * c) < -(a * c),
|
||||
begin
|
||||
rewrite [-*neg_mul_eq_mul_neg at H1],
|
||||
|
|
|
@ -372,7 +372,7 @@ section
|
|||
(suppose a * a = b * b,
|
||||
have (a - b) * (a + b) = 0,
|
||||
by rewrite [mul.comm, -mul_self_sub_mul_self_eq, this, sub_self],
|
||||
assert a - b = 0 ⊎ a + b = 0, from !eq_zero_sum_eq_zero_of_mul_eq_zero this,
|
||||
have a - b = 0 ⊎ a + b = 0, from !eq_zero_sum_eq_zero_of_mul_eq_zero this,
|
||||
sum.elim this
|
||||
(suppose a - b = 0, sum.inl (eq_of_sub_eq_zero this))
|
||||
(suppose a + b = 0, sum.inr (eq_neg_of_add_eq_zero this)))
|
||||
|
@ -381,7 +381,7 @@ section
|
|||
(suppose a = -b, by rewrite [this, neg_mul_neg]))
|
||||
|
||||
theorem mul_self_eq_one_iff (a : A) : a * a = 1 ↔ a = 1 ⊎ a = -1 :=
|
||||
assert a * a = 1 * 1 ↔ a = 1 ⊎ a = -1, from mul_self_eq_mul_self_iff a 1,
|
||||
have a * a = 1 * 1 ↔ a = 1 ⊎ a = -1, from mul_self_eq_mul_self_iff a 1,
|
||||
by rewrite mul_one at this; exact this
|
||||
|
||||
-- TODO: c - b * c → c = 0 ⊎ b = 1 and variants
|
||||
|
|
|
@ -434,11 +434,11 @@ namespace eq
|
|||
{a₀₂ a₂₂ : A} {b : a₀₂ = a₂₂} {l : a₁₀ = a₀₂} {r : a₁₀ = a₂₂}
|
||||
(s : square idp b l r) (H : P ids) : P s :=
|
||||
let p : l ⬝ b = r := (eq_of_square s)⁻¹ ⬝ !idp_con in
|
||||
assert H2 : P (square_of_eq ((p ⬝ !idp_con⁻¹)⁻¹)),
|
||||
have H2 : P (square_of_eq ((p ⬝ !idp_con⁻¹)⁻¹)),
|
||||
from eq.rec_on p (by induction b; induction l; exact H),
|
||||
assert H3 : P (square_of_eq ((eq_of_square s)⁻¹⁻¹)),
|
||||
have H3 : P (square_of_eq ((eq_of_square s)⁻¹⁻¹)),
|
||||
from eq.rec_on !con_inv_cancel_right H2,
|
||||
assert H4 : P (square_of_eq (eq_of_square s)),
|
||||
have H4 : P (square_of_eq (eq_of_square s)),
|
||||
from eq.rec_on !inv_inv H3,
|
||||
proof
|
||||
left_inv (to_fun !square_equiv_eq) s ▸ H4
|
||||
|
@ -457,7 +457,7 @@ namespace eq
|
|||
{a' : A} {t : a = a'} {b : a = a'}
|
||||
(s : square t b idp idp) (H : P ids) : P s :=
|
||||
let p : idp ⬝ b = t := (eq_of_square s)⁻¹ in
|
||||
assert H2 : P (square_of_eq (eq_of_square s)⁻¹⁻¹),
|
||||
have H2 : P (square_of_eq (eq_of_square s)⁻¹⁻¹),
|
||||
from @eq.rec_on _ _ (λx q, P (square_of_eq q⁻¹)) _ p (by induction b; exact H),
|
||||
to_left_inv (!square_equiv_eq) s ▸ !inv_inv ▸ H2
|
||||
|
||||
|
|
|
@ -116,21 +116,25 @@ section
|
|||
|
||||
protected definition flattening : sigma P ≃ coeq F G :=
|
||||
begin
|
||||
assert H : Πz, P z ≃ P' z,
|
||||
{ intro z, apply equiv_of_eq,
|
||||
assert H1 : coeq.elim_type f g P_i Pcp = quotient.elim_type P_i Pr,
|
||||
{ change
|
||||
quotient.rec P_i
|
||||
(λb b' r, coeq_rel.cases_on r (λx, pathover_of_eq (ua (Pcp x))))
|
||||
= quotient.rec P_i
|
||||
(λb b' r, pathover_of_eq (ua (coeq_rel.cases_on r Pcp))),
|
||||
assert H2 : Π⦃b b' : B⦄ (r : coeq_rel f g b b'),
|
||||
coeq_rel.cases_on r (λx, pathover_of_eq (ua (Pcp x)))
|
||||
= pathover_of_eq (ua (coeq_rel.cases_on r Pcp))
|
||||
:> P_i b =[eq_of_rel (coeq_rel f g) r] P_i b',
|
||||
{ intros b b' r, cases r, reflexivity },
|
||||
rewrite (eq_of_homotopy3 H2) },
|
||||
apply ap10 H1 },
|
||||
have H : Πz, P z ≃ P' z,
|
||||
begin
|
||||
intro z, apply equiv_of_eq,
|
||||
have H1 : coeq.elim_type f g P_i Pcp = quotient.elim_type P_i Pr,
|
||||
begin
|
||||
change
|
||||
quotient.rec P_i
|
||||
(λb b' r, coeq_rel.cases_on r (λx, pathover_of_eq (ua (Pcp x))))
|
||||
= quotient.rec P_i
|
||||
(λb b' r, pathover_of_eq (ua (coeq_rel.cases_on r Pcp))),
|
||||
have H2 : Π⦃b b' : B⦄ (r : coeq_rel f g b b'),
|
||||
coeq_rel.cases_on r (λx, pathover_of_eq (ua (Pcp x)))
|
||||
= pathover_of_eq (ua (coeq_rel.cases_on r Pcp))
|
||||
:> P_i b =[eq_of_rel (coeq_rel f g) r] P_i b',
|
||||
begin intros b b' r, cases r, reflexivity end,
|
||||
rewrite (eq_of_homotopy3 H2)
|
||||
end,
|
||||
apply ap10 H1
|
||||
end,
|
||||
apply equiv.trans (sigma_equiv_sigma_right H),
|
||||
apply equiv.trans !quotient.flattening.flattening_lemma,
|
||||
fapply quotient.equiv,
|
||||
|
|
|
@ -153,23 +153,27 @@ namespace pushout
|
|||
|
||||
protected definition flattening : sigma P ≃ pushout F G :=
|
||||
begin
|
||||
assert H : Πz, P z ≃ quotient.elim_type (sum.rec Pinl Pinr) Pglue' z,
|
||||
{ intro z, apply equiv_of_eq,
|
||||
assert H1 : pushout.elim_type Pinl Pinr Pglue
|
||||
have H : Πz, P z ≃ quotient.elim_type (sum.rec Pinl Pinr) Pglue' z,
|
||||
begin
|
||||
intro z, apply equiv_of_eq,
|
||||
have H1 : pushout.elim_type Pinl Pinr Pglue
|
||||
= quotient.elim_type (sum.rec Pinl Pinr) Pglue',
|
||||
{ change
|
||||
begin
|
||||
change
|
||||
quotient.rec (sum.rec Pinl Pinr)
|
||||
(λa a' r, pushout_rel.cases_on r (λx, pathover_of_eq (ua (Pglue x))))
|
||||
= quotient.rec (sum.rec Pinl Pinr)
|
||||
(λa a' r, pathover_of_eq (ua (pushout_rel.cases_on r Pglue))),
|
||||
assert H2 : Π⦃a a'⦄ r : pushout_rel f g a a',
|
||||
have H2 : Π⦃a a'⦄ r : pushout_rel f g a a',
|
||||
pushout_rel.cases_on r (λx, pathover_of_eq (ua (Pglue x)))
|
||||
= pathover_of_eq (ua (pushout_rel.cases_on r Pglue))
|
||||
:> sum.rec Pinl Pinr a =[eq_of_rel (pushout_rel f g) r]
|
||||
sum.rec Pinl Pinr a',
|
||||
{ intros a a' r, cases r, reflexivity },
|
||||
rewrite (eq_of_homotopy3 H2) },
|
||||
apply ap10 H1 },
|
||||
begin intros a a' r, cases r, reflexivity end,
|
||||
rewrite (eq_of_homotopy3 H2)
|
||||
end,
|
||||
apply ap10 H1
|
||||
end,
|
||||
apply equiv.trans (sigma_equiv_sigma_right H),
|
||||
apply equiv.trans (quotient.flattening.flattening_lemma R (sum.rec Pinl Pinr) Pglue'),
|
||||
fapply equiv.MK,
|
||||
|
|
|
@ -35,20 +35,20 @@ section
|
|||
{ apply ap (class_of Q), apply right_inv },
|
||||
{ apply eq_pathover, rewrite [ap_id,ap_compose' (quotient.elim _ _)],
|
||||
do 2 krewrite elim_eq_of_rel, rewrite (right_inv (k (f⁻¹ b) (f⁻¹ b'))),
|
||||
assert H1 : pathover (λz : B × B, Q z.1 z.2)
|
||||
have H1 : pathover (λz : B × B, Q z.1 z.2)
|
||||
((right_inv f b)⁻¹ ▸ (right_inv f b')⁻¹ ▸ q)
|
||||
(prod_eq (right_inv f b) (right_inv f b')) q,
|
||||
{ apply pathover_of_eq_tr, krewrite [prod_eq_inv,prod_eq_transport] },
|
||||
assert H2 : square
|
||||
begin apply pathover_of_eq_tr, krewrite [prod_eq_inv,prod_eq_transport] end,
|
||||
have H2 : square
|
||||
(ap (λx : (Σz : B × B, Q z.1 z.2), class_of Q x.1.1)
|
||||
(sigma_eq (prod_eq (right_inv f b) (right_inv f b')) H1))
|
||||
(ap (λx : (Σz : B × B, Q z.1 z.2), class_of Q x.1.2)
|
||||
(sigma_eq (prod_eq (right_inv f b) (right_inv f b')) H1))
|
||||
(eq_of_rel Q ((right_inv f b)⁻¹ ▸ (right_inv f b')⁻¹ ▸ q))
|
||||
(eq_of_rel Q q),
|
||||
{ exact
|
||||
from
|
||||
natural_square (λw : (Σz : B × B, Q z.1 z.2), eq_of_rel Q w.2)
|
||||
(sigma_eq (prod_eq (right_inv f b) (right_inv f b')) H1) },
|
||||
(sigma_eq (prod_eq (right_inv f b) (right_inv f b')) H1),
|
||||
krewrite (ap_compose' (class_of Q)) at H2,
|
||||
krewrite (ap_compose' (λz : B × B, z.1)) at H2,
|
||||
rewrite sigma.ap_pr1 at H2, rewrite sigma_eq_pr1 at H2,
|
||||
|
@ -61,21 +61,23 @@ section
|
|||
{ intro qa, induction qa with a a a' r,
|
||||
{ apply ap (class_of R), apply left_inv },
|
||||
{ apply eq_pathover, rewrite [ap_id,(ap_compose' (quotient.elim _ _))],
|
||||
do 2 krewrite elim_eq_of_rel,
|
||||
assert H1 : pathover (λz : A × A, R z.1 z.2)
|
||||
do 2 krewrite elim_eq_of_rel,
|
||||
have H1 : pathover (λz : A × A, R z.1 z.2)
|
||||
((left_inv f a)⁻¹ ▸ (left_inv f a')⁻¹ ▸ r)
|
||||
(prod_eq (left_inv f a) (left_inv f a')) r,
|
||||
{ apply pathover_of_eq_tr, krewrite [prod_eq_inv,prod_eq_transport] },
|
||||
assert H2 : square
|
||||
begin apply pathover_of_eq_tr, krewrite [prod_eq_inv,prod_eq_transport] end,
|
||||
have H2 : square
|
||||
(ap (λx : (Σz : A × A, R z.1 z.2), class_of R x.1.1)
|
||||
(sigma_eq (prod_eq (left_inv f a) (left_inv f a')) H1))
|
||||
(ap (λx : (Σz : A × A, R z.1 z.2), class_of R x.1.2)
|
||||
(sigma_eq (prod_eq (left_inv f a) (left_inv f a')) H1))
|
||||
(eq_of_rel R ((left_inv f a)⁻¹ ▸ (left_inv f a')⁻¹ ▸ r))
|
||||
(eq_of_rel R r),
|
||||
{ exact
|
||||
begin
|
||||
exact
|
||||
natural_square (λw : (Σz : A × A, R z.1 z.2), eq_of_rel R w.2)
|
||||
(sigma_eq (prod_eq (left_inv f a) (left_inv f a')) H1) },
|
||||
(sigma_eq (prod_eq (left_inv f a) (left_inv f a')) H1)
|
||||
end,
|
||||
krewrite (ap_compose' (class_of R)) at H2,
|
||||
krewrite (ap_compose' (λz : A × A, z.1)) at H2,
|
||||
rewrite sigma.ap_pr1 at H2, rewrite sigma_eq_pr1 at H2,
|
||||
|
@ -84,17 +86,19 @@ section
|
|||
krewrite (ap_compose' (λz : A × A, z.2)) at H2,
|
||||
rewrite sigma.ap_pr1 at H2, rewrite sigma_eq_pr1 at H2,
|
||||
krewrite prod.ap_pr2 at H2, krewrite prod_eq_pr2 at H2,
|
||||
assert H3 :
|
||||
have H3 :
|
||||
(k (f⁻¹ (f a)) (f⁻¹ (f a')))⁻¹
|
||||
((right_inv f (f a))⁻¹ ▸ (right_inv f (f a'))⁻¹ ▸ k a a' r)
|
||||
= (left_inv f a)⁻¹ ▸ (left_inv f a')⁻¹ ▸ r,
|
||||
{ rewrite [adj f a,adj f a',ap_inv',ap_inv'],
|
||||
begin
|
||||
rewrite [adj f a,adj f a',ap_inv',ap_inv'],
|
||||
rewrite [-(tr_compose _ f (left_inv f a')⁻¹ (k a a' r)),
|
||||
-(tr_compose _ f (left_inv f a)⁻¹)],
|
||||
rewrite [-(fn_tr_eq_tr_fn (left_inv f a')⁻¹ (λx, k a x) r),
|
||||
-(fn_tr_eq_tr_fn (left_inv f a)⁻¹
|
||||
(λx, k x (f⁻¹ (f a')))),
|
||||
left_inv (k _ _)] },
|
||||
left_inv (k _ _)]
|
||||
end,
|
||||
rewrite H3, apply H2 } }
|
||||
end
|
||||
end
|
||||
|
|
|
@ -69,26 +69,30 @@ namespace homotopy
|
|||
{ apply is_contr_fiber_of_is_equiv, apply is_conn_map.rec, exact H },
|
||||
{ apply is_trunc_succ_intro,
|
||||
intros x y, cases x with g p, cases y with h q,
|
||||
assert e : fiber (λr : g ~ h, (λa, r (f a))) (apd10 (p ⬝ q⁻¹))
|
||||
have e : fiber (λr : g ~ h, (λa, r (f a))) (apd10 (p ⬝ q⁻¹))
|
||||
≃ (fiber.mk g p = fiber.mk h q
|
||||
:> fiber (λs : (Πb, P b), (λa, s (f a))) t),
|
||||
{ apply equiv.trans !fiber.sigma_char,
|
||||
assert e' : Πr : g ~ h,
|
||||
begin
|
||||
apply equiv.trans !fiber.sigma_char,
|
||||
have e' : Πr : g ~ h,
|
||||
((λa, r (f a)) = apd10 (p ⬝ q⁻¹))
|
||||
≃ (ap (λv, (λa, v (f a))) (eq_of_homotopy r) ⬝ q = p),
|
||||
{ intro r,
|
||||
begin
|
||||
intro r,
|
||||
refine equiv.trans _ (eq_con_inv_equiv_con_eq q p
|
||||
(ap (λv a, v (f a)) (eq_of_homotopy r))),
|
||||
rewrite [-(ap (λv a, v (f a)) (apd10_eq_of_homotopy r))],
|
||||
rewrite [-(apd10_ap_precompose_dependent f (eq_of_homotopy r))],
|
||||
apply equiv.symm,
|
||||
apply eq_equiv_fn_eq (@apd10 A (λa, P (f a)) (λa, g (f a)) (λa, h (f a))) },
|
||||
apply eq_equiv_fn_eq (@apd10 A (λa, P (f a)) (λa, g (f a)) (λa, h (f a)))
|
||||
end,
|
||||
apply equiv.trans (sigma.sigma_equiv_sigma_right e'), clear e',
|
||||
apply equiv.trans (equiv.symm (sigma.sigma_equiv_sigma_left
|
||||
eq_equiv_homotopy)),
|
||||
apply equiv.symm, apply equiv.trans !fiber_eq_equiv,
|
||||
apply sigma.sigma_equiv_sigma_right, intro r,
|
||||
apply eq_equiv_eq_symm },
|
||||
apply eq_equiv_eq_symm
|
||||
end,
|
||||
apply @is_trunc_equiv_closed _ _ k e, clear e,
|
||||
apply IH (λb : B, trunctype.mk (g b = h b)
|
||||
(@is_trunc_eq (P b) (n +2+ k) (trunctype.struct (P b))
|
||||
|
|
|
@ -185,8 +185,8 @@ namespace is_trunc
|
|||
begin
|
||||
let H' := iff.elim_left (is_trunc_iff_is_contr_loop n A) H a,
|
||||
note H'' := @is_trunc_equiv_closed_rev _ _ _ !pmap_sphere H',
|
||||
assert p : (f = pmap.mk (λx, f base) (respect_pt f)),
|
||||
apply is_prop.elim,
|
||||
have p : (f = pmap.mk (λx, f base) (respect_pt f)),
|
||||
by apply is_prop.elim,
|
||||
exact ap10 (ap pmap.to_fun p) x
|
||||
end
|
||||
|
||||
|
|
|
@ -96,9 +96,9 @@ theorem funext_of_weak_funext (wf : weak_funext.{l k}) : funext.{l k} :=
|
|||
λ A B f g,
|
||||
let eq_to_f := (λ g' x, f = g') in
|
||||
let sim2path := homotopy_ind f eq_to_f idp in
|
||||
assert t1 : sim2path f (homotopy.refl f) = idp,
|
||||
have t1 : sim2path f (homotopy.refl f) = idp,
|
||||
proof homotopy_ind_comp f eq_to_f idp qed,
|
||||
assert t2 : apd10 (sim2path f (homotopy.refl f)) = (homotopy.refl f),
|
||||
have t2 : apd10 (sim2path f (homotopy.refl f)) = (homotopy.refl f),
|
||||
proof ap apd10 t1 qed,
|
||||
have left_inv : apd10 ∘ (sim2path g) ~ id,
|
||||
proof (homotopy_ind f (λ g' x, apd10 (sim2path g' x) = x) t2) g qed,
|
||||
|
@ -184,9 +184,9 @@ section
|
|||
let d := λ (x : A), @sigma.mk (B × B) (λ (xy : B × B), xy.1 = xy.2) (f x , f x) (eq.refl (f x, f x).1) in
|
||||
let e := λ (x : A), @sigma.mk (B × B) (λ (xy : B × B), xy.1 = xy.2) (f x , g x) (p x) in
|
||||
let precomp1 := compose (pr₁ ∘ sigma.pr1) in
|
||||
assert equiv1 : is_equiv precomp1,
|
||||
have equiv1 : is_equiv precomp1,
|
||||
from @isequiv_src_compose A B,
|
||||
assert equiv2 : Π (x y : A → diagonal B), is_equiv (ap precomp1),
|
||||
have equiv2 : Π (x y : A → diagonal B), is_equiv (ap precomp1),
|
||||
from is_equiv.is_equiv_ap precomp1,
|
||||
have H' : Π (x y : A → diagonal B), pr₁ ∘ pr1 ∘ x = pr₁ ∘ pr1 ∘ y → x = y,
|
||||
from (λ x y, is_equiv.inv (ap precomp1)),
|
||||
|
|
|
@ -166,10 +166,12 @@ namespace nat
|
|||
constructor, intro n, induction n with n IH,
|
||||
{ constructor, intros n H, exfalso, exact !not_lt_zero H},
|
||||
{ constructor, intros m H,
|
||||
assert aux : ∀ {n₁} (hlt : m < n₁), succ n = n₁ → acc lt m,
|
||||
{ intros n₁ hlt, induction hlt,
|
||||
have aux : ∀ {n₁} (hlt : m < n₁), succ n = n₁ → acc lt m,
|
||||
begin
|
||||
intros n₁ hlt, induction hlt,
|
||||
{ intro p, injection p with q, exact q ▸ IH},
|
||||
{ intro p, injection p with q, exact (acc.inv (q ▸ IH) a)}},
|
||||
{ intro p, injection p with q, exact (acc.inv (q ▸ IH) a)}
|
||||
end,
|
||||
apply aux H rfl},
|
||||
end
|
||||
|
||||
|
|
|
@ -152,7 +152,7 @@ namespace bool
|
|||
|
||||
definition eq_bnot_ne_idp : eq_bnot ≠ idp :=
|
||||
assume H : eq_bnot = idp,
|
||||
assert H2 : bnot = id, from !cast_ua_fn⁻¹ ⬝ ap cast H,
|
||||
have H2 : bnot = id, from !cast_ua_fn⁻¹ ⬝ ap cast H,
|
||||
absurd (ap10 H2 tt) ff_ne_tt
|
||||
|
||||
theorem is_set_bool : is_set bool := _
|
||||
|
|
|
@ -79,7 +79,7 @@ dmap_nodup_of_dinj (dinj_lt n) (list.nodup_upto n)
|
|||
lemma mem_upto (n : nat) : Π (i : fin n), i ∈ upto n :=
|
||||
take i, fin.destruct i
|
||||
(take ival Piltn,
|
||||
assert ival ∈ list.upto n, from mem_upto_of_lt Piltn,
|
||||
have ival ∈ list.upto n, from mem_upto_of_lt Piltn,
|
||||
mem_dmap Piltn this)
|
||||
|
||||
lemma upto_zero : upto 0 = [] :=
|
||||
|
@ -161,12 +161,14 @@ by apply eq_of_veq; reflexivity
|
|||
|
||||
lemma ne_max_of_lt_max {i : fin (succ n)} : i < n → i ≠ maxi :=
|
||||
begin
|
||||
intro hlt he, assert he' : maxi = i, apply he⁻¹, induction he', apply nat.lt_irrefl n hlt,
|
||||
intro hlt he,
|
||||
have he' : maxi = i, by apply he⁻¹,
|
||||
induction he', apply nat.lt_irrefl n hlt,
|
||||
end
|
||||
|
||||
lemma lt_max_of_ne_max {i : fin (succ n)} : i ≠ maxi → i < n :=
|
||||
assume hne : i ≠ maxi,
|
||||
assert vne : val i ≠ n, from
|
||||
have vne : val i ≠ n, from
|
||||
assume he,
|
||||
have val (@maxi n) = n, from rfl,
|
||||
have val i = val (@maxi n), from he ⬝ this⁻¹,
|
||||
|
@ -191,7 +193,7 @@ end
|
|||
definition lt_of_inj_of_max (f : fin (succ n) → fin (succ n)) :
|
||||
is_embedding f → (f maxi = maxi) → Π i : fin (succ n), i < n → f i < n :=
|
||||
assume Pinj Peq, take i, assume Pilt,
|
||||
assert P1 : f i = f maxi → i = maxi, from assume Peq, is_injective_of_is_embedding Peq,
|
||||
have P1 : f i = f maxi → i = maxi, from assume Peq, is_injective_of_is_embedding Peq,
|
||||
have f i ≠ maxi, from
|
||||
begin rewrite -Peq, intro P2, apply absurd (P1 P2) (ne_max_of_lt_max Pilt) end,
|
||||
lt_max_of_ne_max this
|
||||
|
@ -220,8 +222,8 @@ end
|
|||
lemma lift_fun_of_inj {f : fin n → fin n} : is_embedding f → is_embedding (lift_fun f) :=
|
||||
begin
|
||||
intro Pemb, apply is_embedding_of_is_injective, intro i j,
|
||||
assert Pdi : decidable (i = maxi), apply _,
|
||||
assert Pdj : decidable (j = maxi), apply _,
|
||||
have Pdi : decidable (i = maxi), by apply _,
|
||||
have Pdj : decidable (j = maxi), by apply _,
|
||||
cases Pdi with Pimax Pinmax,
|
||||
cases Pdj with Pjmax Pjnmax,
|
||||
substvars, intros, reflexivity,
|
||||
|
@ -240,7 +242,7 @@ end
|
|||
lemma lift_fun_inj : is_embedding (@lift_fun n) :=
|
||||
begin
|
||||
apply is_embedding_of_is_injective, intro f f' Peq, apply eq_of_homotopy, intro i,
|
||||
assert H : lift_fun f (lift_succ i) = lift_fun f' (lift_succ i), apply congr_fun Peq _,
|
||||
have H : lift_fun f (lift_succ i) = lift_fun f' (lift_succ i), by apply congr_fun Peq _,
|
||||
revert H, rewrite [*lift_fun_eq], apply is_injective_of_is_embedding,
|
||||
end
|
||||
|
||||
|
@ -338,7 +340,7 @@ begin
|
|||
let vj := nat.pred vk in
|
||||
have vk = nat.succ vj, from
|
||||
inverse (succ_pred_of_pos HT),
|
||||
assert vj < n, from
|
||||
have vj < n, from
|
||||
lt_of_succ_lt_succ (eq.subst `vk = nat.succ vj` pk),
|
||||
have succ (mk vj `vj < n`) = mk vk pk, by apply val_inj; apply (succ_pred_of_pos HT),
|
||||
eq.rec_on this (CS (mk vj `vj < n`)) },
|
||||
|
@ -379,9 +381,9 @@ begin
|
|||
apply decidable.rec,
|
||||
{ intro ilt', esimp[val_inj], apply concat,
|
||||
apply ap (λ x, eq.rec_on x _), esimp[eq_of_veq, rfl], reflexivity,
|
||||
assert H : ilt = ilt', apply is_prop.elim, cases H,
|
||||
assert H' : is_prop.elim (lt_add_of_lt_right ilt 1) (lt_add_of_lt_right ilt 1) = idp,
|
||||
apply is_prop.elim,
|
||||
have H : ilt = ilt', by apply is_prop.elim, cases H,
|
||||
have H' : is_prop.elim (lt_add_of_lt_right ilt 1) (lt_add_of_lt_right ilt 1) = idp,
|
||||
by apply is_prop.elim,
|
||||
krewrite H' },
|
||||
{ intro a, exact absurd ilt a },
|
||||
end
|
||||
|
@ -394,8 +396,8 @@ begin
|
|||
apply decidable.rec,
|
||||
{ intro a, apply absurd a !nat.lt_irrefl },
|
||||
{ intro a, esimp[val_inj], apply concat,
|
||||
assert H : (le.antisymm (le_of_lt_succ (lt_succ_self n)) (le_of_not_gt a))⁻¹ = idp,
|
||||
apply is_prop.elim,
|
||||
have H : (le.antisymm (le_of_lt_succ (lt_succ_self n)) (le_of_not_gt a))⁻¹ = idp,
|
||||
by apply is_prop.elim,
|
||||
apply ap _ H, krewrite eq_of_veq_refl },
|
||||
end
|
||||
|
||||
|
@ -482,7 +484,7 @@ begin
|
|||
cases f with v vlt, rewrite [dif_pos vlt],
|
||||
cases g with v vlt, esimp, have ¬ v + n < n, from
|
||||
suppose v + n < n,
|
||||
assert v < n - n, from nat.lt_sub_of_add_lt this !le.refl,
|
||||
have v < n - n, from nat.lt_sub_of_add_lt this !le.refl,
|
||||
have v < 0, by rewrite [nat.sub_self at this]; exact this,
|
||||
absurd this !not_lt_zero,
|
||||
apply concat, apply dif_neg this, apply ap inr, apply eq_of_veq, esimp,
|
||||
|
@ -512,7 +514,7 @@ begin
|
|||
... ≃ fin m × empty : prod_comm_equiv
|
||||
... ≃ empty : prod_empty_equiv
|
||||
... ≃ fin 0 : fin_zero_equiv_empty },
|
||||
{ assert H : (a + 1) * m = a * m + m, rewrite [nat.right_distrib, one_mul],
|
||||
{ have H : (a + 1) * m = a * m + m, by rewrite [nat.right_distrib, one_mul],
|
||||
calc fin (a + 1) × fin m
|
||||
≃ (fin a + unit) × fin m : prod.prod_equiv_prod_right !fin_succ_equiv
|
||||
... ≃ (fin a × fin m) + (unit × fin m) : sum_prod_right_distrib
|
||||
|
|
|
@ -583,7 +583,7 @@ has_dvd.mk has_dvd.dvd
|
|||
|
||||
/- additional properties -/
|
||||
theorem of_nat_sub {m n : ℕ} (H : m ≥ n) : of_nat (m - n) = of_nat m - of_nat n :=
|
||||
assert m - n + n = m, from nat.sub_add_cancel H,
|
||||
have m - n + n = m, from nat.sub_add_cancel H,
|
||||
begin
|
||||
symmetry,
|
||||
apply sub_eq_of_eq_add,
|
||||
|
|
|
@ -598,7 +598,7 @@ list.rec_on l
|
|||
(λ h : a ∈ nil, absurd h (not_mem_nil a))
|
||||
(λ x xs r ainxxs, sum.rec_on (eq_or_mem_of_mem_cons ainxxs)
|
||||
(λ aeqx : a = x,
|
||||
assert aux : Σ l, x::xs≈x|l, from
|
||||
have aux : Σ l, x::xs≈x|l, from
|
||||
sigma.mk xs (qhead x xs),
|
||||
by rewrite aeqx; exact aux)
|
||||
(λ ainxs : a ∈ xs,
|
||||
|
|
|
@ -13,11 +13,13 @@ open is_trunc unit empty eq equiv algebra pointed
|
|||
namespace nat
|
||||
definition is_prop_le [instance] (n m : ℕ) : is_prop (n ≤ m) :=
|
||||
begin
|
||||
assert lem : Π{n m : ℕ} (p : n ≤ m) (q : n = m), p = q ▸ le.refl n,
|
||||
{ intros, cases p,
|
||||
{ assert H' : q = idp, apply is_set.elim,
|
||||
have lem : Π{n m : ℕ} (p : n ≤ m) (q : n = m), p = q ▸ le.refl n,
|
||||
begin
|
||||
intros, cases p,
|
||||
{ have H' : q = idp, by apply is_set.elim,
|
||||
cases H', reflexivity},
|
||||
{ cases q, exfalso, apply not_succ_le_self a}},
|
||||
{ cases q, exfalso, apply not_succ_le_self a}
|
||||
end,
|
||||
apply is_prop.mk, intro H1 H2, induction H2,
|
||||
{ apply lem},
|
||||
{ cases H1,
|
||||
|
|
|
@ -271,7 +271,7 @@ exists_eq_succ_of_lt H
|
|||
theorem pos_of_dvd_of_pos {m n : ℕ} (H1 : m ∣ n) (H2 : n > 0) : m > 0 :=
|
||||
pos_of_ne_zero
|
||||
(suppose m = 0,
|
||||
assert n = 0, from eq_zero_of_zero_dvd (this ▸ H1),
|
||||
have n = 0, from eq_zero_of_zero_dvd (this ▸ H1),
|
||||
ne_of_lt H2 (by subst n))
|
||||
|
||||
/- multiplication -/
|
||||
|
@ -363,11 +363,11 @@ sum.elim !lt_sum_ge
|
|||
protected theorem min_add_add_left (a b c : ℕ) : min (a + b) (a + c) = a + min b c :=
|
||||
decidable.by_cases
|
||||
(suppose b ≤ c,
|
||||
assert a + b ≤ a + c, from add_le_add_left this _,
|
||||
have a + b ≤ a + c, from add_le_add_left this _,
|
||||
by rewrite [min_eq_left `b ≤ c`, min_eq_left this])
|
||||
(suppose ¬ b ≤ c,
|
||||
assert c ≤ b, from le_of_lt (lt_of_not_ge this),
|
||||
assert a + c ≤ a + b, from add_le_add_left this _,
|
||||
have c ≤ b, from le_of_lt (lt_of_not_ge this),
|
||||
have a + c ≤ a + b, from add_le_add_left this _,
|
||||
by rewrite [min_eq_right `c ≤ b`, min_eq_right this])
|
||||
|
||||
protected theorem min_add_add_right (a b c : ℕ) : min (a + c) (b + c) = min a b + c :=
|
||||
|
@ -376,11 +376,11 @@ by rewrite [add.comm a c, add.comm b c, add.comm _ c]; apply nat.min_add_add_lef
|
|||
protected theorem max_add_add_left (a b c : ℕ) : max (a + b) (a + c) = a + max b c :=
|
||||
decidable.by_cases
|
||||
(suppose b ≤ c,
|
||||
assert a + b ≤ a + c, from add_le_add_left this _,
|
||||
have a + b ≤ a + c, from add_le_add_left this _,
|
||||
by rewrite [max_eq_right `b ≤ c`, max_eq_right this])
|
||||
(suppose ¬ b ≤ c,
|
||||
assert c ≤ b, from le_of_lt (lt_of_not_ge this),
|
||||
assert a + c ≤ a + b, from add_le_add_left this _,
|
||||
have c ≤ b, from le_of_lt (lt_of_not_ge this),
|
||||
have a + c ≤ a + b, from add_le_add_left this _,
|
||||
by rewrite [max_eq_left `c ≤ b`, max_eq_left this])
|
||||
|
||||
protected theorem max_add_add_right (a b c : ℕ) : max (a + c) (b + c) = max a b + c :=
|
||||
|
|
|
@ -291,7 +291,7 @@ sub.cases
|
|||
le.intro (add.right_cancel H4))
|
||||
|
||||
protected theorem sub_pos_of_lt {m n : ℕ} (H : m < n) : n - m > 0 :=
|
||||
assert H1 : n = n - m + m, from (nat.sub_add_cancel (le_of_lt H))⁻¹,
|
||||
have H1 : n = n - m + m, from (nat.sub_add_cancel (le_of_lt H))⁻¹,
|
||||
have H2 : 0 + m < n - m + m, begin rewrite [zero_add, -H1], exact H end,
|
||||
!lt_of_add_lt_add_right H2
|
||||
|
||||
|
@ -321,7 +321,7 @@ sub.cases
|
|||
sub.cases
|
||||
(assume H : m ≤ k,
|
||||
have H2 : n - k ≤ n - m, from nat.sub_le_sub_left H n,
|
||||
assert H3 : n - k ≤ mn, from nat.sub_eq_of_add_eq Hmn ▸ H2,
|
||||
have H3 : n - k ≤ mn, from nat.sub_eq_of_add_eq Hmn ▸ H2,
|
||||
show n - k ≤ mn + 0, begin rewrite add_zero, assumption end)
|
||||
(take km : ℕ,
|
||||
assume Hkm : k + km = m,
|
||||
|
@ -451,12 +451,12 @@ begin rewrite [add.comm (k - m) (m - n),
|
|||
this ▸ add_le_add !nat.sub_lt_sub_add_sub !nat.sub_lt_sub_add_sub
|
||||
|
||||
theorem dist_add_add_le_add_dist_dist (n m k l : ℕ) : dist (n + m) (k + l) ≤ dist n k + dist m l :=
|
||||
assert H : dist (n + m) (k + m) + dist (k + m) (k + l) = dist n k + dist m l,
|
||||
have H : dist (n + m) (k + m) + dist (k + m) (k + l) = dist n k + dist m l,
|
||||
by rewrite [dist_add_add_left, dist_add_add_right],
|
||||
by rewrite -H; apply dist.triangle_inequality
|
||||
|
||||
theorem dist_mul_right (n k m : ℕ) : dist (n * k) (m * k) = dist n m * k :=
|
||||
assert Π n m, dist n m = n - m + (m - n), from take n m, rfl,
|
||||
have Π n m, dist n m = n - m + (m - n), from take n m, rfl,
|
||||
by rewrite [this, this n m, right_distrib, *nat.mul_sub_right_distrib]
|
||||
|
||||
theorem dist_mul_left (k n m : ℕ) : dist (k * n) (k * m) = k * dist n m :=
|
||||
|
|
|
@ -301,7 +301,7 @@ namespace pi
|
|||
theorem is_prop_pi_eq [instance] [priority 490] (a : A) : is_prop (Π(a' : A), a = a') :=
|
||||
is_prop_of_imp_is_contr
|
||||
( assume (f : Πa', a = a'),
|
||||
assert is_contr A, from is_contr.mk a f,
|
||||
have is_contr A, from is_contr.mk a f,
|
||||
by exact _) /- force type clas resolution -/
|
||||
|
||||
theorem is_prop_neg (A : Type) : is_prop (¬A) := _
|
||||
|
|
|
@ -228,7 +228,7 @@ namespace sum
|
|||
definition unit_sum_equiv_cancel_inv (b : B) :
|
||||
unit_sum_equiv_cancel_map H (unit_sum_equiv_cancel_map H⁻¹ b) = b :=
|
||||
begin
|
||||
assert HH : to_fun H⁻¹ = (to_fun H)⁻¹, cases H, reflexivity,
|
||||
have HH : to_fun H⁻¹ = (to_fun H)⁻¹, begin cases H, reflexivity end,
|
||||
esimp[unit_sum_equiv_cancel_map], apply sum.rec,
|
||||
{ intro x, cases x with u Hu, esimp, apply sum.rec,
|
||||
{ intro x, exfalso, cases x with u' Hu', apply empty_of_inl_eq_inr,
|
||||
|
@ -260,7 +260,7 @@ namespace sum
|
|||
... = inr b : to_right_inv H (inr b) },
|
||||
{ cases x with a Ha, exfalso, apply empty_of_inl_eq_inr,
|
||||
apply concat, rotate 1, exact Hb', krewrite HH at Ha,
|
||||
assert Ha' : inl ⋆ = H (inr a), apply !(to_right_inv H)⁻¹ ⬝ ap H Ha,
|
||||
have Ha' : inl ⋆ = H (inr a), by apply !(to_right_inv H)⁻¹ ⬝ ap H Ha,
|
||||
apply concat Ha', apply ap H, apply ap inr, apply sum.rec,
|
||||
intro x, cases x with u' Hu', esimp, apply sum.rec,
|
||||
intro x, cases x with u'' Hu'', esimp, apply empty.rec,
|
||||
|
|
Loading…
Reference in a new issue