mirror of
				https://github.com/nmvdw/HITs-Examples
				synced 2025-11-03 23:23:51 +01:00 
			
		
		
		
	Separated lemmas for extensionality for properties, added tactic toHProp
This commit is contained in:
		@@ -9,9 +9,9 @@ fsets/operations_cons_repr.v
 | 
				
			|||||||
fsets/properties_cons_repr.v
 | 
					fsets/properties_cons_repr.v
 | 
				
			||||||
fsets/isomorphism.v
 | 
					fsets/isomorphism.v
 | 
				
			||||||
fsets/operations.v
 | 
					fsets/operations.v
 | 
				
			||||||
fsets/properties.v
 | 
					 | 
				
			||||||
fsets/operations_decidable.v
 | 
					fsets/operations_decidable.v
 | 
				
			||||||
fsets/extensionality.v
 | 
					fsets/extensionality.v
 | 
				
			||||||
 | 
					fsets/properties.v
 | 
				
			||||||
fsets/properties_decidable.v
 | 
					fsets/properties_decidable.v
 | 
				
			||||||
fsets/length.v
 | 
					fsets/length.v
 | 
				
			||||||
fsets/monad.v
 | 
					fsets/monad.v
 | 
				
			||||||
 
 | 
				
			|||||||
@@ -1,12 +1,55 @@
 | 
				
			|||||||
(** Extensionality of the FSets *)
 | 
					(** Extensionality of the FSets *)
 | 
				
			||||||
Require Import HoTT HitTactics.
 | 
					Require Import HoTT HitTactics.
 | 
				
			||||||
From representations Require Import definition.
 | 
					Require Import representations.definition fsets.operations.
 | 
				
			||||||
From fsets Require Import operations properties.
 | 
					 | 
				
			||||||
 | 
					
 | 
				
			||||||
Section ext.
 | 
					Section ext.
 | 
				
			||||||
  Context {A : Type}.
 | 
					  Context {A : Type}.
 | 
				
			||||||
  Context `{Univalence}.
 | 
					  Context `{Univalence}.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					  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.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Lemma subset_union_equiv
 | 
					  Lemma subset_union_equiv
 | 
				
			||||||
    : forall X Y : FSet A, X ⊆ Y <~> X ∪ Y = Y.
 | 
					    : forall X Y : FSet A, X ⊆ Y <~> X ∪ Y = Y.
 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
 
 | 
				
			|||||||
