]> matita.cs.unibo.it Git - helm.git/commitdiff
snapshot
authorEnrico Tassi <enrico.tassi@inria.fr>
Mon, 21 Jan 2008 18:05:11 +0000 (18:05 +0000)
committerEnrico Tassi <enrico.tassi@inria.fr>
Mon, 21 Jan 2008 18:05:11 +0000 (18:05 +0000)
helm/software/matita/dama/depends
helm/software/matita/dama/excess.ma
helm/software/matita/dama/lattice.ma

index 81e17aafc7be49b91fca074e2c7ceb3e96f460b6..e96472aa927e914fa867866ead5bb99728b55229 100644 (file)
@@ -1,4 +1,4 @@
-metric_lattice.ma lattice.ma metric_space.ma
+metric_lattice.ma excess.ma lattice.ma metric_space.ma
 metric_space.ma ordered_divisible_group.ma
 sandwich.ma metric_lattice.ma nat/orders.ma nat/plus.ma sequence.ma
 premetric_lattice.ma lattice.ma metric_space.ma
@@ -7,6 +7,7 @@ divisible_group.ma group.ma nat/orders.ma
 ordered_divisible_group.ma divisible_group.ma nat/orders.ma nat/times.ma ordered_group.ma
 sequence.ma excess.ma nat/orders.ma ordered_group.ma
 constructive_connectives.ma logic/connectives.ma
+lattice.TF.ma excess.ma lattice.ma nat/nat.ma
 group.ma excess.ma
 prevalued_lattice.ma ordered_group.ma
 excess.ma constructive_connectives.ma constructive_higher_order_relations.ma higher_order_defs/relations.ma nat/plus.ma
index 0f236dd06a3ea4b3463740b52be6f23d9620086d..959e866fd6bd66bf436b287c5fd023652da20628 100644 (file)
@@ -19,29 +19,14 @@ include "nat/plus.ma".
 include "constructive_higher_order_relations.ma".
 include "constructive_connectives.ma".
 
