refactor(hott,tests): make sure HoTT library and tests still work if we introduce checkpoints in have-expressions

This commit is contained in:
Leonardo de Moura 2016-02-04 12:20:50 -08:00
parent a08bc408c8
commit 30d6853ffd
7 changed files with 18 additions and 22 deletions

View file

@ -115,7 +115,7 @@ section division_ring
symm (eq_one_div_of_mul_eq_one this)
theorem division_ring.one_div_neg_eq_neg_one_div (H : a ≠ 0) : 1 / (- a) = - (1 / a) :=
have -1 ≠ 0, from
have -1 ≠ (0:A), from
(suppose -1 = 0, absurd (symm (calc
1 = -(-1) : neg_neg
... = -0 : this

View file

@ -181,15 +181,14 @@ section
theorem nondep_funext_from_ua {A : Type} {B : Type}
: Π {f g : A → B}, f ~ g → f = g :=
(λ (f g : A → B) (p : f ~ g),
let d := λ (x : A), sigma.mk (f x , f x) idp in
let e := λ (x : A), sigma.mk (f x , g x) (p x) in
let precomp1 := compose (pr₁ ∘ pr1) in
have equiv1 [visible] : is_equiv precomp1,
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,
from @isequiv_src_compose A B,
have equiv2 [visible] : Π x y, is_equiv (ap precomp1),
assert 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,
have H' : Π (x y : A → diagonal B), pr₁ ∘ pr1 ∘ x = pr₁ ∘ pr1 ∘ y → x = y,
from (λ x y, is_equiv.inv (ap precomp1)),
have eq2 : pr₁ ∘ pr1 ∘ d = pr₁ ∘ pr1 ∘ e,
from idp,

View file

@ -301,8 +301,8 @@ namespace pi
theorem is_hprop_pi_eq [instance] [priority 490] (a : A) : is_hprop (Π(a' : A), a = a') :=
is_hprop_of_imp_is_contr
( assume (f : Πa', a = a'),
assert H : is_contr A, from is_contr.mk a f,
_)
assert is_contr A, from is_contr.mk a f,
by exact _) /- force type clas resolution -/
theorem is_hprop_neg (A : Type) : is_hprop (¬A) := _
local attribute ne [reducible]

View file

@ -291,7 +291,7 @@ private lemma max_count_eq (l₁ l₂ : list A) : ∀ {a : A} {l : list A}, a
| a (b::l) h₁ h₂ :=
assert nodup l, from nodup_of_nodup_cons h₂,
assert b ∉ l, from not_mem_of_nodup_cons h₂,
or.elim h₁
or.elim (eq_or_mem_of_mem_cons h₁)
(suppose a = b,
have a ∉ l, by rewrite this; assumption,
assert a ∉ max_count l₁ l₂ l, from not_mem_max_count_of_not_mem l₁ l₂ this,
@ -337,7 +337,7 @@ private lemma min_count_eq (l₁ l₂ : list A) : ∀ {a : A} {l : list A}, a
| a (b::l) h₁ h₂ :=
assert nodup l, from nodup_of_nodup_cons h₂,
assert b ∉ l, from not_mem_of_nodup_cons h₂,
or.elim h₁
or.elim (eq_or_mem_of_mem_cons h₁)
(suppose a = b,
have a ∉ l, by rewrite this; assumption,
assert a ∉ min_count l₁ l₂ l, from not_mem_min_count_of_not_mem l₁ l₂ this,

View file

@ -19,7 +19,9 @@ u : nodup_list A,
l : list A,
a : A,
l' : list A,
IH : ∀ (x : @nodup A l'), P (@quot.mk (nodup_list A) (nodup_list_setoid A) (@subtype.tag (list A) (@nodup A) l' x)),
IH :
∀ (has_property : @nodup A l'),
P (@quot.mk (nodup_list A) (nodup_list_setoid A) (@subtype.tag (list A) (@nodup A) l' has_property)),
nodup_al' : @nodup A (@cons A a l'),
anl' : not (@list.mem A a l'),
H3 : @eq (list A) (@list.insert A (λ (a b : A), h a b) a l') (@cons A a l'),
@ -41,7 +43,9 @@ u : nodup_list A,
l : list A,
a : A,
l' : list A,
IH : ∀ (x : @nodup A l'), P (@quot.mk (nodup_list A) (nodup_list_setoid A) (@subtype.tag (list A) (@nodup A) l' x)),
IH :
∀ (has_property : @nodup A l'),
P (@quot.mk (nodup_list A) (nodup_list_setoid A) (@subtype.tag (list A) (@nodup A) l' has_property)),
nodup_al' : @nodup A (@cons A a l'),
anl' : not (@list.mem A a l'),
H3 : @eq (list A) (@list.insert A (λ (a b : A), h a b) a l') (@cons A a l'),

View file

@ -30,7 +30,7 @@ set_option trace.class_instances true
theorem T1 {A B C D : Type} {P : C → Prop} (a : A) (H1 : inh B) (H2 : ∃x, P x) : inh ((A → A) × B × (D → C) × Prop) :=
have h1 [visible] : inh A, from inh.intro a,
have h2 [visible] : inh C, from inh_exists H2,
_
by exact _
reveal T1
(*

View file

@ -23,10 +23,3 @@ theorem dcongr {A : Type} {B : A → Type} {a b : A} (f : Π x, B x) (p : a = b)
:= have H1 : ∀ p1 : a = a, transport p1 (f a) = f a, from
assume p1 : a = a, transport_eq p1 (f a),
eq.rec H1 p p
theorem transport_trans {A : Type} {a b c : A} {P : A → Type} (p1 : a = b) (p2 : b = c) (H : P a) :
transport p1 (transport p2 H) = transport (trans p1 p2) H
:= have H1 : ∀ p, transport p1 (transport p H) = transport (trans p1 p) H, from
take p, calc transport p1 (transport p H) = transport p1 H : {transport_eq p H}
... = transport (trans p1 p) H : refl (transport p1 H),
eq.rec H1 p2 p2