@@ -19,14 +19,16 @@ Section Length.
 | 
				
			|||||||
      pose (y' := FSetC_to_FSet y).
 | 
					      pose (y' := FSetC_to_FSet y).
 | 
				
			||||||
      exact (if isIn_b a y' then n else (S n)).
 | 
					      exact (if isIn_b a y' then n else (S n)).
 | 
				
			||||||
    - intros. rewrite transport_const. cbn.
 | 
					    - intros. rewrite transport_const. cbn.
 | 
				
			||||||
      simplify_isIn. simpl. reflexivity.
 | 
					      simplify_isIn_b. simpl. reflexivity.
 | 
				
			||||||
    - intros. rewrite transport_const. cbn.
 | 
					    - intros. rewrite transport_const. cbn.
 | 
				
			||||||
      simplify_isIn.
 | 
					      simplify_isIn_b.
 | 
				
			||||||
      destruct (dec (a = b)) as [Hab | Hab].
 | 
					      destruct (dec (a = b)) as [Hab | Hab].
 | 
				
			||||||
      + rewrite Hab. simplify_isIn. simpl. reflexivity.
 | 
					      + rewrite Hab. simplify_isIn_b. simpl. reflexivity.
 | 
				
			||||||
      + rewrite ?L_isIn_b_false; auto. simpl. 
 | 
					      + rewrite ?L_isIn_b_false; auto.
 | 
				
			||||||
        destruct (isIn_b a (FSetC_to_FSet x0)), (isIn_b b (FSetC_to_FSet x0)) ; reflexivity.
 | 
					        ++ simpl. 
 | 
				
			||||||
        intro p. contradiction (Hab p^).
 | 
					           destruct (isIn_b a (FSetC_to_FSet x0)), (isIn_b b (FSetC_to_FSet x0))
 | 
				
			||||||
 | 
					           ; reflexivity.
 | 
				
			||||||
 | 
					        ++ intro p. contradiction (Hab p^).
 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Definition length_FSet (x: FSet A) := length (FSet_to_FSetC x).
 | 
					  Definition length_FSet (x: FSet A) := length (FSet_to_FSetC x).
 | 
				
			||||||
 
 | 
				
			|||||||
@@ -76,24 +76,7 @@ Section operations.
 | 
				
			|||||||
    intros ; symmetry ; eauto with lattice_hints typeclass_instances.
 | 
					    intros ; symmetry ; eauto with lattice_hints typeclass_instances.
 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  Lemma union_idemL Z : forall x: FSet Z, x ∪ x = x.
 | 
					  Definition single_product {B : Type} (a : A) : FSet B -> FSet (A * B).
 | 
				
			||||||
  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.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Context {B : Type}.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Definition single_product (a : A) : FSet B -> FSet (A * B).
 | 
					 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    hrecursion.
 | 
					    hrecursion.
 | 
				
			||||||
    - apply ∅.
 | 
					    - apply ∅.
 | 
				
			||||||
@@ -107,7 +90,7 @@ Section operations.
 | 
				
			|||||||
    - intros ; apply idem.
 | 
					    - intros ; apply idem.
 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
        
 | 
					        
 | 
				
			||||||
  Definition product : FSet A -> FSet B -> FSet (A * B).
 | 
					  Definition product {B : Type} : FSet A -> FSet B -> FSet (A * B).
 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    intros X Y.
 | 
					    intros X Y.
 | 
				
			||||||
    hrecursion X.
 | 
					    hrecursion X.
 | 
				
			||||||
@@ -119,7 +102,7 @@ Section operations.
 | 
				
			|||||||
    - intros ; apply comm.
 | 
					    - intros ; apply comm.
 | 
				
			||||||
    - intros ; apply nl.
 | 
					    - intros ; apply nl.
 | 
				
			||||||
    - intros ; apply nr.
 | 
					    - intros ; apply nr.
 | 
				
			||||||
    - intros ; apply union_idemL.
 | 
					    - intros ; apply union_idem.
 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
  
 | 
					  
 | 
				
			||||||
End operations.
 | 
					End operations.
 | 
				
			||||||
 
 | 
				
			|||||||
@@ -1,102 +1,9 @@
 | 
				
			|||||||
Require Import HoTT HitTactics.
 | 
					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 *)
 | 
					(* Lemmas relating operations to the membership predicate *)
 | 
				
			||||||
Section operations_isIn.
 | 
					Section characterize_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.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
  Context {A : Type}.
 | 
					  Context {A : Type}.
 | 
				
			||||||
  Context `{Univalence}.
 | 
					  Context `{Univalence}.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
@@ -196,39 +103,42 @@ Section properties.
 | 
				
			|||||||
        ** left ; split ; assumption.
 | 
					        ** left ; split ; assumption.
 | 
				
			||||||
        ** right ; split ; assumption.
 | 
					        ** right ; split ; assumption.
 | 
				
			||||||
  Defined.
 | 
					  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 *)
 | 
					  (** comprehension properties *)
 | 
				
			||||||
  Lemma comprehension_false Y : comprehension (fun (_ : A) => false) Y = ∅.
 | 
					  Lemma comprehension_false Y : comprehension (fun (_ : A) => false) Y = ∅.
 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    hrecursion Y; try (intros; apply set_path2).
 | 
					    toHProp.
 | 
				
			||||||
    - reflexivity.
 | 
					 | 
				
			||||||
    - reflexivity.
 | 
					 | 
				
			||||||
    - intros x y IHa IHb.
 | 
					 | 
				
			||||||
      rewrite IHa, IHb.
 | 
					 | 
				
			||||||
      apply union_idem.
 | 
					 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  (* Can be simplified using extensionality *)
 | 
					 | 
				
			||||||
  Lemma comprehension_subset : forall ϕ (X : FSet A),
 | 
					  Lemma comprehension_subset : forall ϕ (X : FSet A),
 | 
				
			||||||
      (comprehension ϕ X) ∪ X = X.
 | 
					      (comprehension ϕ X) ∪ X = X.
 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
    intros ϕ.
 | 
					    toHProp.
 | 
				
			||||||
    hrecursion; try (intros ; apply set_path2) ; cbn.
 | 
					    destruct (ϕ a) ; eauto with lattice_hints typeclass_instances.
 | 
				
			||||||
    - apply union_idem.
 | 
					  Defined.
 | 
				
			||||||
    - intro a.
 | 
					
 | 
				
			||||||
      destruct (ϕ a).
 | 
					  Lemma comprehension_or : forall ϕ ψ (x: FSet A),
 | 
				
			||||||
      * apply union_idem.
 | 
					      comprehension (fun a => orb (ϕ a) (ψ a)) x
 | 
				
			||||||
      * apply nl.
 | 
					      = (comprehension ϕ x) ∪ (comprehension ψ x).
 | 
				
			||||||
    - intros X Y P Q.
 | 
					  Proof.
 | 
				
			||||||
      rewrite assoc.
 | 
					    toHProp.
 | 
				
			||||||
      rewrite <- (assoc (comprehension ϕ X) (comprehension ϕ Y) X).
 | 
					    symmetry ; destruct (ϕ a) ; destruct (ψ a)
 | 
				
			||||||
      rewrite (comm (comprehension ϕ Y) X).
 | 
					    ; eauto with lattice_hints typeclass_instances.
 | 
				
			||||||
      rewrite assoc.
 | 
					 | 
				
			||||||
      rewrite P.
 | 
					 | 
				
			||||||
      rewrite <- assoc.
 | 
					 | 
				
			||||||
      rewrite Q.
 | 
					 | 
				
			||||||
      reflexivity.
 | 
					 | 
				
			||||||
  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)).
 | 
				
			||||||