-record excess : Type ≝ {
+record excess_base : Type ≝ {
   exc_carr:> Type;
-  exc_relation: exc_carr → exc_carr → Type;
-  exc_coreflexive: coreflexive ? exc_relation;
-  exc_cotransitive: cotransitive ? exc_relation 
+  exc_excess: exc_carr → exc_carr → Type;
+  exc_coreflexive: coreflexive ? exc_excess;
+  exc_cotransitive: cotransitive ? exc_excess 
 }.
 
-interpretation "excess" 'nleq a b =
- (cic:/matita/excess/exc_relation.con _ a b). 
-
-definition le ≝ λE:excess.λa,b:E. ¬ (a ≰ b).
-
-interpretation "ordered sets less or equal than" 'leq a b = 
- (cic:/matita/excess/le.con _ a b).
-
-lemma le_reflexive: ∀E.reflexive ? (le E).
-intros (E); unfold; cases E; simplify; intros (x); apply (H x);
-qed.
-
-lemma le_transitive: ∀E.transitive ? (le E).
-intros (E); unfold; cases E; simplify; unfold Not; intros (x y z Rxy Ryz H2); 
-cases (c x z y H2) (H4 H5); clear H2; [exact (Rxy H4)|exact (Ryz H5)] 
-qed.
+interpretation "excess" 'nleq a b = (cic:/matita/excess/exc_excess.con _ a b). 
 
 record apartness : Type ≝ {
   ap_carr:> Type;
@@ -52,15 +37,10 @@ record apartness : Type ≝ {
 }.
 
 notation "hvbox(a break # b)" non associative with precedence 50 for @{ 'apart $a $b}.
-interpretation "apartness" 'apart x y = 
-  (cic:/matita/excess/ap_apart.con _ x y).
+interpretation "apartness" 'apart x y = (cic:/matita/excess/ap_apart.con _ x y).
 
-definition strong_ext ≝ λA:apartness.λop:A→A.∀x,y. op x # op y → x # y.
-
-definition apart ≝ λE:excess.λa,b:E. a ≰ b ∨ b ≰ a.
-
-definition apart_of_excess: excess → apartness.
-intros (E); apply (mk_apartness E (apart E));  
+definition apartness_of_excess_base: excess_base → apartness.
+intros (E); apply (mk_apartness E (λa,b:E. a ≰ b ∨ b ≰ a));  
 [1: unfold; cases E; simplify; clear E; intros (x); unfold;
     intros (H1); apply (H x); cases H1; assumption;
 |2: unfold; intros(x y H); cases H; clear H; [right|left] assumption;
@@ -69,7 +49,42 @@ intros (E); apply (mk_apartness E (apart E));
     [left; left|right; left|right; right|left; right] assumption]
 qed.
 
-coercion cic:/matita/excess/apart_of_excess.con.
+record excess_ : Type ≝ {
+  exc_exc:> excess_base;
+  exc_ap: apartness;
+  exc_with: ap_carr exc_ap = exc_carr exc_exc
+}.
+
+definition apart_of_excess_: excess_ → apartness.
+intro E; apply (mk_apartness E); unfold Type_OF_excess_; 
+cases (exc_with E); simplify;
+[apply (ap_apart (exc_ap E));
+|apply ap_coreflexive;|apply ap_symmetric;|apply ap_cotransitive] 
+qed.
+
+coercion cic:/matita/excess/apart_of_excess_.con.
+
+record excess : Type ≝ {
+  excess_carr:> excess_;
+  ap2exc: ∀y,x:excess_carr. y # x → y ≰ x ∨ x ≰ y;
+  exc2ap: ∀y,x:excess_carr.y ≰ x ∨ x ≰ y →  y # x 
+}.
+
+definition strong_ext ≝ λA:apartness.λop:A→A.∀x,y. op x # op y → x # y.
+
+definition le ≝ λE:excess.λa,b:E. ¬ (a ≰ b).
+
+interpretation "ordered sets less or equal than" 'leq a b = 
+ (cic:/matita/excess/le.con _ a b).
+
+lemma le_reflexive: ∀E.reflexive ? (le E).
+unfold reflexive; intros 3 (E x H); apply (exc_coreflexive ?? H);
+qed.
+
+lemma le_transitive: ∀E.transitive ? (le E).
+unfold transitive; intros 7 (E x y z H1 H2 H3); cases (exc_cotransitive ??? y H3) (H4 H4);
+[cases (H1 H4)|cases (H2 H4)]
+qed.
 
 definition eq ≝ λA:apartness.λa,b:A. ¬ (a # b).
 
@@ -82,12 +97,12 @@ intros (E); unfold; intros (x); apply ap_coreflexive;
 qed.
 
 lemma eq_sym_:∀E.symmetric ? (eq E).
-intros (E); unfold; intros (x y Exy); unfold; unfold; intros (Ayx); apply Exy;
-apply ap_symmetric; assumption; 
+unfold symmetric; intros 5 (E x y H H1); cases (H (ap_symmetric ??? H1)); 
 qed.
 
 lemma eq_sym:∀E:apartness.∀x,y:E.x ≈ y → y ≈ x ≝ eq_sym_.
 
+(* SETOID REWRITE *)
 coercion cic:/matita/excess/eq_sym.con.
 
 lemma eq_trans_: ∀E.transitive ? (eq E).
@@ -99,23 +114,18 @@ qed.
 lemma eq_trans:∀E:apartness.∀x,z,y:E.x ≈ y → y ≈ z → x ≈ z ≝ 
   λE,x,y,z.eq_trans_ E x z y.
 
-notation > "'Eq'≈" non associative with precedence 50 for
- @{'eqrewrite}.
-interpretation "eq_rew" 'eqrewrite = 
- (cic:/matita/excess/eq_trans.con _ _ _).
+notation > "'Eq'≈" non associative with precedence 50 for @{'eqrewrite}.
+interpretation "eq_rew" 'eqrewrite = (cic:/matita/excess/eq_trans.con _ _ _).
 
-(* BUG: vedere se ricapita *)
 alias id "antisymmetric" = "cic:/matita/constructive_higher_order_relations/antisymmetric.con".
 lemma le_antisymmetric: ∀E.antisymmetric ? (le E) (eq ?).
-intros 5 (E x y Lxy Lyx); intro H;
-cases H; [apply Lxy;|apply Lyx] assumption;
+intros 5 (E x y Lxy Lyx); intro H; 
+cases (ap2exc ??? H); [apply Lxy;|apply Lyx] assumption;
 qed.
 
 definition lt ≝ λE:excess.λa,b:E. a ≤ b ∧ a # b.
 
-interpretation "ordered sets less than" 'lt a b =
- (cic:/matita/excess/lt.con _ a b).
+interpretation "ordered sets less than" 'lt a b = (cic:/matita/excess/lt.con _ a b).
 
 lemma lt_coreflexive: ∀E.coreflexive ? (lt E).
 intros 2 (E x); intro H; cases H (_ ABS); 
@@ -125,43 +135,32 @@ qed.
 lemma lt_transitive: ∀E.transitive ? (lt E).
 intros (E); unfold; intros (x y z H1 H2); cases H1 (Lxy Axy); cases H2 (Lyz Ayz); 
 split; [apply (le_transitive ???? Lxy Lyz)] clear H1 H2;
-cases Axy (H1 H1); cases Ayz (H2 H2); [1:cases (Lxy H1)|3:cases (Lyz H2)]
+elim (ap2exc ??? Axy) (H1 H1); elim (ap2exc ??? Ayz) (H2 H2); [1:cases (Lxy H1)|3:cases (Lyz H2)]
 clear Axy Ayz;lapply (exc_cotransitive E) as c; unfold cotransitive in c;
 lapply (exc_coreflexive E) as r; unfold coreflexive in r;
-[1: lapply (c ?? y H1) as H3; cases H3 (H4 H4); [cases (Lxy H4)|cases (r ? H4)]
-|2: lapply (c ?? x H2) as H3; cases H3 (H4 H4); [right; assumption|cases (Lxy H4)]]
+[1: lapply (c ?? y H1) as H3; elim H3 (H4 H4); [cases (Lxy H4)|cases (r ? H4)]
+|2: lapply (c ?? x H2) as H3; elim H3 (H4 H4); [apply exc2ap; right; assumption|cases (Lxy H4)]]
 qed.
 
 theorem lt_to_excess: ∀E:excess.∀a,b:E. (a < b) → (b ≰ a).
-intros (E a b Lab); cases Lab (LEab Aab);
-cases Aab (H H); [cases (LEab H)] fold normalize (b ≰ a); assumption; (* BUG *)  
-qed.
-
-lemma unfold_apart: ∀E:excess. ∀x,y:E. x ≰ y ∨ y ≰ x → x # y.
-intros; assumption;
+intros (E a b Lab); elim Lab (LEab Aab);
+elim (ap2exc ??? Aab) (H H); [cases (LEab H)] fold normalize (b ≰ a); assumption; (* BUG *)  
 qed.
 
 lemma le_rewl: ∀E:excess.∀z,y,x:E. x ≈ y → x ≤ z → y ≤ z.
 intros (E z y x Exy Lxz); apply (le_transitive ???? ? Lxz);
-intro Xyz; apply Exy; apply unfold_apart; right; assumption;
+intro Xyz; apply Exy; apply exc2ap; right; assumption;
 qed.
 
 lemma le_rewr: ∀E:excess.∀z,y,x:E. x ≈ y → z ≤ x → z ≤ y.
 intros (E z y x Exy Lxz); apply (le_transitive ???? Lxz);
-intro Xyz; apply Exy; apply unfold_apart; left; assumption;
+intro Xyz; apply Exy; apply exc2ap; left; assumption;
 qed.
 
-notation > "'Le'≪" non associative with precedence 50 for
- @{'lerewritel}.
-interpretation "le_rewl" 'lerewritel = 
- (cic:/matita/excess/le_rewl.con _ _ _).
-
-notation > "'Le'≫" non associative with precedence 50 for
- @{'lerewriter}.
-interpretation "le_rewr" 'lerewriter = 
- (cic:/matita/excess/le_rewr.con _ _ _).
+notation > "'Le'≪" non associative with precedence 50 for @{'lerewritel}.
+interpretation "le_rewl" 'lerewritel = (cic:/matita/excess/le_rewl.con _ _ _).
+notation > "'Le'≫" non associative with precedence 50 for @{'lerewriter}.
+interpretation "le_rewr" 'lerewriter = (cic:/matita/excess/le_rewr.con _ _ _).
 
 lemma ap_rewl: ∀A:apartness.∀x,z,y:A. x ≈ y → y # z → x # z.
 intros (A x z y Exy Ayz); cases (ap_cotransitive ???x Ayz); [2:assumption]
@@ -173,39 +172,25 @@ intros (A x z y Exy Azy); apply ap_symmetric; apply (ap_rewl ???? Exy);
 apply ap_symmetric; assumption;
 qed.
 
-notation > "'Ap'≪" non associative with precedence 50 for
- @{'aprewritel}.
-interpretation "ap_rewl" 'aprewritel = 
- (cic:/matita/excess/ap_rewl.con _ _ _).
-
-notation > "'Ap'≫" non associative with precedence 50 for
- @{'aprewriter}.
-interpretation "ap_rewr" 'aprewriter = 
- (cic:/matita/excess/ap_rewr.con _ _ _).
+notation > "'Ap'≪" non associative with precedence 50 for @{'aprewritel}.
+interpretation "ap_rewl" 'aprewritel = (cic:/matita/excess/ap_rewl.con _ _ _).
+notation > "'Ap'≫" non associative with precedence 50 for @{'aprewriter}.
+interpretation "ap_rewr" 'aprewriter = (cic:/matita/excess/ap_rewr.con _ _ _).
 
 lemma exc_rewl: ∀A:excess.∀x,z,y:A. x ≈ y → y ≰ z → x ≰ z.
 intros (A x z y Exy Ayz); elim (exc_cotransitive ???x Ayz); [2:assumption]
-cases Exy; right; assumption;
+cases Exy; apply exc2ap; right; assumption;
 qed.
   
 lemma exc_rewr: ∀A:excess.∀x,z,y:A. x ≈ y → z ≰ y → z ≰ x.
 intros (A x z y Exy Azy); elim (exc_cotransitive ???x Azy); [assumption]
-elim (Exy); left; assumption;
+elim (Exy); apply exc2ap; left; assumption;
 qed.
 
-notation > "'Ex'≪" non associative with precedence 50 for
- @{'excessrewritel}.
-interpretation "exc_rewl" 'excessrewritel = 
- (cic:/matita/excess/exc_rewl.con _ _ _).
-
-notation > "'Ex'≫" non associative with precedence 50 for
- @{'excessrewriter}.
-interpretation "exc_rewr" 'excessrewriter = 
- (cic:/matita/excess/exc_rewr.con _ _ _).
+notation > "'Ex'≪" non associative with precedence 50 for @{'excessrewritel}.
+interpretation "exc_rewl" 'excessrewritel = (cic:/matita/excess/exc_rewl.con _ _ _).
+notation > "'Ex'≫" non associative with precedence 50 for @{'excessrewriter}.
+interpretation "exc_rewr" 'excessrewriter = (cic:/matita/excess/exc_rewr.con _ _ _).
 
 lemma lt_rewr: ∀A:excess.∀x,z,y:A. x ≈ y → z < y → z < x.
 intros (A x y z E H); split; elim H; 
@@ -217,51 +202,53 @@ intros (A x y z E H); split; elim H;
 [apply (Le≪ ? (eq_sym ??? E));| apply (Ap≪ ? E);] assumption;
 qed.
 
-notation > "'Lt'≪" non associative with precedence 50 for
- @{'ltrewritel}.
-interpretation "lt_rewl" 'ltrewritel = 
- (cic:/matita/excess/lt_rewl.con _ _ _).
-
-notation > "'Lt'≫" non associative with precedence 50 for
- @{'ltrewriter}.
-interpretation "lt_rewr" 'ltrewriter = 
- (cic:/matita/excess/lt_rewr.con _ _ _).
+notation > "'Lt'≪" non associative with precedence 50 for @{'ltrewritel}.
+interpretation "lt_rewl" 'ltrewritel = (cic:/matita/excess/lt_rewl.con _ _ _).
+notation > "'Lt'≫" non associative with precedence 50 for @{'ltrewriter}.
+interpretation "lt_rewr" 'ltrewriter = (cic:/matita/excess/lt_rewr.con _ _ _).
 
 lemma lt_le_transitive: ∀A:excess.∀x,y,z:A.x < y → y ≤ z → x < z.
 intros (A x y z LT LE); cases LT (LEx APx); split; [apply (le_transitive ???? LEx LE)]
-whd in LE LEx APx; cases APx (EXx EXx); [cases (LEx EXx)]
+apply exc2ap; cases (ap2exc ??? APx) (EXx EXx); [cases (LEx EXx)]
 cases (exc_cotransitive ??? z EXx) (EXz EXz); [cases (LE EXz)]
 right; assumption;
 qed.
 
 lemma le_lt_transitive: ∀A:excess.∀x,y,z:A.x ≤ y → y < z → x < z.
 intros (A x y z LE LT); cases LT (LEx APx); split; [apply (le_transitive ???? LE LEx)]
-whd in LE LEx APx; cases APx (EXx EXx); [cases (LEx EXx)]
-cases (exc_cotransitive ??? x EXx) (EXz EXz); [right; assumption]
+elim (ap2exc ??? APx) (EXx EXx); [cases (LEx EXx)]
+elim (exc_cotransitive ??? x EXx) (EXz EXz); [apply exc2ap; right; assumption]
 cases LE; assumption;
 qed.
     
 lemma le_le_eq: ∀E:excess.∀a,b:E. a ≤ b → b ≤ a → a ≈ b.
-intros (E x y L1 L2); intro H; cases H; [apply L1|apply L2] assumption;
+intros (E x y L1 L2); intro H; cases (ap2exc ??? H); [apply L1|apply L2] assumption;
 qed.
 
 lemma eq_le_le: ∀E:excess.∀a,b:E. a ≈ b → a ≤ b ∧ b ≤ a.
-intros (E x y H); unfold apart_of_excess in H; unfold apart in H;
-simplify in H; split; intro; apply H; [left|right] assumption.
+intros (E x y H); whd in H;
+split; intro; apply H; apply exc2ap; [left|right] assumption.
 qed.
 
 lemma ap_le_to_lt: ∀E:excess.∀a,c:E.c # a → c ≤ a → c < a.
 intros; split; assumption;
 qed.
 
-definition total_order_property : ∀E:excess. Type ≝
+definition total_order_property : ∀E:excess. Type ≝ 
   λE:excess. ∀a,b:E. a ≰ b → b < a.
 
+(* E(#,≰) → E(#,sym(≰)) *)
 lemma dual_exc: excess→ excess.
-intro E; apply mk_excess; 
-[apply (exc_carr E);|apply (λx,y:E.y≰x);|apply exc_coreflexive;
-|unfold cotransitive; simplify;intros;cases (exc_cotransitive E ??z e);
-  [right|left]assumption]
+intro E; apply mk_excess;
+[1: apply mk_excess_;
+  [1: apply (mk_excess_base (exc_carr (excess_carr E)));
+      [ apply (λx,y:E.y≰x);|apply exc_coreflexive;
+      | unfold cotransitive; simplify; intros (x y z H);
+        cases (exc_cotransitive E ??z H);[right|left]assumption]
+  |2: apply (exc_ap E);
+  |3: apply (exc_with E);]
+|2: simplify; intros (y x H); fold simplify (y#x) in H;
+    apply ap2exc; apply ap_symmetric; apply H;
+|3: simplify; intros; fold simplify (y#x); apply exc2ap; 
+    cases o; [right|left]assumption]
 qed.
index 1f605c257e2e13820be92d875ab5249da7cb09a0..83d138526ae9fea77477599aa4d3ab7b6e5bce5e 100644 (file)
 
 include "excess.ma".
 
-record directed : Type ≝ {
-  dir_carr: apartness;
-  dir_op: dir_carr → dir_carr → dir_carr;
-  dir_op_refl: ∀x.dir_op x x ≈ x;  
-  dir_op_comm: ∀x,y:dir_carr. dir_op x y ≈ dir_op y x;
-  dir_op_assoc: ∀x,y,z:dir_carr. dir_op x (dir_op y z) ≈ dir_op (dir_op x y) z;
-  dir_strong_extop: ∀x.strong_ext ? (dir_op x)  
+record semi_lattice_base : Type ≝ {
+  sl_carr:> apartness;
+  sl_op: sl_carr → sl_carr → sl_carr;
+  sl_op_refl: ∀x.sl_op x x ≈ x;  
+  sl_op_comm: ∀x,y:sl_carr. sl_op x y ≈ sl_op y x;
+  sl_op_assoc: ∀x,y,z:sl_carr. sl_op x (sl_op y z) ≈ sl_op (sl_op x y) z;
+  sl_strong_extop: ∀x.strong_ext ? (sl_op x)  
 }.
 
-definition excl ≝ 
-  λl:directed.λa,b:dir_carr l.ap_apart (dir_carr l) a (dir_op l a b).
-
-lemma excess_of_directed: directed → excess.
-intro l; apply (mk_excess (dir_carr l) (excl l));
-[ intro x; unfold; intro H; unfold in H; apply (ap_coreflexive (dir_carr l) x);
-  apply (ap_rewr ??? (dir_op l x x) (dir_op_refl ? x)); assumption;
-| intros 3 (x y z); unfold excl; intro H;
-  cases (ap_cotransitive ??? (dir_op l (dir_op l x z) y) H) (H1 H2); [2:
-    left; apply ap_symmetric; apply (dir_strong_extop ? y); 
-    apply (ap_rewl ???? (dir_op_comm ???));
-    apply (ap_rewr ???? (dir_op_comm ???));
-    assumption]
-  cases (ap_cotransitive ??? (dir_op l x z) H1) (H2 H3); [left; assumption]
-  right; apply (dir_strong_extop ? x); apply (ap_rewr ???? (dir_op_assoc ????));
-  assumption]
+notation "a \cross b" left associative with precedence 50 for @{ 'op $a $b }.
+interpretation "semi lattice base operation" 'op a b = (cic:/matita/lattice/sl_op.con _ a b).
+
+lemma excess_of_semi_lattice_base: semi_lattice_base → excess.
+intro l;
+apply mk_excess;
+[1: apply mk_excess_;
+    [1: 
+    
+  apply (mk_excess_base (sl_carr l));
+    [1: apply (λa,b:sl_carr l.a # (a ✗ b));
+    |2: unfold; intros 2 (x H); simplify in H;
+        lapply (Ap≪ ? (sl_op_refl ??) H) as H1; clear H;
+        apply (ap_coreflexive ?? H1);
+    |3: unfold; simplify; intros (x y z H1);
+        cases (ap_cotransitive ??? ((x ✗ z) ✗ y) H1) (H2 H2);[2:
+          lapply (Ap≪ ? (sl_op_comm ???) H2) as H21;
+          lapply (Ap≫ ? (sl_op_comm ???) H21) as H22; clear H21 H2;
+          lapply (sl_strong_extop ???? H22); clear H22; 
+          left; apply ap_symmetric; assumption;]
+        cases (ap_cotransitive ??? (x ✗ z) H2) (H3 H3);[left;assumption]
+        right; lapply (Ap≫ ? (sl_op_assoc ????) H3) as H31;
+        apply (sl_strong_extop ???? H31);]
+
+    |2:
+    apply apartness_of_excess_base; 
+    
+  apply (mk_excess_base (sl_carr l));
+    [1: apply (λa,b:sl_carr l.a # (a ✗ b));
+    |2: unfold; intros 2 (x H); simplify in H;
+        lapply (Ap≪ ? (sl_op_refl ??) H) as H1; clear H;
+        apply (ap_coreflexive ?? H1);
+    |3: unfold; simplify; intros (x y z H1);
+        cases (ap_cotransitive ??? ((x ✗ z) ✗ y) H1) (H2 H2);[2:
+          lapply (Ap≪ ? (sl_op_comm ???) H2) as H21;
+          lapply (Ap≫ ? (sl_op_comm ???) H21) as H22; clear H21 H2;
+          lapply (sl_strong_extop ???? H22); clear H22; 
+          left; apply ap_symmetric; assumption;]
+        cases (ap_cotransitive ??? (x ✗ z) H2) (H3 H3);[left;assumption]
+        right; lapply (Ap≫ ? (sl_op_assoc ????) H3) as H31;
+        apply (sl_strong_extop ???? H31);]
+    
+    |3: apply refl_eq;]
+|2,3: intros (x y H); assumption;]         
 qed.    
 
-record prelattice : Type ≝ {
-  pl_carr:> excess;
-  meet: pl_carr → pl_carr → pl_carr;
+record semi_lattice : Type ≝ {
+  sl_exc:> excess;
+  meet: sl_exc → sl_exc → sl_exc;
   meet_refl: ∀x.meet x x ≈ x;  
-  meet_comm: ∀x,y:pl_carr. meet x y ≈ meet y x;
-  meet_assoc: ∀x,y,z:pl_carr. meet x (meet y z) ≈ meet (meet x y) z;
+  meet_comm: ∀x,y. meet x y ≈ meet y x;
+  meet_assoc: ∀x,y,z. meet x (meet y z) ≈ meet (meet x y) z;
   strong_extm: ∀x.strong_ext ? (meet x);
   le_to_eqm: ∀x,y.x ≤ y → x ≈ meet x y;
   lem: ∀x,y.(meet x y) ≤ y 
 }.
  
-interpretation "prelattice meet" 'and a b =
-  (cic:/matita/lattice/meet.con _ a b).
+interpretation "semi lattice meet" 'and a b = (cic:/matita/lattice/meet.con _ a b).
 
-lemma feq_ml: ∀ml:prelattice.∀a,b,c:ml. a ≈ b → (c ∧ a) ≈ (c ∧ b).
+lemma feq_ml: ∀ml:semi_lattice.∀a,b,c:ml. a ≈ b → (c ∧ a) ≈ (c ∧ b).
 intros (l a b c H); unfold eq in H ⊢ %; unfold Not in H ⊢ %;
 intro H1; apply H; clear H; apply (strong_extm ???? H1);
 qed.
 
-lemma feq_mr: ∀ml:prelattice.∀a,b,c:ml. a ≈ b → (a ∧ c) ≈ (b ∧ c).
+lemma feq_mr: ∀ml:semi_lattice.∀a,b,c:ml. a ≈ b → (a ∧ c) ≈ (b ∧ c).
 intros (l a b c H); 
 apply (Eq≈ ? (meet_comm ???)); apply (Eq≈ ?? (meet_comm ???));
 apply feq_ml; assumption;
 qed.
  
-lemma prelattice_of_directed: directed → prelattice.
-intro l_; apply (mk_prelattice (excess_of_directed l_)); [apply (dir_op l_);]
+(*
+lemma semi_lattice_of_semi_lattice_base: semi_lattice_base → semi_lattice.
+intro slb; apply (mk_semi_lattice (excess_of_semi_lattice_base slb));
+[1: apply (sl_op slb);
+|2: intro x; apply (eq_trans (excess_of_semi_lattice_base slb)); [2: 
+      apply (sl_op_refl slb);|1:skip] (sl_op slb x x)); ? (sl_op_refl slb x));
+
+ unfold excess_of_semi_lattice_base; simplify;
+    intro H; elim H;
+    [ 
+    
+    
+    lapply (ap_rewl (excess_of_semi_lattice_base slb) x ? (sl_op slb x x) 
+      (eq_sym (excess_of_semi_lattice_base slb) ?? (sl_op_refl slb x)) t);
+    change in x with (sl_carr slb);
+    apply (Ap≪ (x ✗ x)); (sl_op_refl slb x)); 
+
+whd in H; elim H; clear H;
+    [ apply (ap_coreflexive (excess_of_semi_lattice_base slb) (x ✗ x) t);
+
+prelattice (excess_of_directed l_)); [apply (sl_op l_);]
 unfold excess_of_directed; try unfold apart_of_excess; simplify;
 unfold excl; simplify;
 [intro x; intro H; elim H; clear H; 
- [apply (dir_op_refl l_ x); 
-  lapply (Ap≫ ? (dir_op_comm ???) t) as H; clear t; 
-  lapply (dir_strong_extop l_ ??? H); apply ap_symmetric; assumption
- | lapply (Ap≪ ? (dir_op_refl ?x) t) as H; clear t;
-   lapply (dir_strong_extop l_ ??? H); apply (dir_op_refl l_ x);
+ [apply (sl_op_refl l_ x); 
+  lapply (Ap≫ ? (sl_op_comm ???) t) as H; clear t; 
+  lapply (sl_strong_extop l_ ??? H); apply ap_symmetric; assumption
+ | lapply (Ap≪ ? (sl_op_refl ?x) t) as H; clear t;
+   lapply (sl_strong_extop l_ ??? H); apply (sl_op_refl l_ x);
    apply ap_symmetric; assumption]
 |intros 3 (x y H); cases H (H1 H2); clear H;
- [lapply (Ap≪ ? (dir_op_refl ? (dir_op l_ x y)) H1) as H; clear H1;
-  lapply (dir_strong_extop l_ ??? H) as H1; clear H;
-  lapply (Ap≪ ? (dir_op_comm ???) H1); apply (ap_coreflexive ?? Hletin);
- |lapply (Ap≪ ? (dir_op_refl ? (dir_op l_ y x)) H2) as H; clear H2;
-  lapply (dir_strong_extop l_ ??? H) as H1; clear H;
-  lapply (Ap≪ ? (dir_op_comm ???) H1);apply (ap_coreflexive ?? Hletin);]
+ [lapply (Ap≪ ? (sl_op_refl ? (sl_op l_ x y)) H1) as H; clear H1;
+  lapply (sl_strong_extop l_ ??? H) as H1; clear H;
+  lapply (Ap≪ ? (sl_op_comm ???) H1); apply (ap_coreflexive ?? Hletin);
+ |lapply (Ap≪ ? (sl_op_refl ? (sl_op l_ y x)) H2) as H; clear H2;
+  lapply (sl_strong_extop l_ ??? H) as H1; clear H;
+  lapply (Ap≪ ? (sl_op_comm ???) H1);apply (ap_coreflexive ?? Hletin);]
 |intros 4 (x y z H); cases H (H1 H2); clear H;
- [lapply (Ap≪ ? (dir_op_refl ? (dir_op l_ x (dir_op l_ y z))) H1) as H; clear H1;
-  lapply (dir_strong_extop l_ ??? H) as H1; clear H;
-  lapply (Ap≪ ? (eq_sym ??? (dir_op_assoc ?x y z)) H1) as H; clear H1;
+ [lapply (Ap≪ ? (sl_op_refl ? (sl_op l_ x (sl_op l_ y z))) H1) as H; clear H1;
+  lapply (sl_strong_extop l_ ??? H) as H1; clear H;
+  lapply (Ap≪ ? (eq_sym ??? (sl_op_assoc ?x y z)) H1) as H; clear H1;
   apply (ap_coreflexive ?? H);
- |lapply (Ap≪ ? (dir_op_refl ? (dir_op l_ (dir_op l_ x y) z)) H2) as H; clear H2;
-  lapply (dir_strong_extop l_ ??? H) as H1; clear H;
-  lapply (Ap≪ ? (dir_op_assoc ?x y z) H1) as H; clear H1;
+ |lapply (Ap≪ ? (sl_op_refl ? (sl_op l_ (sl_op l_ x y) z)) H2) as H; clear H2;
+  lapply (sl_strong_extop l_ ??? H) as H1; clear H;
+  lapply (Ap≪ ? (sl_op_assoc ?x y z) H1) as H; clear H1;
   apply (ap_coreflexive ?? H);]
 |intros (x y z H); elim H (H1 H1); clear H;
- lapply (Ap≪ ? (dir_op_refl ??) H1) as H; clear H1;
- lapply (dir_strong_extop l_ ??? H) as H1; clear H;
- lapply (dir_strong_extop l_ ??? H1) as H; clear H1;
- cases (ap_cotransitive ??? (dir_op l_ y z) H);[left|right|right|left] try assumption;
- [apply ap_symmetric;apply (Ap≪ ? (dir_op_comm ???));
- |apply (Ap≫ ? (dir_op_comm ???));
+ lapply (Ap≪ ? (sl_op_refl ??) H1) as H; clear H1;
+ lapply (sl_strong_extop l_ ??? H) as H1; clear H;
+ lapply (sl_strong_extop l_ ??? H1) as H; clear H1;
+ cases (ap_cotransitive ??? (sl_op l_ y z) H);[left|right|right|left] try assumption;
+ [apply ap_symmetric;apply (Ap≪ ? (sl_op_comm ???));
+ |apply (Ap≫ ? (sl_op_comm ???));
  |apply ap_symmetric;] assumption;
 |intros 4 (x y H H1); apply H; clear H; elim H1 (H H);
- lapply (Ap≪ ? (dir_op_refl ??) H) as H1; clear H;
- lapply (dir_strong_extop l_ ??? H1) as H; clear H1;[2: apply ap_symmetric]
+ lapply (Ap≪ ? (sl_op_refl ??) H) as H1; clear H;
+ lapply (sl_strong_extop l_ ??? H1) as H; clear H1;[2: apply ap_symmetric]
  assumption
 |intros 3 (x y H); 
- cut (dir_op l_ x y ≈ dir_op l_ x (dir_op l_ y y)) as H1;[2:
-   intro; lapply (dir_strong_extop ???? a); apply (dir_op_refl l_ y);
+ cut (sl_op l_ x y ≈ sl_op l_ x (sl_op l_ y y)) as H1;[2:
+   intro; lapply (sl_strong_extop ???? a); apply (sl_op_refl l_ y);
    apply ap_symmetric; assumption;]
- lapply (Ap≪ ? (eq_sym ??? H1) H); apply (dir_op_assoc l_ x y y);
+ lapply (Ap≪ ? (eq_sym ??? H1) H); apply (sl_op_assoc l_ x y y);
  assumption; ]
 qed.
+*)
+
 
 record lattice_ : Type ≝ {
-  latt_mcarr:> prelattice;
-  latt_jcarr_: prelattice;
-  latt_with: pl_carr latt_jcarr_ = dual_exc (pl_carr latt_mcarr)
+  latt_mcarr:> semi_lattice;
+  latt_jcarr_: semi_lattice;
+  latt_with:  sl_exc latt_jcarr_ = dual_exc (sl_exc latt_mcarr)
 }.
 
-lemma latt_jcarr : lattice_ → prelattice.
+lemma latt_jcarr : lattice_ → semi_lattice.
 intro l;
-apply (mk_prelattice (dual_exc l)); unfold excess_OF_lattice_;
+apply (mk_semi_lattice (dual_exc l)); 
+unfold excess_OF_lattice_;
 cases (latt_with l); simplify;
 [apply meet|apply meet_refl|apply meet_comm|apply meet_assoc|
 apply strong_extm| apply le_to_eqm|apply lem]