1
0
mirror of https://github.com/nmvdw/HITs-Examples synced 2025-11-04 07:33:51 +01:00

Port the codebase to HottClasses

Initial work + use the latest version of HoTT
This commit is contained in:
2017-11-01 17:47:41 +01:00
parent c6f756a856
commit 4fafbf175e
17 changed files with 389 additions and 495 deletions

View File

@@ -8,9 +8,9 @@ Section finite_hott.
Context `{Univalence}.
(* A subobject is B-finite if its extension is B-finite as a type *)
Definition Bfin (X : Sub A) : hProp := BuildhProp (Finite {a : A & a X}).
Definition Bfin (X : Sub A) : hProp := BuildhProp (Finite {a : A | a X}).
Global Instance singleton_contr a `{IsHSet A} : Contr {b : A & b {|a|}}.
Global Instance singleton_contr a `{IsHSet A} : Contr {b : A | b {|a|}}.
Proof.
exists (a; tr idpath).
intros [b p].
@@ -20,7 +20,7 @@ Section finite_hott.
apply p^.
Defined.
Definition singleton_fin_equiv' a : Fin 1 -> {b : A & b {|a|}}.
Definition singleton_fin_equiv' a : Fin 1 -> {b : A | b {|a|}}.
Proof.
intros _. apply (a; tr idpath).
Defined.
@@ -66,7 +66,8 @@ Section finite_hott.
: DecidablePaths A.
Proof.
intros a b.
destruct (U {|a|} {|b|} (singleton a) (singleton b)) as [n pn].
specialize (U {|a|} {|b|} (singleton a) (singleton b)).
destruct U as [n pn].
strip_truncations.
unfold Sect in *.
destruct pn as [f [g fg gf _]], n as [ | [ | n]].
@@ -75,9 +76,9 @@ Section finite_hott.
apply (tr(inl(tr idpath))).
- refine (inl _).
pose (s1 := (a;tr(inl(tr idpath)))
: {c : A & Trunc (-1) (Trunc (-1) (c = a) + Trunc (-1) (c = b))}).
: {c : A | Trunc (-1) (Trunc (-1) (c = a) + Trunc (-1) (c = b))}).
pose (s2 := (b;tr(inr(tr idpath)))
: {c : A & Trunc (-1) (Trunc (-1) (c = a) + Trunc (-1) (c = b))}).
: {c : A | Trunc (-1) (Trunc (-1) (c = a) + Trunc (-1) (c = b))}).
refine (ap (fun x => x.1) (gf s1)^ @ _ @ (ap (fun x => x.1) (gf s2))).
assert (fs_eq : f s1 = f s2).
{ by apply path_ishprop. }
@@ -116,9 +117,9 @@ Section singleton_set.
Variable (A : Type).
Context `{Univalence}.
Variable (HA : forall a, {b : A & b {|a|}} <~> Fin 1).
Variable (HA : forall a, {b : A | b {|a|}} <~> Fin 1).
Definition el x : {b : A & b {|x|}} := (x;tr idpath).
Definition el x : {b : A | b {|x|}} := (x;tr idpath).
Theorem single_bfin_set : forall (x : A) (p : x = x), p = idpath.
Proof.
@@ -132,7 +133,7 @@ Section singleton_set.
rewrite <- ap_compose.
enough(forall r,
(eissect HA X)^
@ ap (fun x0 : {b : A & b {|x|}} => HA^-1 (HA x0)) r
@ ap (fun x0 : {b : A | b {|x|}} => HA^-1 (HA x0)) r
@ eissect HA X = r
) as H2.
{
@@ -166,7 +167,7 @@ End singleton_set.
Section empty.
Variable (A : Type).
Variable (X : A -> hProp)
(Xequiv : {a : A & a X} <~> Fin 0).
(Xequiv : {a : A | a X} <~> Fin 0).
Context `{Univalence}.
Lemma X_empty : X = .
@@ -185,10 +186,10 @@ Section split.
Variable (A : Type).
Variable (P : A -> hProp)
(n : nat)
(f : {a : A & P a } <~> Fin n + Unit).
(f : {a : A | P a } <~> Fin n + Unit).
Definition split : exists P' : Sub A, exists b : A,
({a : A & P' a} <~> Fin n) * (forall x, P x = (P' x merely (x = b))).
prod ({a : A | P' a} <~> Fin n) (forall x, P x = (P' x merely (x = b))).
Proof.
pose (fun x : A => sig (fun y : Fin n => x = (f^-1 (inl y)).1)) as P'.
assert (forall x, IsHProp (P' x)).
@@ -267,7 +268,7 @@ Arguments Bfin {_} _.
Arguments split {_} {_} _ _ _.
Section Bfin_no_singletons.
Definition S1toSig (x : S1) : {x : S1 & merely(x = base)}.
Definition S1toSig (x : S1) : {x : S1 | merely(x = base)}.
Proof.
exists x.
simple refine (S1_ind (fun z => merely(z = base)) (tr idpath) _ x).
@@ -382,7 +383,7 @@ Section kfin_bfin.
Lemma notIn_ext_union_singleton (b : A) (Y : Sub A) :
~ (b Y) ->
{a : A & a ({|b|} Y)} <~> {a : A & a Y} + Unit.
{a : A | a ({|b|} Y)} <~> {a : A | a Y} + Unit.
Proof.
intros HYb.
unshelve eapply BuildEquiv.
@@ -422,7 +423,7 @@ Section kfin_bfin.
strip_truncations.
revert fX. revert X.
induction n; intros X fX.
- rewrite (X_empty _ X fX), (neutralL_max (Sub A)).
- rewrite (X_empty _ X fX). rewrite left_identity.
apply HY.
- destruct (split X n fX) as
(X' & b & HX' & HX).
@@ -431,10 +432,10 @@ Section kfin_bfin.
+ cut (X Y = X' Y).
{ intros HXY. rewrite HXY. assumption. }
apply path_forall. intro a.
unfold union, sub_union, max_fun.
unfold union, sub_union, join, join_fun.
rewrite HX.
rewrite (commutativity (X' a)).
rewrite (associativity _ (X' a)).
rewrite <- (associativity _ (X' a)).
apply path_iff_hprop.
* intros Ha.
strip_truncations.
@@ -448,15 +449,14 @@ Section kfin_bfin.
strip_truncations.
exists (n'.+1).
apply tr.
transitivity ({a : A & a (fun a => merely (a = b)) (X' Y)}).
transitivity ({a : A | a (fun a => merely (a = b)) (X' Y)}).
* apply equiv_functor_sigma_id.
intro a.
rewrite <- (associative_max (Sub A)).
rewrite (associativity (fun a => merely (a = b)) X' Y).
assert (X = X' (fun a => merely (a = b))) as HX_.
** apply path_forall. intros ?.
unfold union, sub_union, max_fun.
apply HX.
** rewrite HX_, <- (commutative_max (Sub A) X').
** rewrite HX_, <- (commutativity X').
reflexivity.
* etransitivity.
{ apply (notIn_ext_union_singleton b _ HX'Yb). }
@@ -466,7 +466,7 @@ Section kfin_bfin.
Definition FSet_to_Bfin : forall (X : FSet A), Bfin (map X).
Proof.
hinduction; try (intros; apply path_ishprop).
- exists 0.
- exists 0%nat.
apply tr.
simple refine (BuildEquiv _ _ _ _).
destruct 1 as [? []].
@@ -490,4 +490,4 @@ Proof.
apply path_ishprop.
* refine (fun a => (a;HY a);fun _ => _).
reflexivity.
Defined.
Defined.

View File

@@ -300,7 +300,7 @@ Section subobjects.
simpl. apply path_ishprop.
Defined.
Fixpoint weaken_list_r (P Q : Sub A) (ls : list (sigT Q)) : list (sigT (max_L P Q)).
Fixpoint weaken_list_r (P Q : Sub A) (ls : list (sigT Q)) : list (sigT (P Q)).
Proof.
destruct ls as [|[x Hx] ls].
- exact nil.
@@ -323,7 +323,7 @@ Section subobjects.
Defined.
Fixpoint concatD {P Q : Sub A}
(ls : list (sigT P)) (ls' : list (sigT Q)) : list (sigT (max_L P Q)).
(ls : list (sigT P)) (ls' : list (sigT Q)) : list (sigT (P Q)).
Proof.
destruct ls as [|[y Hy] ls].
- apply weaken_list_r. exact ls'.
@@ -356,7 +356,7 @@ Section subobjects.
Defined.
Lemma enumeratedS_union (P Q : Sub A) :
enumeratedS P -> enumeratedS Q -> enumeratedS (max_L P Q).
enumeratedS P -> enumeratedS Q -> enumeratedS (P Q).
Proof.
intros HP HQ.
strip_truncations; apply tr.
@@ -388,7 +388,7 @@ Section subobjects.
Transparent enumeratedS.
Instance hprop_sub_fset (P : Sub A) :
IsHProp {X : FSet A & k_finite.map X = P}.
IsHProp {X : FSet A | k_finite.map X = P}.
Proof.
apply hprop_allpath. intros [X HX] [Y HY].
assert (X = Y) as HXY.
@@ -433,7 +433,7 @@ Section subobjects.
Definition enumeratedS_to_FSet :
forall (P : Sub A), enumeratedS P ->
{X : FSet A & k_finite.map X = P}.
{X : FSet A | k_finite.map X = P}.
Proof.
intros P HP.
strip_truncations.

View File

@@ -33,7 +33,7 @@ Section k_finite.
Definition Kf : hProp := Kf_sub (fun x => True).
Instance: IsHProp {X : FSet A & forall a : A, map X a}.
Instance: IsHProp {X : FSet A | forall a : A, map X a}.
Proof.
apply hprop_allpath.
intros [X PX] [Y PY].
@@ -77,19 +77,16 @@ Section structure_k_finite.
Context (A : Type).
Context `{Univalence}.
Lemma map_union : forall X Y : FSet A, map (X Y) = max_fun (map X) (map Y).
Lemma map_union : forall X Y : FSet A, map (X Y) = (map X) (map Y).
Proof.
intros.
unfold map, max_fun.
reflexivity.
intros.
reflexivity.
Defined.
Lemma k_finite_union : closedUnion (Kf_sub A).
Proof.
unfold closedUnion, Kf_sub, Kf_sub_intern.
intros.
destruct X0 as [SX XP].
destruct X1 as [SY YP].
intros X Y [SX XP] [SY YP].
exists (SX SY).
rewrite map_union.
rewrite XP, YP.
@@ -179,9 +176,9 @@ Section k_properties.
| inl a => ({|a|} : FSet A)
| inr b =>
end).
exists (bind _ (fset_fmap f X)).
exists (mjoin (fset_fmap f X)).
intro a.
apply bind_isIn.
apply mjoin_isIn.
specialize (HX (inl a)).
exists {|a|}. split; [ | apply tr; reflexivity ].
apply (fmap_isIn f (inl a) X).
@@ -259,33 +256,35 @@ Section alternative_definition.
Local Ltac help_solve :=
repeat (let x := fresh in intro x ; destruct x) ; intros
; try (simple refine (path_sigma _ _ _ _ _)) ; try (apply path_ishprop) ; simpl
; unfold union, sub_union, max_fun
; unfold union, sub_union, join, join_fun
; apply path_forall
; intro z
; eauto with lattice_hints typeclass_instances.
Definition fset_to_k : FSet A -> {P : A -> hProp & kf_sub P}.
Definition fset_to_k : FSet A -> {P : A -> hProp | kf_sub P}.
Proof.
hinduction.
assert (IsHSet {P : A -> hProp | kf_sub P}) as Hs.
{ apply trunc_sigma. }
simple refine (FSet_rec A {P : A -> hProp | kf_sub P} Hs _ _ _ _ _ _ _ _).
- exists .
auto.
simpl. auto.
- intros a.
exists {|a|}.
auto.
simpl. auto.
- intros [P1 HP1] [P2 HP2].
exists (P1 P2).
intros ? ? ? HP.
apply HP.
* apply HP1 ; assumption.
* apply HP2 ; assumption.
- help_solve.
- help_solve.
- help_solve.
- help_solve.
- help_solve.
- help_solve. (* TODO: eauto *) apply associativity.
- help_solve. apply commutativity.
- help_solve. apply left_identity.
- help_solve. apply right_identity.
- help_solve. apply binary_idempotent.
Defined.
Definition k_to_fset : {P : A -> hProp & kf_sub P} -> FSet A.
Definition k_to_fset : {P : A -> hProp | kf_sub P} -> FSet A.
Proof.
intros [P HP].
destruct (HP (Kf_sub _) (k_finite_empty _) (k_finite_singleton _) (k_finite_union _)).
@@ -299,7 +298,7 @@ Section alternative_definition.
refine ((ap (fun z => _ z) HX2^)^ @ (ap (fun z => z X2) HX1^)^).
Defined.
Lemma k_to_fset_to_k (X : {P : A -> hProp & kf_sub P}) : fset_to_k(k_to_fset X) = X.
Lemma k_to_fset_to_k (X : {P : A -> hProp | kf_sub P}) : fset_to_k(k_to_fset X) = X.
Proof.
simple refine (path_sigma _ _ _ _ _) ; try (apply path_ishprop).
apply path_forall.

View File

@@ -6,9 +6,9 @@ Section subobjects.
Definition Sub := A -> hProp.
Global Instance sub_empty : hasEmpty Sub := fun _ => False_hp.
Global Instance sub_union : hasUnion Sub := max_fun.
Global Instance sub_intersection : hasIntersection Sub := min_fun.
Global Instance sub_empty : hasEmpty Sub := fun _ => .
Global Instance sub_union : hasUnion Sub := join.
Global Instance sub_intersection : hasIntersection Sub := meet.
Global Instance sub_singleton : hasSingleton Sub A
:= fun a b => BuildhProp (Trunc (-1) (b = a)).
Global Instance sub_membership : hasMembership Sub A := fun a X => X a.