mirror of
https://github.com/nmvdw/HITs-Examples
synced 2025-12-16 23:23:50 +01:00
Separated lemmas for extensionality for properties, added tactic toHProp
This commit is contained in:
@@ -1,102 +1,9 @@
|
||||
Require Import HoTT HitTactics.
|
||||
Require Export representations.definition disjunction fsets.operations.
|
||||
From fsets Require Import operations extensionality.
|
||||
Require Export representations.definition disjunction.
|
||||
|
||||
(* Lemmas relating operations to the membership predicate *)
|
||||
Section operations_isIn.
|
||||
|
||||
Context {A : Type}.
|
||||
Context `{Univalence}.
|
||||
|
||||
Lemma union_idem : forall x: FSet A, x ∪ x = x.
|
||||
Proof.
|
||||
hinduction ; try (intros ; apply set_path2).
|
||||
- apply nl.
|
||||
- apply idem.
|
||||
- intros x y P Q.
|
||||
rewrite assoc.
|
||||
rewrite (comm x y).
|
||||
rewrite <- (assoc y x x).
|
||||
rewrite P.
|
||||
rewrite (comm y x).
|
||||
rewrite <- (assoc x y y).
|
||||
f_ap.
|
||||
Defined.
|
||||
|
||||
(** ** Properties about subset relation. *)
|
||||
Lemma subset_union (X Y : FSet A) :
|
||||
X ⊆ Y -> X ∪ Y = Y.
|
||||
Proof.
|
||||
hinduction X ; try (intros; apply path_forall; intro; apply set_path2).
|
||||
- intros. apply nl.
|
||||
- intros a.
|
||||
hinduction Y ; try (intros; apply path_forall; intro; apply set_path2).
|
||||
+ intro.
|
||||
contradiction.
|
||||
+ intro a0.
|
||||
simple refine (Trunc_ind _ _).
|
||||
intro p ; simpl.
|
||||
rewrite p; apply idem.
|
||||
+ intros X1 X2 IH1 IH2.
|
||||
simple refine (Trunc_ind _ _).
|
||||
intros [e1 | e2].
|
||||
++ rewrite assoc.
|
||||
rewrite (IH1 e1).
|
||||
reflexivity.
|
||||
++ rewrite comm.
|
||||
rewrite <- assoc.
|
||||
rewrite (comm X2).
|
||||
rewrite (IH2 e2).
|
||||
reflexivity.
|
||||
- intros X1 X2 IH1 IH2 [G1 G2].
|
||||
rewrite <- assoc.
|
||||
rewrite (IH2 G2).
|
||||
apply (IH1 G1).
|
||||
Defined.
|
||||
|
||||
Lemma subset_union_l (X : FSet A) :
|
||||
forall Y, X ⊆ X ∪ Y.
|
||||
Proof.
|
||||
hinduction X ; try (repeat (intro; intros; apply path_forall);
|
||||
intro ; apply equiv_hprop_allpath ; apply _).
|
||||
- apply (fun _ => tt).
|
||||
- intros a Y.
|
||||
apply (tr(inl(tr idpath))).
|
||||
- intros X1 X2 HX1 HX2 Y.
|
||||
split.
|
||||
* rewrite <- assoc. apply HX1.
|
||||
* rewrite (comm X1 X2). rewrite <- assoc. apply HX2.
|
||||
Defined.
|
||||
|
||||
(* simplify it using extensionality *)
|
||||
Lemma comprehension_or : forall ϕ ψ (x: FSet A),
|
||||
comprehension (fun a => orb (ϕ a) (ψ a)) x
|
||||
= (comprehension ϕ x) ∪ (comprehension ψ x).
|
||||
Proof.
|
||||
intros ϕ ψ.
|
||||
hinduction ; try (intros; apply set_path2).
|
||||
- apply (union_idem _)^.
|
||||
- intros.
|
||||
destruct (ϕ a) ; destruct (ψ a) ; symmetry.
|
||||
* apply union_idem.
|
||||
* apply nr.
|
||||
* apply nl.
|
||||
* apply union_idem.
|
||||
- simpl. intros x y P Q.
|
||||
rewrite P.
|
||||
rewrite Q.
|
||||
rewrite <- assoc.
|
||||
rewrite (assoc (comprehension ψ x)).
|
||||
rewrite (comm (comprehension ψ x) (comprehension ϕ y)).
|
||||
rewrite <- assoc.
|
||||
rewrite <- assoc.
|
||||
reflexivity.
|
||||
Defined.
|
||||
|
||||
End operations_isIn.
|
||||
|
||||
(* Other properties *)
|
||||
Section properties.
|
||||
|
||||
Section characterize_isIn.
|
||||
Context {A : Type}.
|
||||
Context `{Univalence}.
|
||||
|
||||
@@ -174,9 +81,9 @@ Section properties.
|
||||
** right.
|
||||
split ; try (apply (tr H1)) ; try (apply Hb2).
|
||||
Defined.
|
||||
|
||||
|
||||
Definition isIn_product (a : A) (b : B) (X : FSet A) (Y : FSet B) :
|
||||
isIn (a,b) (product X Y) = land (isIn a X) (isIn b Y).
|
||||
isIn (a,b) (product X Y) = land (isIn a X) (isIn b Y).
|
||||
Proof.
|
||||
hinduction X ; try (intros ; apply path_ishprop).
|
||||
- apply path_hprop ; symmetry ; apply prod_empty_l.
|
||||
@@ -196,39 +103,42 @@ Section properties.
|
||||
** left ; split ; assumption.
|
||||
** right ; split ; assumption.
|
||||
Defined.
|
||||
End characterize_isIn.
|
||||
|
||||
Ltac simplify_isIn :=
|
||||
repeat (rewrite union_isIn
|
||||
|| rewrite comprehension_isIn).
|
||||
|
||||
Ltac toHProp :=
|
||||
repeat intro;
|
||||
apply fset_ext ; intros ;
|
||||
simplify_isIn ; eauto with lattice_hints typeclass_instances.
|
||||
|
||||
(* Other properties *)
|
||||
Section properties.
|
||||
Context {A : Type}.
|
||||
Context `{Univalence}.
|
||||
|
||||
(* The proof can be simplified using extensionality *)
|
||||
(** comprehension properties *)
|
||||
Lemma comprehension_false Y : comprehension (fun (_ : A) => false) Y = ∅.
|
||||
Proof.
|
||||
hrecursion Y; try (intros; apply set_path2).
|
||||
- reflexivity.
|
||||
- reflexivity.
|
||||
- intros x y IHa IHb.
|
||||
rewrite IHa, IHb.
|
||||
apply union_idem.
|
||||
toHProp.
|
||||
Defined.
|
||||
|
||||
(* Can be simplified using extensionality *)
|
||||
Lemma comprehension_subset : forall ϕ (X : FSet A),
|
||||
(comprehension ϕ X) ∪ X = X.
|
||||
Proof.
|
||||
intros ϕ.
|
||||
hrecursion; try (intros ; apply set_path2) ; cbn.
|
||||
- apply union_idem.
|
||||
- intro a.
|
||||
destruct (ϕ a).
|
||||
* apply union_idem.
|
||||
* apply nl.
|
||||
- intros X Y P Q.
|
||||
rewrite assoc.
|
||||
rewrite <- (assoc (comprehension ϕ X) (comprehension ϕ Y) X).
|
||||
rewrite (comm (comprehension ϕ Y) X).
|
||||
rewrite assoc.
|
||||
rewrite P.
|
||||
rewrite <- assoc.
|
||||
rewrite Q.
|
||||
reflexivity.
|
||||
toHProp.
|
||||
destruct (ϕ a) ; eauto with lattice_hints typeclass_instances.
|
||||
Defined.
|
||||
|
||||
Lemma comprehension_or : forall ϕ ψ (x: FSet A),
|
||||
comprehension (fun a => orb (ϕ a) (ψ a)) x
|
||||
= (comprehension ϕ x) ∪ (comprehension ψ x).
|
||||
Proof.
|
||||
toHProp.
|
||||
symmetry ; destruct (ϕ a) ; destruct (ψ a)
|
||||
; eauto with lattice_hints typeclass_instances.
|
||||
Defined.
|
||||
|
||||
Lemma merely_choice : forall X : FSet A, hor (X = ∅) (hexists (fun a => a ∈ X)).
|
||||
@@ -257,6 +167,7 @@ Section properties.
|
||||
apply (tr (inl Xa)).
|
||||
Defined.
|
||||
|
||||
(*
|
||||
Lemma separation : forall (X : FSet A) (f : {a | a ∈ X} -> B),
|
||||
hexists (fun Y : FSet B => forall (b : B),
|
||||
b ∈ Y = hexists (fun a => hexists (fun (p : a ∈ X) => f (a;p) = b))).
|
||||
@@ -358,5 +269,6 @@ Section properties.
|
||||
repeat f_ap.
|
||||
apply path_ishprop.
|
||||
Defined.
|
||||
*)
|
||||
|
||||
End properties.
|
||||
|
||||
Reference in New Issue
Block a user