@@ -257,6 +167,7 @@ Section properties.
 | 
				
			|||||||
        apply (tr (inl Xa)).
 | 
					        apply (tr (inl Xa)).
 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					(*
 | 
				
			||||||
  Lemma separation : forall (X : FSet A) (f : {a | a ∈ X} -> B),
 | 
					  Lemma separation : forall (X : FSet A) (f : {a | a ∈ X} -> B),
 | 
				
			||||||
      hexists (fun Y : FSet B => forall (b : B),
 | 
					      hexists (fun Y : FSet B => forall (b : B),
 | 
				
			||||||
                   b ∈ Y = hexists (fun a => hexists (fun (p : a ∈ X) => f (a;p) = b))).
 | 
					                   b ∈ Y = hexists (fun a => hexists (fun (p : a ∈ X) => f (a;p) = b))).
 | 
				
			||||||
@@ -358,5 +269,6 @@ Section properties.
 | 
				
			|||||||
           repeat f_ap.
 | 
					           repeat f_ap.
 | 
				
			||||||
           apply path_ishprop.
 | 
					           apply path_ishprop.
 | 
				
			||||||
  Defined.
 | 
					  Defined.
 | 
				
			||||||
 | 
					*)
 | 
				
			||||||
  
 | 
					  
 | 
				
			||||||
End properties.
 | 
					End properties.
 | 
				
			||||||
 
 | 
				
			|||||||
@@ -87,10 +87,8 @@ Section operations_isIn.
 | 
				
			|||||||
 | 
					
 | 
				
			||||||
End operations_isIn.
 | 
					End operations_isIn.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
Global Opaque isIn_b.
 | 
					 | 
				
			||||||
 | 
					 | 
				
			||||||
(* Some suporting tactics *)
 | 
					(* Some suporting tactics *)
 | 
				
			||||||
Ltac simplify_isIn :=
 | 
					Ltac simplify_isIn_b :=
 | 
				
			||||||
  repeat (rewrite union_isIn_b
 | 
					  repeat (rewrite union_isIn_b
 | 
				
			||||||
        || rewrite L_isIn_b_aa
 | 
					        || rewrite L_isIn_b_aa
 | 
				
			||||||
        || rewrite intersection_isIn_b
 | 
					        || rewrite intersection_isIn_b
 | 
				
			||||||
@@ -99,7 +97,7 @@ Ltac simplify_isIn :=
 | 
				
			|||||||
Ltac toBool :=
 | 
					Ltac toBool :=
 | 
				
			||||||
  repeat intro;
 | 
					  repeat intro;
 | 
				
			||||||
  apply ext ; intros ;
 | 
					  apply ext ; intros ;
 | 
				
			||||||
  simplify_isIn ; eauto with bool_lattice_hints typeclass_instances.
 | 
					  simplify_isIn_b ; eauto with bool_lattice_hints typeclass_instances.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
Section SetLattice.
 | 
					Section SetLattice.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 
 | 
				
			|||||||
@@ -197,3 +197,18 @@ End FSet.
 | 
				
			|||||||
Notation "{| x |}" :=  (L x).
 | 
					Notation "{| x |}" :=  (L x).
 | 
				
			||||||
Infix "∪" := U (at level 8, right associativity).
 | 
					Infix "∪" := U (at level 8, right associativity).
 | 
				
			||||||
Notation "∅" := E.
 | 
					Notation "∅" := E.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					Lemma union_idem {A : Type} : 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.
 | 
				
			||||||
		Reference in New Issue
	
	Block a user