]> matita.cs.unibo.it Git - helm.git/commitdiff
more stuff
authorEnrico Tassi <enrico.tassi@inria.fr>
Wed, 28 Nov 2007 18:12:33 +0000 (18:12 +0000)
committerEnrico Tassi <enrico.tassi@inria.fr>
Wed, 28 Nov 2007 18:12:33 +0000 (18:12 +0000)
helm/software/matita/dama/metric_lattice.ma

index 56120525d8cf3d794f6889000b7cbc1e12f8a995..6028ada59748c3659533fb48b5fd7c02ec8b694a 100644 (file)
@@ -43,23 +43,31 @@ record mlattice (R : ogroup) : Type ≝ {
   ml_props:> is_mlattice R ml_carr
 }.
 
-axiom meq_joinl: ∀R.∀ml:mlattice R.∀x,y,z:ml. x≈z → δx y ≈ δz y.
+lemma eq_to_zero: ∀R.∀ml:mlattice R.∀x,y:ml.x ≈ y → δ x y ≈ 0.
+intros (R ml x y H); intro H1; apply H; clear H; 
+apply (ml_prop1 ?? ml); split [apply mpositive] apply ap_symmetric;
+assumption;
+qed.
+
+lemma meq_joinl: ∀R.∀ml:mlattice R.∀x,y,z:ml. x≈z → δx y ≈ δz y.
+intros (R ml x y z); apply le_le_eq;
+[ apply (le_transitive ???? (mtineq ???y z)); 
+  apply (le_rewl ??? (0+δz y) (eq_to_zero ???? H));
+  apply (le_rewl ??? (δz y) (zero_neutral ??)); apply le_reflexive;
+| apply (le_transitive ???? (mtineq ???y x));
+  apply (le_rewl ??? (0+δx y) (eq_to_zero ??z x H));
+  apply (le_rewl ??? (δx y) (zero_neutral ??)); apply le_reflexive;]
+qed.
 
 lemma meq_joinr: ∀R.∀ml:mlattice R.∀x,y,z:ml. x≈z → δy x ≈ δy z.
 intros; apply (eq_trans ???? (msymmetric ??y x));
 apply (eq_trans ????? (msymmetric ??z y)); apply meq_joinl; assumption;
-qed.
-
-(*
-lemma foo: ∀R.∀ml:mlattice R.∀x,y,z:ml. δx y ≈ δ(y∨x) (y∨z).
-intros; apply le_le_eq; [
-  apply (le_rewr ???? (meq_joinl ????? (join_comm ???)));
-  apply (le_rewr ???? (meq_joinr ????? (join_comm ???)));
-*) 
+qed. 
 
 (* 3.11 *)
 lemma le_mtri: 
   ∀R.∀ml:mlattice R.∀x,y,z:ml. x ≤ y → y ≤ z → δ x z ≈ δ x y + δ y z.
 intros (R ml x y z Lxy Lyz); apply le_le_eq; [apply mtineq]
-apply (le_transitive ????? (ml_prop2 ?? ml (y∧x) ??)); 
+apply (le_transitive ????? (ml_prop2 ?? ml (y) ??)); 
 (* auto type. assert failure *)
+whd;