1
0
mirror of https://github.com/nmvdw/HITs-Examples synced 2025-11-03 15:13:51 +01:00

Changed lattice

This commit is contained in:
Niels
2017-08-03 12:21:34 +02:00
parent 77a449e68b
commit 7d74b45fc3
7 changed files with 375 additions and 577 deletions

View File

@@ -75,11 +75,11 @@ Proof.
- apply true.
- apply (fun _ => false).
- apply andb.
- intros. symmetry. eauto with bool_lattice_hints.
- eauto with bool_lattice_hints.
- eauto with bool_lattice_hints.
- eauto with bool_lattice_hints.
- eauto with bool_lattice_hints.
- intros. symmetry. eauto with lattice_hints typeclass_instances.
- eauto with bool_lattice_hints typeclass_instances.
- eauto with bool_lattice_hints typeclass_instances.
- eauto with bool_lattice_hints typeclass_instances.
- eauto with bool_lattice_hints typeclass_instances.
Defined.
End operations.

View File

@@ -12,7 +12,14 @@ Section decidable_A.
intros a.
hinduction ; try (intros ; apply path_ishprop).
- apply _.
- intros. apply _.
- intros.
unfold Decidable.
destruct (dec (a = a0)) as [p | np].
* apply (inl (tr p)).
* right.
intro p.
strip_truncations.
contradiction.
- intros. apply _.
Defined.

View File

