]> matita.cs.unibo.it Git - helm.git/commitdiff
the is_structure is a bad idea if you don't have canonical structures.
authorEnrico Tassi <enrico.tassi@inria.fr>
Sat, 8 Dec 2007 10:36:00 +0000 (10:36 +0000)
committerEnrico Tassi <enrico.tassi@inria.fr>
Sat, 8 Dec 2007 10:36:00 +0000 (10:36 +0000)
helm/software/matita/dama/metric_lattice.ma

index b5261b9dc44698b021b875f288180b741c614fb4..0bfc3db678179c3ceedaab49b34ff762961f6f75 100644 (file)
@@ -33,24 +33,20 @@ qed.
 
 coercion cic:/matita/metric_lattice/ml_mspace.con.
 
-record is_mlattice (R : todgroup) (ml: mlattice_ R) : Type ≝ {
-  ml_prop1: ∀a,b:ml. 0 < δ a b → a # b;
-  ml_prop2: ∀a,b,c:ml. δ (a∨b) (a∨c) + δ (a∧b) (a∧c) ≤ δ b c
-}.
-
 record mlattice (R : todgroup) : Type ≝ {
   ml_carr :> mlattice_ R;
-  ml_props:> is_mlattice R ml_carr
+  ml_prop1: ∀a,b:ml_carr. 0 < δ a b → a # b;
+  ml_prop2: ∀a,b,c:ml_carr. δ (a∨b) (a∨c) + δ (a∧b) (a∧c) ≤ δ b c
 }.
 
 lemma eq_to_ndlt0: ∀R.∀ml:mlattice R.∀a,b:ml. a ≈ b → ¬ 0 < δ a b.
-intros (R ml a b E); intro H; apply E; apply (ml_prop1 ?? ml);
+intros (R ml a b E); intro H; apply E; apply ml_prop1;
 assumption;
 qed.
 
 lemma eq_to_dzero: ∀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;
+apply ml_prop1; split [apply mpositive] apply ap_symmetric;
 assumption;
 qed.
 
@@ -76,7 +72,7 @@ intros; split [apply mpositive] apply ap_symmetric; assumption;
 qed.
 
 lemma dap_to_ap: ∀R.∀ml:mlattice R.∀x,y:ml. δ x y # 0 → x # y.
-intros (R ml x y H); apply (ml_prop1 ?? ml); split; [apply mpositive;]
+intros (R ml x y H); apply ml_prop1; split; [apply mpositive;]
 apply ap_symmetric; assumption;
 qed.
 
@@ -84,7 +80,7 @@ qed.
 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) ??)); 
+apply (le_transitive ????? (ml_prop2 ?? (y) ??)); 
 cut ( δx y+ δy z ≈ δ(y∨x) (y∨z)+ δ(y∧x) (y∧z)); [
   apply (le_rewr ??? (δx y+ δy z)); [assumption] apply le_reflexive]
 lapply (le_to_eqm ??? Lxy) as Dxm; lapply (le_to_eqm ??? Lyz) as Dym;