mirror of https://github.com/nmvdw/HITs-Examples
Clean up trailing whitespaces and an unused definition.
This commit is contained in:
parent
31e08af1d1
commit
33808928db
|
@ -12,7 +12,7 @@ Open Scope logic_scope.
|
||||||
Section lor_props.
|
Section lor_props.
|
||||||
Context `{Univalence}.
|
Context `{Univalence}.
|
||||||
Variable X Y Z : hProp.
|
Variable X Y Z : hProp.
|
||||||
|
|
||||||
Lemma lor_assoc : (X ∨ Y) ∨ Z = X ∨ Y ∨ Z.
|
Lemma lor_assoc : (X ∨ Y) ∨ Z = X ∨ Y ∨ Z.
|
||||||
Proof.
|
Proof.
|
||||||
apply path_iff_hprop ; cbn.
|
apply path_iff_hprop ; cbn.
|
||||||
|
@ -20,15 +20,15 @@ Section lor_props.
|
||||||
intros [xy | z] ; cbn.
|
intros [xy | z] ; cbn.
|
||||||
+ simple refine (Trunc_ind _ _ xy).
|
+ simple refine (Trunc_ind _ _ xy).
|
||||||
intros [x | y].
|
intros [x | y].
|
||||||
++ apply (tr (inl x)).
|
++ apply (tr (inl x)).
|
||||||
++ apply (tr (inr (tr (inl y)))).
|
++ apply (tr (inr (tr (inl y)))).
|
||||||
+ apply (tr (inr (tr (inr z)))).
|
+ apply (tr (inr (tr (inr z)))).
|
||||||
* simple refine (Trunc_ind _ _).
|
* simple refine (Trunc_ind _ _).
|
||||||
intros [x | yz] ; cbn.
|
intros [x | yz] ; cbn.
|
||||||
+ apply (tr (inl (tr (inl x)))).
|
+ apply (tr (inl (tr (inl x)))).
|
||||||
+ simple refine (Trunc_ind _ _ yz).
|
+ simple refine (Trunc_ind _ _ yz).
|
||||||
intros [y | z].
|
intros [y | z].
|
||||||
++ apply (tr (inl (tr (inr y)))).
|
++ apply (tr (inl (tr (inr y)))).
|
||||||
++ apply (tr (inr z)).
|
++ apply (tr (inr z)).
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
|
@ -131,7 +131,7 @@ Section hPropLattice.
|
||||||
- intros [a b] ; apply a.
|
- intros [a b] ; apply a.
|
||||||
- intros a ; apply (pair a a).
|
- intros a ; apply (pair a a).
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Instance lor_neutrall : NeutralL lor lfalse.
|
Instance lor_neutrall : NeutralL lor lfalse.
|
||||||
Proof.
|
Proof.
|
||||||
unfold NeutralL.
|
unfold NeutralL.
|
||||||
|
@ -169,7 +169,7 @@ Section hPropLattice.
|
||||||
* assumption.
|
* assumption.
|
||||||
* apply (tr (inl X)).
|
* apply (tr (inl X)).
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Global Instance lattice_hprop : Lattice hProp :=
|
Global Instance lattice_hprop : Lattice hProp :=
|
||||||
{ commutative_min := _ ;
|
{ commutative_min := _ ;
|
||||||
commutative_max := _ ;
|
commutative_max := _ ;
|
||||||
|
@ -182,5 +182,5 @@ Section hPropLattice.
|
||||||
absorption_min_max := _ ;
|
absorption_min_max := _ ;
|
||||||
absorption_max_min := _
|
absorption_max_min := _
|
||||||
}.
|
}.
|
||||||
|
|
||||||
End hPropLattice.
|
End hPropLattice.
|
||||||
|
|
|
@ -6,7 +6,7 @@ Section ext.
|
||||||
Context {A : Type}.
|
Context {A : Type}.
|
||||||
Context `{Univalence}.
|
Context `{Univalence}.
|
||||||
|
|
||||||
Lemma subset_union (X Y : FSet A) :
|
Lemma subset_union (X Y : FSet A) :
|
||||||
X ⊆ Y -> X ∪ Y = Y.
|
X ⊆ Y -> X ∪ Y = Y.
|
||||||
Proof.
|
Proof.
|
||||||
hinduction X ; try (intros; apply path_forall; intro; apply set_path2).
|
hinduction X ; try (intros; apply path_forall; intro; apply set_path2).
|
||||||
|
@ -17,7 +17,7 @@ Section ext.
|
||||||
contradiction.
|
contradiction.
|
||||||
+ intro a0.
|
+ intro a0.
|
||||||
simple refine (Trunc_ind _ _).
|
simple refine (Trunc_ind _ _).
|
||||||
intro p ; simpl.
|
intro p ; simpl.
|
||||||
rewrite p; apply idem.
|
rewrite p; apply idem.
|
||||||
+ intros X1 X2 IH1 IH2.
|
+ intros X1 X2 IH1 IH2.
|
||||||
simple refine (Trunc_ind _ _).
|
simple refine (Trunc_ind _ _).
|
||||||
|
@ -112,8 +112,8 @@ Section ext.
|
||||||
unshelve esplit.
|
unshelve esplit.
|
||||||
{ intros [H1 H2]. etransitivity. apply H1^.
|
{ intros [H1 H2]. etransitivity. apply H1^.
|
||||||
rewrite comm. apply H2. }
|
rewrite comm. apply H2. }
|
||||||
intro; apply path_prod; apply set_path2.
|
intro; apply path_prod; apply set_path2.
|
||||||
all: intro; apply set_path2.
|
all: intro; apply set_path2.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Lemma eq_subset (X Y : FSet A) :
|
Lemma eq_subset (X Y : FSet A) :
|
||||||
|
@ -149,4 +149,4 @@ Section ext.
|
||||||
split ; apply subset_isIn.
|
split ; apply subset_isIn.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
End ext.
|
End ext.
|
||||||
|
|
|
@ -14,7 +14,7 @@ Section Iso.
|
||||||
- apply E.
|
- apply E.
|
||||||
- intros a x.
|
- intros a x.
|
||||||
apply ({|a|} ∪ x).
|
apply ({|a|} ∪ x).
|
||||||
- intros. cbn.
|
- intros. cbn.
|
||||||
etransitivity. apply assoc.
|
etransitivity. apply assoc.
|
||||||
apply (ap (∪ x)).
|
apply (ap (∪ x)).
|
||||||
apply idem.
|
apply idem.
|
||||||
|
@ -22,7 +22,7 @@ Section Iso.
|
||||||
etransitivity. apply assoc.
|
etransitivity. apply assoc.
|
||||||
etransitivity. refine (ap (∪ x) _ ).
|
etransitivity. refine (ap (∪ x) _ ).
|
||||||
apply FSet.comm.
|
apply FSet.comm.
|
||||||
symmetry.
|
symmetry.
|
||||||
apply assoc.
|
apply assoc.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
|
@ -39,10 +39,10 @@ Section Iso.
|
||||||
- apply singleton_idem.
|
- apply singleton_idem.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Lemma append_union: forall (x y: FSetC A),
|
Lemma append_union: forall (x y: FSetC A),
|
||||||
FSetC_to_FSet (x ∪ y) = (FSetC_to_FSet x) ∪ (FSetC_to_FSet y).
|
FSetC_to_FSet (x ∪ y) = (FSetC_to_FSet x) ∪ (FSetC_to_FSet y).
|
||||||
Proof.
|
Proof.
|
||||||
intros x.
|
intros x.
|
||||||
hrecursion x; try (intros; apply path_forall; intro; apply set_path2).
|
hrecursion x; try (intros; apply path_forall; intro; apply set_path2).
|
||||||
- intros. symmetry. apply nl.
|
- intros. symmetry. apply nl.
|
||||||
- intros a x HR y. unfold union, fsetc_union in *. rewrite HR. apply assoc.
|
- intros a x HR y. unfold union, fsetc_union in *. rewrite HR. apply assoc.
|
||||||
|
|
|
@ -34,7 +34,7 @@ Section operations.
|
||||||
- apply comm.
|
- apply comm.
|
||||||
- apply nl.
|
- apply nl.
|
||||||
- apply nr.
|
- apply nr.
|
||||||
- intros; simpl.
|
- intros; simpl.
|
||||||
destruct (P x).
|
destruct (P x).
|
||||||
+ apply idem.
|
+ apply idem.
|
||||||
+ apply nl.
|
+ apply nl.
|
||||||
|
@ -61,7 +61,7 @@ Section operations.
|
||||||
- intros ; apply nr.
|
- intros ; apply nr.
|
||||||
- intros ; apply idem.
|
- intros ; apply idem.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Definition product {A B : Type} : FSet A -> FSet B -> FSet (A * B).
|
Definition product {A B : Type} : FSet A -> FSet B -> FSet (A * B).
|
||||||
Proof.
|
Proof.
|
||||||
intros X Y.
|
intros X Y.
|
||||||
|
@ -76,7 +76,7 @@ Section operations.
|
||||||
- intros ; apply nr.
|
- intros ; apply nr.
|
||||||
- intros ; apply union_idem.
|
- intros ; apply union_idem.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Global Instance fset_subset : forall A, hasSubset (FSet A).
|
Global Instance fset_subset : forall A, hasSubset (FSet A).
|
||||||
Proof.
|
Proof.
|
||||||
intros A X Y.
|
intros A X Y.
|
||||||
|
@ -103,21 +103,6 @@ Section operations.
|
||||||
split ; apply p.
|
split ; apply p.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Definition map (A B : Type) (f : A -> B) : FSet A -> FSet B.
|
|
||||||
Proof.
|
|
||||||
hrecursion.
|
|
||||||
- apply ∅.
|
|
||||||
- intro a.
|
|
||||||
apply {|f a|}.
|
|
||||||
- apply (∪).
|
|
||||||
- apply assoc.
|
|
||||||
- apply comm.
|
|
||||||
- apply nl.
|
|
||||||
- apply nr.
|
|
||||||
- intros.
|
|
||||||
apply idem.
|
|
||||||
Defined.
|
|
||||||
|
|
||||||
Local Ltac remove_transport
|
Local Ltac remove_transport
|
||||||
:= intros ; apply path_forall ; intro Z ; rewrite transport_arrow
|
:= intros ; apply path_forall ; intro Z ; rewrite transport_arrow
|
||||||
; rewrite transport_const ; simpl.
|
; rewrite transport_const ; simpl.
|
||||||
|
@ -155,6 +140,6 @@ Section operations.
|
||||||
- remove_transport.
|
- remove_transport.
|
||||||
rewrite <- (idem (Z (x; tr 1%path))).
|
rewrite <- (idem (Z (x; tr 1%path))).
|
||||||
pointwise.
|
pointwise.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
End operations.
|
End operations.
|
||||||
|
|
|
@ -10,7 +10,7 @@ Section characterize_isIn.
|
||||||
|
|
||||||
(** isIn properties *)
|
(** isIn properties *)
|
||||||
Definition empty_isIn (a: A) : a ∈ ∅ -> Empty := idmap.
|
Definition empty_isIn (a: A) : a ∈ ∅ -> Empty := idmap.
|
||||||
|
|
||||||
Definition singleton_isIn (a b: A) : a ∈ {|b|} -> Trunc (-1) (a = b) := idmap.
|
Definition singleton_isIn (a b: A) : a ∈ {|b|} -> Trunc (-1) (a = b) := idmap.
|
||||||
|
|
||||||
Definition union_isIn (X Y : FSet A) (a : A)
|
Definition union_isIn (X Y : FSet A) (a : A)
|
||||||
|
@ -20,7 +20,7 @@ Section characterize_isIn.
|
||||||
{|X ∪ Y & ϕ|} = {|X & ϕ|} ∪ {|Y & ϕ|}.
|
{|X ∪ Y & ϕ|} = {|X & ϕ|} ∪ {|Y & ϕ|}.
|
||||||
Proof.
|
Proof.
|
||||||
reflexivity.
|
reflexivity.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Lemma comprehension_isIn (ϕ : A -> Bool) (a : A) : forall X : FSet A,
|
Lemma comprehension_isIn (ϕ : A -> Bool) (a : A) : forall X : FSet A,
|
||||||
a ∈ {|X & ϕ|} = if ϕ a then a ∈ X else False_hp.
|
a ∈ {|X & ϕ|} = if ϕ a then a ∈ X else False_hp.
|
||||||
|
@ -37,7 +37,7 @@ Section characterize_isIn.
|
||||||
destruct c ; destruct d ; rewrite Hc, Hd ; try reflexivity
|
destruct c ; destruct d ; rewrite Hc, Hd ; try reflexivity
|
||||||
; apply path_iff_hprop ; try contradiction ; intros ; strip_truncations
|
; apply path_iff_hprop ; try contradiction ; intros ; strip_truncations
|
||||||
; apply (false_ne_true).
|
; apply (false_ne_true).
|
||||||
* apply (Hd^ @ ap ϕ X^ @ Hc).
|
* apply (Hd^ @ ap ϕ X^ @ Hc).
|
||||||
* apply (Hc^ @ ap ϕ X @ Hd).
|
* apply (Hc^ @ ap ϕ X @ Hd).
|
||||||
}
|
}
|
||||||
apply (X (ϕ a) (ϕ b) idpath idpath).
|
apply (X (ϕ a) (ϕ b) idpath idpath).
|
||||||
|
@ -57,7 +57,7 @@ End characterize_isIn.
|
||||||
Section product_isIn.
|
Section product_isIn.
|
||||||
Context {A B : Type}.
|
Context {A B : Type}.
|
||||||
Context `{Univalence}.
|
Context `{Univalence}.
|
||||||
|
|
||||||
Lemma isIn_singleproduct (a : A) (b : B) (c : A) : forall (Y : FSet B),
|
Lemma isIn_singleproduct (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) = land (BuildhProp (Trunc (-1) (a = c))) (b ∈ Y).
|
||||||
Proof.
|
Proof.
|
||||||
|
@ -65,7 +65,7 @@ Section product_isIn.
|
||||||
- apply path_hprop ; symmetry ; apply prod_empty_r.
|
- apply path_hprop ; symmetry ; apply prod_empty_r.
|
||||||
- intros d.
|
- intros d.
|
||||||
apply path_iff_hprop.
|
apply path_iff_hprop.
|
||||||
* intros.
|
* intros.
|
||||||
strip_truncations.
|
strip_truncations.
|
||||||
split ; apply tr ; try (apply (ap fst X)) ; try (apply (ap snd X)).
|
split ; apply tr ; try (apply (ap fst X)) ; try (apply (ap snd X)).
|
||||||
* intros [Z1 Z2].
|
* intros [Z1 Z2].
|
||||||
|
@ -93,7 +93,7 @@ Section product_isIn.
|
||||||
** right.
|
** right.
|
||||||
split ; try (apply (tr H1)) ; try (apply Hb2).
|
split ; try (apply (tr H1)) ; try (apply Hb2).
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Definition isIn_product (a : A) (b : B) (X : FSet A) (Y : FSet B) :
|
Definition isIn_product (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) = land (a ∈ X) (b ∈ Y).
|
||||||
Proof.
|
Proof.
|
||||||
|
@ -147,7 +147,7 @@ Section properties.
|
||||||
Global Instance joinsemilattice_fset : JoinSemiLattice (FSet A).
|
Global Instance joinsemilattice_fset : JoinSemiLattice (FSet A).
|
||||||
Proof.
|
Proof.
|
||||||
split ; toHProp.
|
split ; toHProp.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
(** comprehension properties *)
|
(** comprehension properties *)
|
||||||
Lemma comprehension_false : forall (X : FSet A), {|X & fun _ => false|} = ∅.
|
Lemma comprehension_false : forall (X : FSet A), {|X & fun _ => false|} = ∅.
|
||||||
|
@ -176,7 +176,7 @@ Section properties.
|
||||||
Proof.
|
Proof.
|
||||||
toHProp.
|
toHProp.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Lemma merely_choice : forall X : FSet A, hor (X = ∅) (hexists (fun a => a ∈ X)).
|
Lemma merely_choice : forall X : FSet A, hor (X = ∅) (hexists (fun a => a ∈ X)).
|
||||||
Proof.
|
Proof.
|
||||||
hinduction; try (intros; apply equiv_hprop_allpath ; apply _).
|
hinduction; try (intros; apply equiv_hprop_allpath ; apply _).
|
||||||
|
@ -295,5 +295,5 @@ Section properties.
|
||||||
repeat f_ap.
|
repeat f_ap.
|
||||||
apply path_ishprop.
|
apply path_ishprop.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
End properties.
|
End properties.
|
||||||
|
|
|
@ -7,7 +7,7 @@ Section properties.
|
||||||
Context {A : Type}.
|
Context {A : Type}.
|
||||||
|
|
||||||
Definition append_nl : forall (x: FSetC A), ∅ ∪ x = x
|
Definition append_nl : forall (x: FSetC A), ∅ ∪ x = x
|
||||||
:= fun _ => idpath.
|
:= fun _ => idpath.
|
||||||
|
|
||||||
Lemma append_nr : forall (x: FSetC A), x ∪ ∅ = x.
|
Lemma append_nr : forall (x: FSetC A), x ∪ ∅ = x.
|
||||||
Proof.
|
Proof.
|
||||||
|
@ -15,20 +15,20 @@ Section properties.
|
||||||
- reflexivity.
|
- reflexivity.
|
||||||
- intros. apply (ap (fun y => a;;y) X).
|
- intros. apply (ap (fun y => a;;y) X).
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Lemma append_assoc {H: Funext}:
|
Lemma append_assoc {H: Funext}:
|
||||||
forall (x y z: FSetC A), x ∪ (y ∪ z) = (x ∪ y) ∪ z.
|
forall (x y z: FSetC A), x ∪ (y ∪ z) = (x ∪ y) ∪ z.
|
||||||
Proof.
|
Proof.
|
||||||
hinduction
|
hinduction
|
||||||
; try (intros ; apply path_forall ; intro
|
; try (intros ; apply path_forall ; intro
|
||||||
; apply path_forall ; intro ; apply set_path2).
|
; apply path_forall ; intro ; apply set_path2).
|
||||||
- reflexivity.
|
- reflexivity.
|
||||||
- intros a x HR y z.
|
- intros a x HR y z.
|
||||||
specialize (HR y z).
|
specialize (HR y z).
|
||||||
apply (ap (fun y => a;;y) HR).
|
apply (ap (fun y => a;;y) HR).
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Lemma append_singleton: forall (a: A) (x: FSetC A),
|
Lemma append_singleton: forall (a: A) (x: FSetC A),
|
||||||
a ;; x = x ∪ (a ;; ∅).
|
a ;; x = x ∪ (a ;; ∅).
|
||||||
Proof.
|
Proof.
|
||||||
intro a. hinduction; try (intros; apply set_path2).
|
intro a. hinduction; try (intros; apply set_path2).
|
||||||
|
@ -38,22 +38,22 @@ Section properties.
|
||||||
+ apply (ap (fun y => b ;; y) HR).
|
+ apply (ap (fun y => b ;; y) HR).
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Lemma append_comm {H: Funext}:
|
Lemma append_comm {H: Funext}:
|
||||||
forall (x1 x2: FSetC A), x1 ∪ x2 = x2 ∪ x1.
|
forall (x1 x2: FSetC A), x1 ∪ x2 = x2 ∪ x1.
|
||||||
Proof.
|
Proof.
|
||||||
hinduction ; try (intros ; apply path_forall ; intro ; apply set_path2).
|
hinduction ; try (intros ; apply path_forall ; intro ; apply set_path2).
|
||||||
- intros. symmetry. apply append_nr.
|
- intros. symmetry. apply append_nr.
|
||||||
- intros a x1 HR x2.
|
- intros a x1 HR x2.
|
||||||
etransitivity.
|
etransitivity.
|
||||||
apply (ap (fun y => a;;y) (HR x2)).
|
apply (ap (fun y => a;;y) (HR x2)).
|
||||||
transitivity ((x2 ∪ x1) ∪ (a;;∅)).
|
transitivity ((x2 ∪ x1) ∪ (a;;∅)).
|
||||||
+ apply append_singleton.
|
+ apply append_singleton.
|
||||||
+ etransitivity.
|
+ etransitivity.
|
||||||
* symmetry. apply append_assoc.
|
* symmetry. apply append_assoc.
|
||||||
* simple refine (ap (x2 ∪) (append_singleton _ _)^).
|
* simple refine (ap (x2 ∪) (append_singleton _ _)^).
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Lemma singleton_idem: forall (a: A),
|
Lemma singleton_idem: forall (a: A),
|
||||||
{|a|} ∪ {|a|} = {|a|}.
|
{|a|} ∪ {|a|} = {|a|}.
|
||||||
Proof.
|
Proof.
|
||||||
unfold singleton.
|
unfold singleton.
|
||||||
|
|
|
@ -60,7 +60,7 @@ Section operations_isIn.
|
||||||
contradiction.
|
contradiction.
|
||||||
- reflexivity.
|
- reflexivity.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
(* Union and membership *)
|
(* Union and membership *)
|
||||||
Lemma union_isIn_b (X Y : FSet A) (a : A) :
|
Lemma union_isIn_b (X Y : FSet A) (a : A) :
|
||||||
a ∈_d (X ∪ Y) = orb (a ∈_d X) (a ∈_d Y).
|
a ∈_d (X ∪ Y) = orb (a ∈_d X) (a ∈_d Y).
|
||||||
|
@ -111,24 +111,24 @@ Section SetLattice.
|
||||||
intros x y.
|
intros x y.
|
||||||
apply (x ∪ y).
|
apply (x ∪ y).
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Instance fset_min : minimum (FSet A).
|
Instance fset_min : minimum (FSet A).
|
||||||
Proof.
|
Proof.
|
||||||
intros x y.
|
intros x y.
|
||||||
apply (x ∩ y).
|
apply (x ∩ y).
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Instance fset_bot : bottom (FSet A).
|
Instance fset_bot : bottom (FSet A).
|
||||||
Proof.
|
Proof.
|
||||||
unfold bottom.
|
unfold bottom.
|
||||||
apply ∅.
|
apply ∅.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Instance lattice_fset : Lattice (FSet A).
|
Instance lattice_fset : Lattice (FSet A).
|
||||||
Proof.
|
Proof.
|
||||||
split; toBool.
|
split; toBool.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
End SetLattice.
|
End SetLattice.
|
||||||
|
|
||||||
(* With extensionality we can prove decidable equality *)
|
(* With extensionality we can prove decidable equality *)
|
||||||
|
@ -142,4 +142,4 @@ Section dec_eq.
|
||||||
apply decidable_prod.
|
apply decidable_prod.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
End dec_eq.
|
End dec_eq.
|
||||||
|
|
|
@ -87,12 +87,12 @@ Section properties.
|
||||||
|
|
||||||
Definition well_defined_filter : forall (A : Type) (ϕ : A -> Bool) (X Y : T A),
|
Definition well_defined_filter : forall (A : Type) (ϕ : A -> Bool) (X Y : T A),
|
||||||
set_eq A X Y -> set_eq A (filter ϕ X) (filter ϕ Y).
|
set_eq A X Y -> set_eq A (filter ϕ X) (filter ϕ Y).
|
||||||
Proof.
|
Proof.
|
||||||
intros A ϕ X Y HXY.
|
intros A ϕ X Y HXY.
|
||||||
simplify.
|
simplify.
|
||||||
by rewrite HXY.
|
by rewrite HXY.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Lemma union_comm : forall A (X Y : T A),
|
Lemma union_comm : forall A (X Y : T A),
|
||||||
set_eq A (X ∪ Y) (Y ∪ X).
|
set_eq A (X ∪ Y) (Y ∪ X).
|
||||||
Proof.
|
Proof.
|
||||||
|
@ -151,7 +151,7 @@ Proof. intros a b c Hab Hbc. apply (Hab @ Hbc). Defined.
|
||||||
|
|
||||||
Instance View_recursion A : HitRecursion (View A) :=
|
Instance View_recursion A : HitRecursion (View A) :=
|
||||||
{
|
{
|
||||||
indTy := _; recTy := forall (P : Type) (HP: IsHSet P) (u : T A -> P), (forall x y : T A, set_eq T (@f) A x y -> u x = u y) -> View A -> P;
|
indTy := _; recTy := forall (P : Type) (HP: IsHSet P) (u : T A -> P), (forall x y : T A, set_eq T (@f) A x y -> u x = u y) -> View A -> P;
|
||||||
H_inductor := quotient_ind (R A); H_recursor := @quotient_rec _ (R A) _
|
H_inductor := quotient_ind (R A); H_recursor := @quotient_rec _ (R A) _
|
||||||
}.
|
}.
|
||||||
|
|
||||||
|
@ -169,7 +169,7 @@ assert (resp1 : forall x y y', set_eq (@f) y y' -> u x y = u x y').
|
||||||
assert (resp2 : forall x x' y, set_eq (@f) x x' -> u x y = u x' y).
|
assert (resp2 : forall x x' y, set_eq (@f) x x' -> u x y = u x' y).
|
||||||
{ intros x x' y Hxx'.
|
{ intros x x' y Hxx'.
|
||||||
apply Hresp. apply Hxx'.
|
apply Hresp. apply Hxx'.
|
||||||
reflexivity. }
|
reflexivity. }
|
||||||
hrecursion.
|
hrecursion.
|
||||||
- intros a.
|
- intros a.
|
||||||
hrecursion.
|
hrecursion.
|
||||||
|
@ -193,7 +193,7 @@ simple refine (View_rec2 _ _ _ _).
|
||||||
apply related_classes_eq.
|
apply related_classes_eq.
|
||||||
unfold R in *.
|
unfold R in *.
|
||||||
eapply well_defined_union; eauto.
|
eapply well_defined_union; eauto.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Ltac reduce :=
|
Ltac reduce :=
|
||||||
intros ;
|
intros ;
|
||||||
|
|
|
@ -8,7 +8,7 @@ Section Operations.
|
||||||
Global Instance list_empty A : hasEmpty (list A) := nil.
|
Global Instance list_empty A : hasEmpty (list A) := nil.
|
||||||
|
|
||||||
Global Instance list_single A: hasSingleton (list A) A := fun a => cons a nil.
|
Global Instance list_single A: hasSingleton (list A) A := fun a => cons a nil.
|
||||||
|
|
||||||
Global Instance list_union A : hasUnion (list A).
|
Global Instance list_union A : hasUnion (list A).
|
||||||
Proof.
|
Proof.
|
||||||
intros l1 l2.
|
intros l1 l2.
|
||||||
|
@ -58,7 +58,7 @@ Section ListToSet.
|
||||||
* strip_truncations ; apply (tr (inl z1)).
|
* strip_truncations ; apply (tr (inl z1)).
|
||||||
* apply (tr (inr z2)).
|
* apply (tr (inr z2)).
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Definition empty_empty : list_to_set A ∅ = ∅ := idpath.
|
Definition empty_empty : list_to_set A ∅ = ∅ := idpath.
|
||||||
|
|
||||||
Lemma filter_comprehension (ϕ : A -> Bool) (l : list A) :
|
Lemma filter_comprehension (ϕ : A -> Bool) (l : list A) :
|
||||||
|
@ -72,7 +72,7 @@ Section ListToSet.
|
||||||
* rewrite nl.
|
* rewrite nl.
|
||||||
apply IHl.
|
apply IHl.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Definition singleton_single (a : A) : list_to_set A (singleton a) = {|a|} :=
|
Definition singleton_single (a : A) : list_to_set A (singleton a) = {|a|} :=
|
||||||
nr {|a|}.
|
nr {|a|}.
|
||||||
|
|
||||||
|
|
|
@ -4,7 +4,7 @@ Require Import notation.
|
||||||
|
|
||||||
Section binary_operation.
|
Section binary_operation.
|
||||||
Variable A : Type.
|
Variable A : Type.
|
||||||
|
|
||||||
Class maximum :=
|
Class maximum :=
|
||||||
max_L : operation A.
|
max_L : operation A.
|
||||||
|
|
||||||
|
@ -75,7 +75,7 @@ Section BoolLattice.
|
||||||
|
|
||||||
Ltac solve_bool :=
|
Ltac solve_bool :=
|
||||||
let x := fresh in
|
let x := fresh in
|
||||||
repeat (intro x ; destruct x)
|
repeat (intro x ; destruct x)
|
||||||
; compute
|
; compute
|
||||||
; auto
|
; auto
|
||||||
; try contradiction.
|
; try contradiction.
|
||||||
|
@ -83,12 +83,12 @@ Section BoolLattice.
|
||||||
Instance maximum_bool : maximum Bool := orb.
|
Instance maximum_bool : maximum Bool := orb.
|
||||||
Instance minimum_bool : minimum Bool := andb.
|
Instance minimum_bool : minimum Bool := andb.
|
||||||
Instance bottom_bool : bottom Bool := false.
|
Instance bottom_bool : bottom Bool := false.
|
||||||
|
|
||||||
Global Instance lattice_bool : Lattice Bool.
|
Global Instance lattice_bool : Lattice Bool.
|
||||||
Proof.
|
Proof.
|
||||||
split ; solve_bool.
|
split ; solve_bool.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Definition and_true : forall b, andb b true = b.
|
Definition and_true : forall b, andb b true = b.
|
||||||
Proof.
|
Proof.
|
||||||
solve_bool.
|
solve_bool.
|
||||||
|
@ -116,7 +116,7 @@ Section BoolLattice.
|
||||||
Proof.
|
Proof.
|
||||||
solve_bool.
|
solve_bool.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
End BoolLattice.
|
End BoolLattice.
|
||||||
|
|
||||||
Section fun_lattice.
|
Section fun_lattice.
|
||||||
|
@ -141,7 +141,7 @@ Section fun_lattice.
|
||||||
Proof.
|
Proof.
|
||||||
split ; solve_fun.
|
split ; solve_fun.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
End fun_lattice.
|
End fun_lattice.
|
||||||
|
|
||||||
Section sub_lattice.
|
Section sub_lattice.
|
||||||
|
@ -152,7 +152,7 @@ Section sub_lattice.
|
||||||
Context {Hbot : P empty_L}.
|
Context {Hbot : P empty_L}.
|
||||||
|
|
||||||
Definition AP : Type := sig P.
|
Definition AP : Type := sig P.
|
||||||
|
|
||||||
Instance botAP : bottom AP := (empty_L ; Hbot).
|
Instance botAP : bottom AP := (empty_L ; Hbot).
|
||||||
|
|
||||||
Instance maxAP : maximum AP :=
|
Instance maxAP : maximum AP :=
|
||||||
|
@ -174,17 +174,17 @@ Section sub_lattice.
|
||||||
|
|
||||||
Ltac solve_sub :=
|
Ltac solve_sub :=
|
||||||
let x := fresh in
|
let x := fresh in
|
||||||
repeat (intro x ; destruct x)
|
repeat (intro x ; destruct x)
|
||||||
; simple refine (path_sigma _ _ _ _ _)
|
; simple refine (path_sigma _ _ _ _ _)
|
||||||
; simpl
|
; simpl
|
||||||
; try (apply hprop_sub)
|
; try (apply hprop_sub)
|
||||||
; eauto 3 with lattice_hints typeclass_instances.
|
; eauto 3 with lattice_hints typeclass_instances.
|
||||||
|
|
||||||
Global Instance lattice_sub : Lattice AP.
|
Global Instance lattice_sub : Lattice AP.
|
||||||
Proof.
|
Proof.
|
||||||
split ; solve_sub.
|
split ; solve_sub.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
End sub_lattice.
|
End sub_lattice.
|
||||||
|
|
||||||
Create HintDb bool_lattice_hints.
|
Create HintDb bool_lattice_hints.
|
||||||
|
|
|
@ -7,7 +7,7 @@ Module Export T.
|
||||||
|
|
||||||
Private Inductive T (B : Type) : Type :=
|
Private Inductive T (B : Type) : Type :=
|
||||||
| b : T B
|
| b : T B
|
||||||
| c : T B.
|
| c : T B.
|
||||||
|
|
||||||
Axiom p : A -> b A = c A.
|
Axiom p : A -> b A = c A.
|
||||||
Axiom setT : IsHSet (T A).
|
Axiom setT : IsHSet (T A).
|
||||||
|
@ -23,7 +23,7 @@ Module Export T.
|
||||||
Variable (bP : P (b A)).
|
Variable (bP : P (b A)).
|
||||||
Variable (cP : P (c A)).
|
Variable (cP : P (c A)).
|
||||||
Variable (pP : forall a : A, (p a) # bP = cP).
|
Variable (pP : forall a : A, (p a) # bP = cP).
|
||||||
|
|
||||||
(* Induction principle *)
|
(* Induction principle *)
|
||||||
Fixpoint T_ind
|
Fixpoint T_ind
|
||||||
(x : T A)
|
(x : T A)
|
||||||
|
@ -31,7 +31,7 @@ Module Export T.
|
||||||
: P x
|
: P x
|
||||||
:= (match x return _ -> _ -> P x with
|
:= (match x return _ -> _ -> P x with
|
||||||
| b => fun _ _ => bP
|
| b => fun _ _ => bP
|
||||||
| c => fun _ _ => cP
|
| c => fun _ _ => cP
|
||||||
end) pP H.
|
end) pP H.
|
||||||
|
|
||||||
Axiom T_ind_beta_p : forall (a : A),
|
Axiom T_ind_beta_p : forall (a : A),
|
||||||
|
@ -68,7 +68,7 @@ Module Export T.
|
||||||
End T_recursion.
|
End T_recursion.
|
||||||
|
|
||||||
Instance T_recursion A : HitRecursion (T A)
|
Instance T_recursion A : HitRecursion (T A)
|
||||||
:= {indTy := _; recTy := _;
|
:= {indTy := _; recTy := _;
|
||||||
H_inductor := T_ind A; H_recursor := T_rec A }.
|
H_inductor := T_ind A; H_recursor := T_rec A }.
|
||||||
|
|
||||||
End T.
|
End T.
|
||||||
|
@ -119,44 +119,44 @@ Section merely_dec_lem.
|
||||||
|
|
||||||
Local Ltac f_prop := apply path_forall ; intro ; apply path_ishprop.
|
Local Ltac f_prop := apply path_forall ; intro ; apply path_ishprop.
|
||||||
|
|
||||||
Lemma transport_code_b_x (a : A) :
|
Lemma transport_code_b_x (a : A) :
|
||||||
transport code_b (p a) = fun _ => a.
|
transport code_b (p a) = fun _ => a.
|
||||||
Proof.
|
Proof.
|
||||||
f_prop.
|
f_prop.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Lemma transport_code_c_x (a : A) :
|
Lemma transport_code_c_x (a : A) :
|
||||||
transport code_c (p a) = fun _ => tt.
|
transport code_c (p a) = fun _ => tt.
|
||||||
Proof.
|
Proof.
|
||||||
f_prop.
|
f_prop.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Lemma transport_code_c_x_V (a : A) :
|
Lemma transport_code_c_x_V (a : A) :
|
||||||
transport code_c (p a)^ = fun _ => a.
|
transport code_c (p a)^ = fun _ => a.
|
||||||
Proof.
|
Proof.
|
||||||
f_prop.
|
f_prop.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Lemma transport_code_x_b (a : A) :
|
Lemma transport_code_x_b (a : A) :
|
||||||
transport (fun x => code x (b A)) (p a) = fun _ => a.
|
transport (fun x => code x (b A)) (p a) = fun _ => a.
|
||||||
Proof.
|
Proof.
|
||||||
f_prop.
|
f_prop.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Lemma transport_code_x_c (a : A) :
|
Lemma transport_code_x_c (a : A) :
|
||||||
transport (fun x => code x (c A)) (p a) = fun _ => tt.
|
transport (fun x => code x (c A)) (p a) = fun _ => tt.
|
||||||
Proof.
|
Proof.
|
||||||
f_prop.
|
f_prop.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Lemma transport_code_x_c_V (a : A) :
|
Lemma transport_code_x_c_V (a : A) :
|
||||||
transport (fun x => code x (c A)) (p a)^ = fun _ => a.
|
transport (fun x => code x (c A)) (p a)^ = fun _ => a.
|
||||||
Proof.
|
Proof.
|
||||||
f_prop.
|
f_prop.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Lemma ap_diag {B : Type} {x y : B} (p : x = y) :
|
Lemma ap_diag {B : Type} {x y : B} (p : x = y) :
|
||||||
ap (fun x : B => (x, x)) p = path_prod' p p.
|
ap (fun x : B => (x, x)) p = path_prod' p p.
|
||||||
Proof.
|
Proof.
|
||||||
by path_induction.
|
by path_induction.
|
||||||
Defined.
|
Defined.
|
||||||
|
@ -217,7 +217,7 @@ Section merely_dec_lem.
|
||||||
refine (transport_arrow _ _ _ @ _).
|
refine (transport_arrow _ _ _ @ _).
|
||||||
refine (transport_paths_FlFr _ _ @ _).
|
refine (transport_paths_FlFr _ _ @ _).
|
||||||
rewrite transport_code_c_x_V.
|
rewrite transport_code_c_x_V.
|
||||||
hott_simpl.
|
hott_simpl.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Lemma transport_paths_FlFr_trunc :
|
Lemma transport_paths_FlFr_trunc :
|
||||||
|
@ -229,7 +229,7 @@ Section merely_dec_lem.
|
||||||
refine (ap tr _).
|
refine (ap tr _).
|
||||||
exact ((concat_1p r)^ @ (concat_p1 (1 @ r))^).
|
exact ((concat_1p r)^ @ (concat_p1 (1 @ r))^).
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Definition decode : forall (x y : T A), code x y -> x = y.
|
Definition decode : forall (x y : T A), code x y -> x = y.
|
||||||
Proof.
|
Proof.
|
||||||
simple refine (T_ind _ _ _ _ _ _); simpl.
|
simple refine (T_ind _ _ _ _ _ _); simpl.
|
||||||
|
@ -248,7 +248,7 @@ Section merely_dec_lem.
|
||||||
f_ap.
|
f_ap.
|
||||||
refine (ap (fun x => (p x)) _).
|
refine (ap (fun x => (p x)) _).
|
||||||
apply path_ishprop.
|
apply path_ishprop.
|
||||||
+ intro.
|
+ intro.
|
||||||
rewrite transport_code_x_c_V.
|
rewrite transport_code_x_c_V.
|
||||||
hott_simpl.
|
hott_simpl.
|
||||||
+ intro b.
|
+ intro b.
|
||||||
|
@ -264,7 +264,7 @@ Section merely_dec_lem.
|
||||||
intros p. induction p.
|
intros p. induction p.
|
||||||
simpl. revert u. simple refine (T_ind _ _ _ _ _ _).
|
simpl. revert u. simple refine (T_ind _ _ _ _ _ _).
|
||||||
+ simpl. reflexivity.
|
+ simpl. reflexivity.
|
||||||
+ simpl. reflexivity.
|
+ simpl. reflexivity.
|
||||||
+ intro a.
|
+ intro a.
|
||||||
apply set_path2.
|
apply set_path2.
|
||||||
Defined.
|
Defined.
|
||||||
|
@ -278,12 +278,12 @@ Section merely_dec_lem.
|
||||||
+ simpl. intro a. apply path_ishprop.
|
+ simpl. intro a. apply path_ishprop.
|
||||||
+ intro a. apply path_forall; intros ?. apply set_path2.
|
+ intro a. apply path_forall; intros ?. apply set_path2.
|
||||||
- simple refine (T_ind _ _ _ _ _ _).
|
- simple refine (T_ind _ _ _ _ _ _).
|
||||||
+ simpl. intro a. apply path_ishprop.
|
+ simpl. intro a. apply path_ishprop.
|
||||||
+ simpl. apply path_ishprop.
|
+ simpl. apply path_ishprop.
|
||||||
+ intro a. apply path_forall; intros ?. apply set_path2.
|
+ intro a. apply path_forall; intros ?. apply set_path2.
|
||||||
- intro a. repeat (apply path_forall; intros ?). apply set_path2.
|
- intro a. repeat (apply path_forall; intros ?). apply set_path2.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
|
|
||||||
Instance T_hprop (a : A) : IsHProp (b A = c A).
|
Instance T_hprop (a : A) : IsHProp (b A = c A).
|
||||||
Proof.
|
Proof.
|
||||||
|
@ -307,7 +307,7 @@ Section merely_dec_lem.
|
||||||
rewrite ?decode_encode in H1.
|
rewrite ?decode_encode in H1.
|
||||||
apply H1.
|
apply H1.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Lemma equiv_pathspace_T : (b A = c A) = A.
|
Lemma equiv_pathspace_T : (b A = c A) = A.
|
||||||
Proof.
|
Proof.
|
||||||
apply path_iff_ishprop.
|
apply path_iff_ishprop.
|
||||||
|
|
|
@ -3,7 +3,7 @@ Require Import HoTT HitTactics.
|
||||||
Require Export notation.
|
Require Export notation.
|
||||||
|
|
||||||
Module Export FSetC.
|
Module Export FSetC.
|
||||||
|
|
||||||
Section FSetC.
|
Section FSetC.
|
||||||
Private Inductive FSetC (A : Type) : Type :=
|
Private Inductive FSetC (A : Type) : Type :=
|
||||||
| Nil : FSetC A
|
| Nil : FSetC A
|
||||||
|
@ -14,9 +14,9 @@ Module Export FSetC.
|
||||||
Variable A : Type.
|
Variable A : Type.
|
||||||
Arguments Cns {_} _ _.
|
Arguments Cns {_} _ _.
|
||||||
Infix ";;" := Cns (at level 8, right associativity).
|
Infix ";;" := Cns (at level 8, right associativity).
|
||||||
|
|
||||||
Axiom dupl : forall (a : A) (x : FSetC A),
|
Axiom dupl : forall (a : A) (x : FSetC A),
|
||||||
a ;; a ;; x = a ;; x.
|
a ;; a ;; x = a ;; x.
|
||||||
|
|
||||||
Axiom comm : forall (a b : A) (x : FSetC A),
|
Axiom comm : forall (a b : A) (x : FSetC A),
|
||||||
a ;; b ;; x = b ;; a ;; x.
|
a ;; b ;; x = b ;; a ;; x.
|
||||||
|
@ -41,9 +41,9 @@ Module Export FSetC.
|
||||||
Variable (duplP : forall (a: A) (x: FSetC A) (px : P x),
|
Variable (duplP : forall (a: A) (x: FSetC A) (px : P x),
|
||||||
dupl a x # cnsP a (a;;x) (cnsP a x px) = cnsP a x px).
|
dupl a x # cnsP a (a;;x) (cnsP a x px) = cnsP a x px).
|
||||||
Variable (commP : forall (a b: A) (x: FSetC A) (px: P x),
|
Variable (commP : forall (a b: A) (x: FSetC A) (px: P x),
|
||||||
comm a b x # cnsP a (b;;x) (cnsP b x px) =
|
comm a b x # cnsP a (b;;x) (cnsP b x px) =
|
||||||
cnsP b (a;;x) (cnsP a x px)).
|
cnsP b (a;;x) (cnsP a x px)).
|
||||||
|
|
||||||
(* Induction principle *)
|
(* Induction principle *)
|
||||||
Fixpoint FSetC_ind
|
Fixpoint FSetC_ind
|
||||||
(x : FSetC A)
|
(x : FSetC A)
|
||||||
|
@ -76,18 +76,18 @@ Module Export FSetC.
|
||||||
(* Recursion principle *)
|
(* Recursion principle *)
|
||||||
Definition FSetC_rec : FSetC A -> P.
|
Definition FSetC_rec : FSetC A -> P.
|
||||||
Proof.
|
Proof.
|
||||||
simple refine (FSetC_ind _ _ _ _ _ _ _ );
|
simple refine (FSetC_ind _ _ _ _ _ _ _ );
|
||||||
try (intros; simple refine ((transport_const _ _) @ _ )); cbn.
|
try (intros; simple refine ((transport_const _ _) @ _ )); cbn.
|
||||||
- apply nil.
|
- apply nil.
|
||||||
- apply (fun a => fun _ => cns a).
|
- apply (fun a => fun _ => cns a).
|
||||||
- apply duplP.
|
- apply duplP.
|
||||||
- apply commP.
|
- apply commP.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Definition FSetC_rec_beta_dupl : forall (a: A) (x : FSetC A),
|
Definition FSetC_rec_beta_dupl : forall (a: A) (x : FSetC A),
|
||||||
ap FSetC_rec (dupl a x) = duplP a (FSetC_rec x).
|
ap FSetC_rec (dupl a x) = duplP a (FSetC_rec x).
|
||||||
Proof.
|
Proof.
|
||||||
intros.
|
intros.
|
||||||
eapply (cancelL (transport_const (dupl a x) _)).
|
eapply (cancelL (transport_const (dupl a x) _)).
|
||||||
simple refine ((apD_const _ _)^ @ _).
|
simple refine ((apD_const _ _)^ @ _).
|
||||||
apply FSetC_ind_beta_dupl.
|
apply FSetC_ind_beta_dupl.
|
||||||
|
@ -96,7 +96,7 @@ Module Export FSetC.
|
||||||
Definition FSetC_rec_beta_comm : forall (a b: A) (x : FSetC A),
|
Definition FSetC_rec_beta_comm : forall (a b: A) (x : FSetC A),
|
||||||
ap FSetC_rec (comm a b x) = commP a b (FSetC_rec x).
|
ap FSetC_rec (comm a b x) = commP a b (FSetC_rec x).
|
||||||
Proof.
|
Proof.
|
||||||
intros.
|
intros.
|
||||||
eapply (cancelL (transport_const (comm a b x) _)).
|
eapply (cancelL (transport_const (comm a b x) _)).
|
||||||
simple refine ((apD_const _ _)^ @ _).
|
simple refine ((apD_const _ _)^ @ _).
|
||||||
apply FSetC_ind_beta_comm.
|
apply FSetC_ind_beta_comm.
|
||||||
|
@ -107,7 +107,7 @@ Module Export FSetC.
|
||||||
|
|
||||||
Instance FSetC_recursion A : HitRecursion (FSetC A) :=
|
Instance FSetC_recursion A : HitRecursion (FSetC A) :=
|
||||||
{
|
{
|
||||||
indTy := _; recTy := _;
|
indTy := _; recTy := _;
|
||||||
H_inductor := FSetC_ind A; H_recursor := FSetC_rec A
|
H_inductor := FSetC_ind A; H_recursor := FSetC_rec A
|
||||||
}.
|
}.
|
||||||
|
|
||||||
|
|
|
@ -14,7 +14,7 @@ Module Export FSet.
|
||||||
Global Instance fset_union : forall A, hasUnion (FSet A) := U.
|
Global Instance fset_union : forall A, hasUnion (FSet A) := U.
|
||||||
|
|
||||||
Variable A : Type.
|
Variable A : Type.
|
||||||
|
|
||||||
Axiom assoc : forall (x y z : FSet A),
|
Axiom assoc : forall (x y z : FSet A),
|
||||||
x ∪ (y ∪ z) = (x ∪ y) ∪ z.
|
x ∪ (y ∪ z) = (x ∪ y) ∪ z.
|
||||||
|
|
||||||
|
@ -38,7 +38,7 @@ Module Export FSet.
|
||||||
Arguments comm {_} _ _.
|
Arguments comm {_} _ _.
|
||||||
Arguments nl {_} _.
|
Arguments nl {_} _.
|
||||||
Arguments nr {_} _.
|
Arguments nr {_} _.
|
||||||
Arguments idem {_} _.
|
Arguments idem {_} _.
|
||||||
|
|
||||||
Section FSet_induction.
|
Section FSet_induction.
|
||||||
Variable A: Type.
|
Variable A: Type.
|
||||||
|
@ -47,22 +47,22 @@ Module Export FSet.
|
||||||
Variable (eP : P ∅).
|
Variable (eP : P ∅).
|
||||||
Variable (lP : forall a: A, P {|a|}).
|
Variable (lP : forall a: A, P {|a|}).
|
||||||
Variable (uP : forall (x y: FSet A), P x -> P y -> P (x ∪ y)).
|
Variable (uP : forall (x y: FSet A), P x -> P y -> P (x ∪ y)).
|
||||||
Variable (assocP : forall (x y z : FSet A)
|
Variable (assocP : forall (x y z : FSet A)
|
||||||
(px: P x) (py: P y) (pz: P z),
|
(px: P x) (py: P y) (pz: P z),
|
||||||
assoc x y z #
|
assoc x y z #
|
||||||
(uP x (y ∪ z) px (uP y z py pz))
|
(uP x (y ∪ z) px (uP y z py pz))
|
||||||
=
|
=
|
||||||
(uP (x ∪ y) z (uP x y px py) pz)).
|
(uP (x ∪ y) z (uP x y px py) pz)).
|
||||||
Variable (commP : forall (x y: FSet A) (px: P x) (py: P y),
|
Variable (commP : forall (x y: FSet A) (px: P x) (py: P y),
|
||||||
comm x y #
|
comm x y #
|
||||||
uP x y px py = uP y x py px).
|
uP x y px py = uP y x py px).
|
||||||
Variable (nlP : forall (x : FSet A) (px: P x),
|
Variable (nlP : forall (x : FSet A) (px: P x),
|
||||||
nl x # uP ∅ x eP px = px).
|
nl x # uP ∅ x eP px = px).
|
||||||
Variable (nrP : forall (x : FSet A) (px: P x),
|
Variable (nrP : forall (x : FSet A) (px: P x),
|
||||||
nr x # uP x ∅ px eP = px).
|
nr x # uP x ∅ px eP = px).
|
||||||
Variable (idemP : forall (x : A),
|
Variable (idemP : forall (x : A),
|
||||||
idem x # uP {|x|} {|x|} (lP x) (lP x) = lP x).
|
idem x # uP {|x|} {|x|} (lP x) (lP x) = lP x).
|
||||||
|
|
||||||
(* Induction principle *)
|
(* Induction principle *)
|
||||||
Fixpoint FSet_ind
|
Fixpoint FSet_ind
|
||||||
(x : FSet A)
|
(x : FSet A)
|
||||||
|
@ -119,7 +119,7 @@ Module Export FSet.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Definition FSet_rec_beta_assoc : forall (x y z : FSet A),
|
Definition FSet_rec_beta_assoc : forall (x y z : FSet A),
|
||||||
ap FSet_rec (assoc x y z)
|
ap FSet_rec (assoc x y z)
|
||||||
=
|
=
|
||||||
assocP (FSet_rec x) (FSet_rec y) (FSet_rec z).
|
assocP (FSet_rec x) (FSet_rec y) (FSet_rec z).
|
||||||
Proof.
|
Proof.
|
||||||
|
@ -131,7 +131,7 @@ Module Export FSet.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Definition FSet_rec_beta_comm : forall (x y : FSet A),
|
Definition FSet_rec_beta_comm : forall (x y : FSet A),
|
||||||
ap FSet_rec (comm x y)
|
ap FSet_rec (comm x y)
|
||||||
=
|
=
|
||||||
commP (FSet_rec x) (FSet_rec y).
|
commP (FSet_rec x) (FSet_rec y).
|
||||||
Proof.
|
Proof.
|
||||||
|
@ -143,7 +143,7 @@ Module Export FSet.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Definition FSet_rec_beta_nl : forall (x : FSet A),
|
Definition FSet_rec_beta_nl : forall (x : FSet A),
|
||||||
ap FSet_rec (nl x)
|
ap FSet_rec (nl x)
|
||||||
=
|
=
|
||||||
nlP (FSet_rec x).
|
nlP (FSet_rec x).
|
||||||
Proof.
|
Proof.
|
||||||
|
@ -155,7 +155,7 @@ Module Export FSet.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Definition FSet_rec_beta_nr : forall (x : FSet A),
|
Definition FSet_rec_beta_nr : forall (x : FSet A),
|
||||||
ap FSet_rec (nr x)
|
ap FSet_rec (nr x)
|
||||||
=
|
=
|
||||||
nrP (FSet_rec x).
|
nrP (FSet_rec x).
|
||||||
Proof.
|
Proof.
|
||||||
|
@ -167,7 +167,7 @@ Module Export FSet.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Definition FSet_rec_beta_idem : forall (a : A),
|
Definition FSet_rec_beta_idem : forall (a : A),
|
||||||
ap FSet_rec (idem a)
|
ap FSet_rec (idem a)
|
||||||
=
|
=
|
||||||
idemP a.
|
idemP a.
|
||||||
Proof.
|
Proof.
|
||||||
|
@ -177,12 +177,12 @@ Module Export FSet.
|
||||||
simple refine ((apD_const _ _)^ @ _).
|
simple refine ((apD_const _ _)^ @ _).
|
||||||
apply FSet_ind_beta_idem.
|
apply FSet_ind_beta_idem.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
End FSet_recursion.
|
End FSet_recursion.
|
||||||
|
|
||||||
Instance FSet_recursion A : HitRecursion (FSet A) :=
|
Instance FSet_recursion A : HitRecursion (FSet A) :=
|
||||||
{
|
{
|
||||||
indTy := _; recTy := _;
|
indTy := _; recTy := _;
|
||||||
H_inductor := FSet_ind A; H_recursor := FSet_rec A
|
H_inductor := FSet_ind A; H_recursor := FSet_rec A
|
||||||
}.
|
}.
|
||||||
|
|
||||||
|
@ -200,5 +200,5 @@ Proof.
|
||||||
rewrite P.
|
rewrite P.
|
||||||
rewrite (comm y x).
|
rewrite (comm y x).
|
||||||
rewrite <- (assoc x y y).
|
rewrite <- (assoc x y y).
|
||||||
f_ap.
|
f_ap.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
|
@ -47,7 +47,7 @@ Proof.
|
||||||
destruct (if P a as b return ((b = true) + (b = false))
|
destruct (if P a as b return ((b = true) + (b = false))
|
||||||
then inl 1%path
|
then inl 1%path
|
||||||
else inr 1%path) as [Pa' | Pa'].
|
else inr 1%path) as [Pa' | Pa'].
|
||||||
- rewrite Pa' in Pa. contradiction (true_ne_false Pa).
|
- rewrite Pa' in Pa. contradiction (true_ne_false Pa).
|
||||||
- reflexivity.
|
- reflexivity.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
|
@ -104,7 +104,7 @@ Defined.
|
||||||
Lemma enumerated_surj (A B : Type) (f : A -> B) :
|
Lemma enumerated_surj (A B : Type) (f : A -> B) :
|
||||||
IsSurjection f -> enumerated A -> enumerated B.
|
IsSurjection f -> enumerated A -> enumerated B.
|
||||||
Proof.
|
Proof.
|
||||||
intros Hsurj HeA. strip_truncations; apply tr.
|
intros Hsurj HeA. strip_truncations; apply tr.
|
||||||
destruct HeA as [eA HeA].
|
destruct HeA as [eA HeA].
|
||||||
exists (map f eA).
|
exists (map f eA).
|
||||||
intros x. specialize (Hsurj x).
|
intros x. specialize (Hsurj x).
|
||||||
|
@ -157,7 +157,7 @@ destruct ys as [|y ys].
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Fixpoint listProd {A B} (xs : list A) (ys : list B) : list (A * B).
|
Fixpoint listProd {A B} (xs : list A) (ys : list B) : list (A * B).
|
||||||
Proof.
|
Proof.
|
||||||
destruct xs as [|x xs].
|
destruct xs as [|x xs].
|
||||||
- exact nil.
|
- exact nil.
|
||||||
- refine (app _ _).
|
- refine (app _ _).
|
||||||
|
@ -165,7 +165,7 @@ destruct xs as [|x xs].
|
||||||
+ exact (listProd _ _ xs ys).
|
+ exact (listProd _ _ xs ys).
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Lemma listExt_prod_sing {A B} (x : A) (y : B) (ys : list B) :
|
Lemma listExt_prod_sing {A B} (x : A) (y : B) (ys : list B) :
|
||||||
listExt ys y -> listExt (listProd_sing x ys) (x, y).
|
listExt ys y -> listExt (listProd_sing x ys) (x, y).
|
||||||
Proof.
|
Proof.
|
||||||
induction ys; simpl.
|
induction ys; simpl.
|
||||||
|
@ -193,11 +193,11 @@ induction xs as [| x' xs]; intros x y.
|
||||||
rewrite <- Hyy' in IHxs.
|
rewrite <- Hyy' in IHxs.
|
||||||
apply listExt_app_l. apply IHxs. assumption.
|
apply listExt_app_l. apply IHxs. assumption.
|
||||||
simpl. apply tr. left. apply tr. reflexivity.
|
simpl. apply tr. left. apply tr. reflexivity.
|
||||||
* right.
|
* right.
|
||||||
apply listExt_app_l.
|
apply listExt_app_l.
|
||||||
apply IHxs. assumption.
|
apply IHxs. assumption.
|
||||||
simpl. apply tr. right. assumption.
|
simpl. apply tr. right. assumption.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
(** Properties of enumerated sets: closed under products *)
|
(** Properties of enumerated sets: closed under products *)
|
||||||
Lemma enumerated_prod (A B : Type) `{Funext} :
|
Lemma enumerated_prod (A B : Type) `{Funext} :
|
||||||
|
@ -221,7 +221,7 @@ Section enumerated_fset.
|
||||||
| nil => ∅
|
| nil => ∅
|
||||||
| cons x xs => {|x|} ∪ (list_to_fset xs)
|
| cons x xs => {|x|} ∪ (list_to_fset xs)
|
||||||
end.
|
end.
|
||||||
|
|
||||||
Lemma list_to_fset_ext (ls : list A) (a : A):
|
Lemma list_to_fset_ext (ls : list A) (a : A):
|
||||||
listExt ls a -> a ∈ (list_to_fset ls).
|
listExt ls a -> a ∈ (list_to_fset ls).
|
||||||
Proof.
|
Proof.
|
||||||
|
@ -250,8 +250,8 @@ End enumerated_fset.
|
||||||
Section fset_dec_enumerated.
|
Section fset_dec_enumerated.
|
||||||
Variable A : Type.
|
Variable A : Type.
|
||||||
Context `{Univalence}.
|
Context `{Univalence}.
|
||||||
|
|
||||||
Definition Kf_fsetc :
|
Definition Kf_fsetc :
|
||||||
Kf A -> exists (X : FSetC A), forall (a : A), k_finite.map (FSetC_to_FSet X) a.
|
Kf A -> exists (X : FSetC A), forall (a : A), k_finite.map (FSetC_to_FSet X) a.
|
||||||
Proof.
|
Proof.
|
||||||
intros [X HX].
|
intros [X HX].
|
||||||
|
@ -260,7 +260,7 @@ Section fset_dec_enumerated.
|
||||||
by rewrite <- HX.
|
by rewrite <- HX.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Definition merely_enumeration_FSetC :
|
Definition merely_enumeration_FSetC :
|
||||||
forall (X : FSetC A),
|
forall (X : FSetC A),
|
||||||
hexists (fun (ls : list A) => forall a, a ∈ (FSetC_to_FSet X) = listExt ls a).
|
hexists (fun (ls : list A) => forall a, a ∈ (FSetC_to_FSet X) = listExt ls a).
|
||||||
Proof.
|
Proof.
|
||||||
|
@ -274,13 +274,13 @@ Section fset_dec_enumerated.
|
||||||
- intros. apply path_ishprop.
|
- intros. apply path_ishprop.
|
||||||
- intros. apply path_ishprop.
|
- intros. apply path_ishprop.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Definition Kf_enumerated : Kf A -> enumerated A.
|
Definition Kf_enumerated : Kf A -> enumerated A.
|
||||||
Proof.
|
Proof.
|
||||||
intros HKf. apply Kf_fsetc in HKf.
|
intros HKf. apply Kf_fsetc in HKf.
|
||||||
destruct HKf as [X HX].
|
destruct HKf as [X HX].
|
||||||
pose (ls' := (merely_enumeration_FSetC X)).
|
pose (ls' := (merely_enumeration_FSetC X)).
|
||||||
simple refine (@Trunc_rec _ _ _ _ _ ls'). clear ls'.
|
simple refine (@Trunc_rec _ _ _ _ _ ls'). clear ls'.
|
||||||
intros [ls Hls].
|
intros [ls Hls].
|
||||||
apply tr. exists ls.
|
apply tr. exists ls.
|
||||||
intros a. rewrite <- Hls. apply (HX a).
|
intros a. rewrite <- Hls. apply (HX a).
|
||||||
|
@ -293,7 +293,7 @@ Section subobjects.
|
||||||
|
|
||||||
Definition enumeratedS (P : Sub A) : hProp :=
|
Definition enumeratedS (P : Sub A) : hProp :=
|
||||||
enumerated (sigT P).
|
enumerated (sigT P).
|
||||||
|
|
||||||
Lemma enumeratedS_empty : closedEmpty enumeratedS.
|
Lemma enumeratedS_empty : closedEmpty enumeratedS.
|
||||||
Proof.
|
Proof.
|
||||||
unfold enumeratedS.
|
unfold enumeratedS.
|
||||||
|
@ -319,7 +319,7 @@ Section subobjects.
|
||||||
- apply (cons (x; tr (inr Hx))).
|
- apply (cons (x; tr (inr Hx))).
|
||||||
apply (weaken_list_r _ _ ls).
|
apply (weaken_list_r _ _ ls).
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Lemma listExt_weaken (P Q : Sub A) (ls : list (sigT Q)) (x : A) (Hx : Q x) :
|
Lemma listExt_weaken (P Q : Sub A) (ls : list (sigT Q)) (x : A) (Hx : Q x) :
|
||||||
listExt ls (x; Hx) -> listExt (weaken_list_r P Q ls) (x; tr (inr Hx)).
|
listExt ls (x; Hx) -> listExt (weaken_list_r P Q ls) (x; tr (inr Hx)).
|
||||||
Proof.
|
Proof.
|
||||||
|
@ -333,7 +333,7 @@ Section subobjects.
|
||||||
exists (Hxy..1). apply path_ishprop.
|
exists (Hxy..1). apply path_ishprop.
|
||||||
+ right. apply IHls. assumption.
|
+ right. apply IHls. assumption.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Fixpoint concatD {P Q : Sub A}
|
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 (max_L P Q)).
|
||||||
Proof.
|
Proof.
|
||||||
|
@ -382,9 +382,9 @@ Section subobjects.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
Opaque enumeratedS.
|
Opaque enumeratedS.
|
||||||
Definition FSet_to_enumeratedS :
|
Definition FSet_to_enumeratedS :
|
||||||
forall (X : FSet A), enumeratedS (k_finite.map X).
|
forall (X : FSet A), enumeratedS (k_finite.map X).
|
||||||
Proof.
|
Proof.
|
||||||
hinduction.
|
hinduction.
|
||||||
- apply enumeratedS_empty.
|
- apply enumeratedS_empty.
|
||||||
- intros a. apply enumeratedS_singleton.
|
- intros a. apply enumeratedS_singleton.
|
||||||
|
|
|
@ -100,7 +100,7 @@ Section structure_k_finite.
|
||||||
exists {|a|}.
|
exists {|a|}.
|
||||||
cbn.
|
cbn.
|
||||||
apply path_forall.
|
apply path_forall.
|
||||||
intro z.
|
intro z.
|
||||||
reflexivity.
|
reflexivity.
|
||||||
Defined.
|
Defined.
|
||||||
|
|
||||||
|
|
Loading…
Reference in New Issue