mirror of
				https://github.com/nmvdw/HITs-Examples
				synced 2025-11-03 23:23:51 +01:00 
			
		
		
		
	Added join-semilattice
This commit is contained in:
		@@ -1,6 +1,7 @@
 | 
				
			|||||||
Require Import HoTT HitTactics.
 | 
					Require Import HoTT HitTactics.
 | 
				
			||||||
From fsets Require Import operations extensionality.
 | 
					From fsets Require Import operations extensionality.
 | 
				
			||||||
Require Export representations.definition disjunction.
 | 
					Require Export representations.definition disjunction.
 | 
				
			||||||
 | 
					Require Import lattice.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
(* Lemmas relating operations to the membership predicate *)
 | 
					(* Lemmas relating operations to the membership predicate *)
 | 
				
			||||||
Section characterize_isIn.
 | 
					Section characterize_isIn.
 | 
				
			||||||
@@ -119,6 +120,14 @@ Section properties.
 | 
				
			|||||||
  Context {A : Type}.
 | 
					  Context {A : Type}.
 | 
				
			||||||
  Context `{Univalence}.
 | 
					  Context `{Univalence}.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					  Instance: bottom(FSet A) := ∅.
 | 
				
			||||||
 | 
					  Instance: maximum(FSet A) := U.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					  Global Instance joinsemilattice_fset : JoinSemiLattice (FSet A).
 | 
				
			||||||
 | 
					  Proof.
 | 
				
			||||||
 | 
					    split ; toHProp.
 | 
				
			||||||
 | 
					  Defined.  
 | 
				
			||||||
 | 
					
 | 
				
			||||||
  (** comprehension properties *)
 | 
					  (** comprehension properties *)
 | 
				
			||||||
  Lemma comprehension_false Y : comprehension (fun (_ : A) => false) Y = ∅.
 | 
					  Lemma comprehension_false Y : comprehension (fun (_ : A) => false) Y = ∅.
 | 
				
			||||||
  Proof.
 | 
					  Proof.
 | 
				
			||||||
 
 | 
				
			|||||||
@@ -55,6 +55,29 @@ Arguments NeutralL {_} _ _.
 | 
				
			|||||||
Arguments NeutralR {_} _ _.
 | 
					Arguments NeutralR {_} _ _.
 | 
				
			||||||
Arguments Absorption {_} _ _.
 | 
					Arguments Absorption {_} _ _.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					Section JoinSemiLattice.
 | 
				
			||||||
 | 
					  Variable A : Type.
 | 
				
			||||||
 | 
					  Context {max_L : maximum A} {empty_L : bottom A}.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					  Class JoinSemiLattice :=
 | 
				
			||||||
 | 
					    {
 | 
				
			||||||
 | 
					      commutative_max_js :> Commutative max_L ;
 | 
				
			||||||
 | 
					      associative_max_js :> Associative max_L ;
 | 
				
			||||||
 | 
					      idempotent_max_js :> Idempotent max_L ;
 | 
				
			||||||
 | 
					      neutralL_max_js :> NeutralL max_L empty_L ;
 | 
				
			||||||
 | 
					      neutralR_max_js :> NeutralR max_L empty_L ;
 | 
				
			||||||
 | 
					    }.
 | 
				
			||||||
 | 
					End JoinSemiLattice.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					Arguments JoinSemiLattice _ {_} {_}.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
 | 
					Create HintDb joinsemilattice_hints.
 | 
				
			||||||
 | 
					Hint Resolve associativity : joinsemilattice_hints.
 | 
				
			||||||
 | 
					Hint Resolve commutative : joinsemilattice_hints.
 | 
				
			||||||
 | 
					Hint Resolve idempotency : joinsemilattice_hints.
 | 
				
			||||||
 | 
					Hint Resolve neutralityL : joinsemilattice_hints.
 | 
				
			||||||
 | 
					Hint Resolve neutralityR : joinsemilattice_hints.
 | 
				
			||||||
 | 
					
 | 
				
			||||||
Section Lattice.
 | 
					Section Lattice.
 | 
				
			||||||
  Variable A : Type.
 | 
					  Variable A : Type.
 | 
				
			||||||
  Context {max_L : maximum A} {min_L : minimum A} {empty_L : bottom A}.
 | 
					  Context {max_L : maximum A} {min_L : minimum A} {empty_L : bottom A}.
 | 
				
			||||||
 
 | 
				
			|||||||
		Reference in New Issue
	
	Block a user