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

Update the code to match the latest HoTT

HoTT commit 3526c344c47d32f5d4d268658031777239ec952b
This commit is contained in:
2017-11-06 15:25:56 +01:00
parent 55136b24e7
commit 8ce4cb760a
10 changed files with 31 additions and 328 deletions

View File

@@ -142,4 +142,4 @@ Section ext.
- symmetry.
eapply equiv_functor_prod' ; apply subset_union_equiv.
Defined.
End ext.
End ext.

View File

@@ -131,25 +131,19 @@ Defined.
Section relations.
Context `{Univalence}.
(** Membership of finite sets. *)
(** Membership of finite sets. *)
Global Instance fset_member : forall A, hasMembership (FSet A) A.
Proof.
intros A a.
hrecursion.
- apply False_hp.
- apply .
- apply (fun a' => merely(a = a')).
- apply lor.
(* TODO *)
- eauto with lattice_hints typeclass_instances.
apply associativity.
- eauto with lattice_hints typeclass_instances.
apply commutativity.
- eauto with lattice_hints typeclass_instances.
apply left_identity.
- eauto with lattice_hints typeclass_instances.
apply right_identity.
- eauto with lattice_hints typeclass_instances.
intros. simpl. apply binary_idempotent.
- apply ().
- eauto 10 with lattice_hints typeclass_instances.
- eauto 10 with lattice_hints typeclass_instances.
- eauto 10 with lattice_hints typeclass_instances.
- eauto 10 with lattice_hints typeclass_instances.
- eauto 8 with lattice_hints typeclass_instances.
Defined.
(** Subset relation of finite sets. *)
@@ -157,17 +151,13 @@ Section relations.
Proof.
intros A X Y.
hrecursion X.
- apply Unit_hp.
- apply .
- apply (fun a => a Y).
- apply land.
(* TODO *)
- eauto with lattice_hints typeclass_instances.
apply associativity.
- eauto with lattice_hints typeclass_instances.
apply commutativity.
- apply left_identity.
- apply right_identity.
- eauto with lattice_hints typeclass_instances.
intros. apply binary_idempotent.
- apply ().
- eauto 10 with lattice_hints typeclass_instances.
- eauto 10 with lattice_hints typeclass_instances.
- eauto 10 with lattice_hints typeclass_instances.
- eauto 10 with lattice_hints typeclass_instances.
- eauto 8 with lattice_hints typeclass_instances.
Defined.
End relations.

View File

@@ -244,7 +244,7 @@ Section properties_membership.
Definition singleton_isIn (a b: A) : a {|b|} -> Trunc (-1) (a = b) := idmap.
Definition union_isIn (X Y : FSet A) (a : A)
: a X Y = a X a Y := idpath.
: a X Y = a X a Y := idpath.
Lemma comprehension_union (ϕ : A -> Bool) : forall X Y : FSet A,
{| (X Y) | ϕ|} = {|X | ϕ|} {|Y | ϕ|}.
@@ -261,7 +261,7 @@ Section properties_membership.
assert (forall c d, ϕ a = c -> ϕ b = d ->
a (if ϕ b then {|b|} else )
=
(if ϕ a then BuildhProp (Trunc (-1) (a = b)) else False_hp)) as X.
(if ϕ a then merely (a = b) else False_hp)) as X.
{
intros c d Hc Hd.
destruct c ; destruct d ; rewrite Hc, Hd ; try reflexivity
@@ -286,7 +286,7 @@ Section properties_membership.
Context {B : Type}.
Lemma singleproduct_isIn (a : A) (b : B) (c : A) : forall (Y : FSet B),
(a, b) (single_product c Y) = land (BuildhProp (Trunc (-1) (a = c))) (b Y).
(a, b) (single_product c Y) = merely (a = c) (b Y).
Proof.
hinduction ; try (intros ; apply path_ishprop).
- apply path_hprop ; symmetry ; apply prod_empty_r.
@@ -325,7 +325,7 @@ Section properties_membership.
Defined.
Definition product_isIn (a : A) (b : B) (X : FSet A) (Y : FSet B) :
(a,b) (product X Y) = land (a X) (b Y).
(a,b) (product X Y) = (a X) (b Y).
Proof.
hinduction X ; try (intros ; apply path_ishprop).
- apply path_hprop ; symmetry ; apply prod_empty_l.
@@ -664,8 +664,9 @@ Ltac simplify_isIn_d :=
Ltac toBool :=
repeat intro;
apply ext ; intros ;
simplify_isIn_d ; eauto with bool_lattice_hints typeclass_instances.
apply ext; intros;
simplify_isIn_d;
eauto 10 with lattice_hints typeclass_instances.
(** If `A` has decidable equality, then `FSet A` is a lattice *)
Section set_lattice.
@@ -680,7 +681,6 @@ Section set_lattice.
repeat split; try apply _;
compute[sg_op meet_is_sg_op meet_fset];
toBool.
apply associativity.
Defined.
End set_lattice.