@@ -4,153 +4,151 @@ Require Export representations.definition disjunction fsets.operations.
(* Lemmas relating operations to the membership predicate *)
Section operations_isIn.
Context {A : Type}.
Context `{Univalence}.
Context {A : Type}.
Context `{Univalence}.
Lemma union_idem : forall x: FSet A, U 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.
Qed.
(** ** Properties about subset relation. *)
Lemma subset_union (X Y : FSet A) :
subset X Y -> U 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 union_idem : forall x: FSet A, U x x = x.
Proof.
hinduction;
try (intros ; apply set_path2) ; cbn.
- 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.
Lemma subset_union_l (X : FSet A) :
forall Y, subset X (U 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 ; left ; apply tr ; reflexivity.
- intros X1 X2 HX1 HX2 Y.
split.
* rewrite <- assoc. apply HX1.
* rewrite (comm X1 X2). rewrite <- assoc. apply HX2.
Defined.
(** ** Properties about subset relation. *)
Lemma subset_union (X Y : FSet A) :
subset X Y -> U 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 ; cbn.
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.
(* Union and membership *)
Lemma union_isIn (X Y : FSet A) (a : A) :
isIn a (U X Y) = isIn a X isIn a Y.
Proof.
reflexivity.
Defined.
Lemma subset_union_l (X : FSet A) :
forall Y, subset X (U 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 ; left ; apply tr ; reflexivity.
- intros X1 X2 HX1 HX2 Y.
split.
* rewrite <- assoc. apply HX1.
* rewrite (comm X1 X2). rewrite <- assoc. apply HX2.
Defined.
(* Union and membership *)
Lemma union_isIn (X Y : FSet A) (a : A) :
isIn a (U X Y) = isIn a X isIn a Y.
Proof.
reflexivity.
Defined.
Lemma comprehension_or : forall ϕ ψ (x: FSet A),
comprehension (fun a => orb (ϕ a) (ψ a)) x = U (comprehension ϕ x)
(comprehension ψ x).
Proof.
intros ϕ ψ.
hinduction; try (intros; apply set_path2).
- cbn. apply (union_idem _)^.
- cbn. 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.
Lemma comprehension_or : forall ϕ ψ (x: FSet A),
comprehension (fun a => orb (ϕ a) (ψ a)) x = U (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.
Context {A : Type}.
Context `{Univalence}.
Context {A : Type}.
Context `{Univalence}.
(** isIn properties *)
Lemma singleton_isIn (a b: A) : isIn a (L b) -> Trunc (-1) (a = b).
Proof.
apply idmap.
Defined.
(** isIn properties *)
Lemma singleton_isIn (a b: A) : isIn a (L b) -> Trunc (-1) (a = b).
Proof.
apply idmap.
Defined.
Lemma empty_isIn (a: A) : isIn a E -> Empty.
Proof.
apply idmap.
Defined.
Lemma empty_isIn (a: A) : isIn a E -> Empty.
Proof.
apply idmap.
Defined.
(** comprehension properties *)
Lemma comprehension_false Y : comprehension (fun (_ : A) => false) Y = E.
Proof.
hrecursion Y; try (intros; apply set_path2).
- reflexivity.
- reflexivity.
- intros x y IHa IHb.
rewrite IHa.
rewrite IHb.
apply union_idem.
Defined.
(** comprehension properties *)
Lemma comprehension_false Y : comprehension (fun (_ : A) => false) Y = E.
Proof.
hrecursion Y; try (intros; apply set_path2).
- reflexivity.
- reflexivity.
- intros x y IHa IHb.
rewrite IHa.
rewrite IHb.
apply union_idem.
Defined.
Lemma comprehension_subset : forall ϕ (X : FSet A),
U (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.
Defined.
Lemma comprehension_subset : forall ϕ (X : FSet A),
U (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.
Defined.
End properties.

View File

@@ -6,143 +6,150 @@ Require Export lattice.
(* Lemmas relating operations to the membership predicate *)
Section operations_isIn.
Context {A : Type} `{DecidablePaths A} `{Univalence}.
Context {A : Type} `{DecidablePaths A} `{Univalence}.
Lemma ext : forall (S T : FSet A), (forall a, isIn_b a S = isIn_b a T) -> S = T.
Proof.
intros X Y H2.
apply fset_ext.
intro a.
specialize (H2 a).
unfold isIn_b, dec in H2.
destruct (isIn_decidable a X) ; destruct (isIn_decidable a Y).
- apply path_iff_hprop ; intro ; assumption.
- contradiction (true_ne_false).
- contradiction (true_ne_false) ; apply H2^.
- apply path_iff_hprop ; intro ; contradiction.
Defined.
Lemma ext : forall (S T : FSet A), (forall a, isIn_b a S = isIn_b a T) -> S = T.
Proof.
intros X Y H2.
apply fset_ext.
intro a.
specialize (H2 a).
unfold isIn_b, dec in H2.
destruct (isIn_decidable a X) ; destruct (isIn_decidable a Y).
- apply path_iff_hprop ; intro ; assumption.
- contradiction (true_ne_false).
- contradiction (true_ne_false) ; apply H2^.
- apply path_iff_hprop ; intro ; contradiction.
Defined.
Lemma empty_isIn (a : A) :
isIn_b a = false.
Proof.
reflexivity.
Defined.
Lemma empty_isIn (a : A) :
isIn_b a = false.
Proof.
reflexivity.
Defined.
Lemma L_isIn (a b : A) :
isIn a {|b|} -> a = b.
Proof.
intros. strip_truncations. assumption.
Defined.
Lemma L_isIn (a b : A) :
isIn a {|b|} -> a = b.
Proof.
intros. strip_truncations. assumption.
Defined.
Lemma L_isIn_b_true (a b : A) (p : a = b) :
isIn_b a {|b|} = true.
Proof.
unfold isIn_b, dec.
destruct (isIn_decidable a {|b|}) as [n | n] .
- reflexivity.
- simpl in n.
contradiction (n (tr p)).
Defined.
Lemma L_isIn_b_true (a b : A) (p : a = b) :
isIn_b a {|b|} = true.
Proof.
unfold isIn_b, dec.
destruct (isIn_decidable a {|b|}) as [n | n] .
- reflexivity.
- simpl in n.
contradiction (n (tr p)).
Defined.
Lemma L_isIn_b_aa (a : A) :
isIn_b a {|a|} = true.
Proof.
apply L_isIn_b_true ; reflexivity.
Defined.
Lemma L_isIn_b_aa (a : A) :
isIn_b a {|a|} = true.
Proof.
apply L_isIn_b_true ; reflexivity.
Defined.
Lemma L_isIn_b_false (a b : A) (p : a <> b) :
isIn_b a {|b|} = false.
Proof.
unfold isIn_b, dec.
destruct (isIn_decidable a {|b|}).
- simpl in t.
strip_truncations.
contradiction.
- reflexivity.
Defined.
(* Union and membership *)
Lemma union_isIn (X Y : FSet A) (a : A) :
isIn_b a (U X Y) = orb (isIn_b a X) (isIn_b a Y).
Proof.
unfold isIn_b ; unfold dec.
simpl.
destruct (isIn_decidable a X) ; destruct (isIn_decidable a Y) ; reflexivity.
Defined.
Lemma L_isIn_b_false (a b : A) (p : a <> b) :
isIn_b a {|b|} = false.
Proof.
unfold isIn_b, dec.
destruct (isIn_decidable a {|b|}).
- simpl in t.
strip_truncations.
contradiction.
- reflexivity.
Defined.
(* Union and membership *)
Lemma union_isIn (X Y : FSet A) (a : A) :
isIn_b a (U X Y) = orb (isIn_b a X) (isIn_b a Y).
Proof.
unfold isIn_b ; unfold dec.
simpl.
destruct (isIn_decidable a X) ; destruct (isIn_decidable a Y) ; reflexivity.
Defined.
Lemma intersection_isIn (X Y: FSet A) (a : A) :
Lemma intersection_isIn (X Y: FSet A) (a : A) :
isIn_b a (intersection X Y) = andb (isIn_b a X) (isIn_b a Y).
Proof.
hinduction X; try (intros ; apply set_path2).
- reflexivity.
- intro b.
destruct (dec (a = b)).
* rewrite p.
destruct (isIn_b b Y) ; symmetry ; eauto with bool_lattice_hints.
* destruct (isIn_b b Y) ; destruct (isIn_b a Y) ; symmetry ; eauto with bool_lattice_hints.
Proof.
hinduction X; try (intros ; apply set_path2).
- reflexivity.
- intro b.
destruct (dec (a = b)).
* rewrite p.
destruct (isIn_b b Y) ; symmetry ; eauto with bool_lattice_hints.
* destruct (isIn_b b Y) ; destruct (isIn_b a Y) ; symmetry ; eauto with bool_lattice_hints.
+ rewrite and_false.
symmetry.
apply (L_isIn_b_false a b n).
+ rewrite and_true.
apply (L_isIn_b_false a b n).
- intros X1 X2 P Q.
rewrite union_isIn ; rewrite union_isIn.
rewrite P.
rewrite Q.
unfold isIn_b, dec.
destruct (isIn_decidable a X1)
; destruct (isIn_decidable a X2)
; destruct (isIn_decidable a Y)
; reflexivity.
Defined.
- intros X1 X2 P Q.
rewrite union_isIn ; rewrite union_isIn.
rewrite P.
rewrite Q.
unfold isIn_b, dec.
destruct (isIn_decidable a X1)
; destruct (isIn_decidable a X2)
; destruct (isIn_decidable a Y)
; reflexivity.
Defined.
Lemma comprehension_isIn (Y : FSet A) (ϕ : A -> Bool) (a : A) :
isIn_b a (comprehension ϕ Y) = andb (isIn_b a Y) (ϕ a).
Proof.
hinduction Y ; try (intros; apply set_path2).
- apply empty_isIn.
- intro b.
destruct (isIn_decidable a {|b|}).
* simpl in t.
strip_truncations.
rewrite t.
destruct (ϕ b).
** rewrite (L_isIn_b_true _ _ idpath).
eauto with bool_lattice_hints.
** rewrite empty_isIn ; rewrite (L_isIn_b_true _ _ idpath).
eauto with bool_lattice_hints.
* destruct (ϕ b).
** rewrite L_isIn_b_false.
*** eauto with bool_lattice_hints.
*** intro.
apply (n (tr X)).
** rewrite empty_isIn.
rewrite L_isIn_b_false.
*** eauto with bool_lattice_hints.
*** intro.
apply (n (tr X)).
- intros.
Opaque isIn_b.
rewrite ?union_isIn.
rewrite X.
rewrite X0.
assert (forall b1 b2 b3,
(b1 && b2 || b3 && b2)%Bool = ((b1 || b3) && b2)%Bool).
{ intros ; destruct b1, b2, b3 ; reflexivity. }
apply X1.
Defined.
Lemma comprehension_isIn (Y : FSet A) (ϕ : A -> Bool) (a : A) :
isIn_b a (comprehension ϕ Y) = andb (isIn_b a Y) (ϕ a).
Proof.
hinduction Y ; try (intros; apply set_path2).
- apply empty_isIn.
- intro b.
destruct (isIn_decidable a {|b|}).
* simpl in t.
strip_truncations.
rewrite t.
destruct (ϕ b).
** rewrite (L_isIn_b_true _ _ idpath).
eauto with bool_lattice_hints.
** rewrite empty_isIn ; rewrite (L_isIn_b_true _ _ idpath).
eauto with bool_lattice_hints.
* destruct (ϕ b).
** rewrite L_isIn_b_false.
*** eauto with bool_lattice_hints.
*** intro.
apply (n (tr X)).
** rewrite empty_isIn.
rewrite L_isIn_b_false.
*** eauto with bool_lattice_hints.
*** intro.
apply (n (tr X)).
- intros.
Opaque isIn_b.
rewrite ?union_isIn.
rewrite X.
rewrite X0.
assert (forall b1 b2 b3,
(b1 && b2 || b3 && b2)%Bool = ((b1 || b3) && b2)%Bool).
{ intros ; destruct b1, b2, b3 ; reflexivity. }
apply X1.
Defined.
End operations_isIn.
Global Opaque isIn_b.
(* Some suporting tactics *)
(*
Ltac simplify_isIn :=
repeat (rewrite union_isIn
|| rewrite L_isIn_b_aa
|| rewrite intersection_isIn
|| rewrite comprehension_isIn).
*)
Ltac simplify_isIn :=
repeat (rewrite union_isIn
|| rewrite L_isIn_b_aa
|| rewrite intersection_isIn
|| rewrite comprehension_isIn).
Ltac toBool := try (intro) ;
intros ; apply ext ; intros ; simplify_isIn ; eauto with bool_lattice_hints.
intros ; apply ext ; intros ; simplify_isIn ; eauto with bool_lattice_hints typeclass_instances.
Section SetLattice.
@@ -150,70 +157,15 @@ Section SetLattice.
Context {A_deceq : DecidablePaths A}.
Context `{Univalence}.
Instance fset_union_com : Commutative (@U A).
Proof.
toBool.
Defined.
Instance fset_max : maximum (FSet A) := U.
Instance fset_min : minimum (FSet A) := intersection.
Instance fset_bot : bottom (FSet A) := E.
Instance fset_intersection_com : Commutative intersection.
Instance lattice_fset : Lattice (FSet A).
Proof.
toBool.
Defined.
Instance fset_union_assoc : Associative (@U A).
Proof.
toBool.
Defined.
Instance fset_intersection_assoc : Associative intersection.
Proof.
toBool.
Defined.
Instance fset_union_idem : Idempotent (@U A).
Proof.
exact union_idem.
Defined.
Instance fset_intersection_idem : Idempotent intersection.
Proof.
toBool.
Defined.
Instance fset_union_nl : NeutralL (@U A) (@E A).
Proof.
toBool.
Defined.
Instance fset_union_nr : NeutralR (@U A) (@E A).
Proof.
toBool.
Defined.
Instance fset_absorption_intersection_union : Absorption (@U A) intersection.
Proof.
toBool.
Defined.
Instance fset_absorption_union_intersection : Absorption intersection (@U A).
Proof.
toBool.
split ; toBool.
Defined.
Instance lattice_fset : Lattice intersection (@U A) (@E A) :=
{
commutative_min := _ ;
commutative_max := _ ;
associative_min := _ ;
associative_max := _ ;
idempotent_min := _ ;
idempotent_max := _ ;
neutralL_min := _ ;
neutralR_min := _ ;
absorption_min_max := _ ;
absorption_max_min := _
}.
End SetLattice.
(* Comprehension properties *)
@@ -236,16 +188,16 @@ Section comprehension_properties.
Lemma comprehension_false Y : comprehension (fun (_ : A) => false) Y = E.
Proof.
toBool.
Defined.
Qed.
Lemma comprehension_all : forall (X : FSet A),
comprehension (fun a => isIn_b a X) X = X.
Proof.
toBool.
Defined.
Qed.
Lemma comprehension_subset : forall ϕ (X : FSet A),
U (comprehension ϕ X) X = X.
U (comprehension ϕ X) X = X.
Proof.
toBool.
Defined.
@@ -266,7 +218,7 @@ Section dec_eq.
- right. intros [p1 p2]. contradiction.
- right. intros [p1 p2]. contradiction.
Defined.
Instance fsets_dec_eq : DecidablePaths (FSet A).
Proof.
intros X Y.