Fixed NeutralL and NeutralR

This commit is contained in:
Niels 2017-08-01 17:25:57 +02:00
parent b3b3e5b6c2
commit e6bf0f9d5d
3 changed files with 21 additions and 19 deletions

View File

@ -9,9 +9,9 @@ Notation "A B" := (lor A B) (at level 20, right associativity) : logic_scope
Arguments lor _%L _%L. Arguments lor _%L _%L.
Open Scope logic_scope. Open Scope logic_scope.
Section lor_props. Section lor_props.
Variable X Y Z : hProp.
Context `{Univalence}. Context `{Univalence}.
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.
@ -45,7 +45,7 @@ Section lor_props.
+ apply (tr (inl x)). + apply (tr (inl x)).
Defined. Defined.
Lemma lor_nr : False_hp X = X. Lemma lor_nl : False_hp X = X.
Proof. Proof.
apply path_iff_hprop ; cbn. apply path_iff_hprop ; cbn.
* simple refine (Trunc_ind _ _). * simple refine (Trunc_ind _ _).
@ -55,7 +55,7 @@ Section lor_props.
* apply (fun x => tr (inr x)). * apply (fun x => tr (inr x)).
Defined. Defined.
Lemma lor_nl : X False_hp = X. Lemma lor_nr : X False_hp = X.
Proof. Proof.
apply path_iff_hprop ; cbn. apply path_iff_hprop ; cbn.
* simple refine (Trunc_ind _ _). * simple refine (Trunc_ind _ _).
@ -89,7 +89,6 @@ Section hPropLattice.
unfold Commutative. unfold Commutative.
intros. intros.
apply lor_comm. apply lor_comm.
apply _.
Defined. Defined.
Instance land_commutative : Commutative land. Instance land_commutative : Commutative land.
@ -105,7 +104,6 @@ Section hPropLattice.
unfold Associative. unfold Associative.
intros. intros.
apply lor_assoc. apply lor_assoc.
apply _.
Defined. Defined.
Instance land_associative : Associative land. Instance land_associative : Associative land.
@ -122,7 +120,6 @@ Section hPropLattice.
unfold Idempotent. unfold Idempotent.
intros. intros.
apply lor_idem. apply lor_idem.
apply _.
Defined. Defined.
Instance land_idempotent : Idempotent land. Instance land_idempotent : Idempotent land.
@ -139,7 +136,6 @@ Section hPropLattice.
unfold NeutralL. unfold NeutralL.
intros. intros.
apply lor_nl. apply lor_nl.
apply _.
Defined. Defined.
Instance lor_neutralr : NeutralR lor False_hp. Instance lor_neutralr : NeutralR lor False_hp.
@ -147,7 +143,6 @@ Section hPropLattice.
unfold NeutralR. unfold NeutralR.
intros. intros.
apply lor_nr. apply lor_nr.
apply _.
Defined. Defined.
Instance bool_absorption_orb_andb : Absorption lor land. Instance bool_absorption_orb_andb : Absorption lor land.
@ -174,7 +169,7 @@ Section hPropLattice.
* apply (tr (inl X)). * apply (tr (inl X)).
Defined. Defined.
Global Instance lattice_bool : Lattice andb orb false := Global Instance lattice_hprop : Lattice land lor False_hp :=
{ commutative_min := _ ; { commutative_min := _ ;
commutative_max := _ ; commutative_max := _ ;
associative_min := _ ; associative_min := _ ;
@ -187,4 +182,11 @@ Section hPropLattice.
absorption_max_min := _ absorption_max_min := _
}. }.
End hPropLattice. End hPropLattice.
Hint Resolve
commutative_min commutative_max associative_min associative_max
idempotent_min idempotent_max
neutralL_min neutralR_min
absorption_min_max absorption_max_min
: lattice_hints.

View File

@ -15,12 +15,12 @@ hrecursion.
- intro a'. - intro a'.
exists (Trunc (-1) (a = a')). exists (Trunc (-1) (a = a')).
exact _. exact _.
- apply lor. - apply lor.
- intros ; apply lor_assoc. exact _. - intros ; symmetry ; apply lor_assoc.
- intros ; apply lor_comm. exact _. - intros ; apply lor_comm.
- intros ; apply lor_nl. exact _. - intros ; apply lor_nl.
- intros ; apply lor_nr. exact _. - intros ; apply lor_nr.
- intros ; apply lor_idem. exact _. - intros ; apply lor_idem.
Defined. Defined.
Definition subset : FSet A -> FSet A -> hProp. Definition subset : FSet A -> FSet A -> hProp.

View File

@ -24,10 +24,10 @@ Section Defs.
Variable n : A. Variable n : A.
Class NeutralL := Class NeutralL :=
neutralityL : forall x, f x n = x. neutralityL : forall x, f n x = x.
Class NeutralR := Class NeutralR :=
neutralityR : forall x, f n x = x. neutralityR : forall x, f x n = x.
End Defs. End Defs.