1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
15 set "baseuri" "cic:/matita/lebesgue/".
20 axiom ltT : ∀R:real.∀a,b:R.Type.
23 alias symbol "or" = "constructive or".
24 definition rapart ≝ λR:real.λa,b:R.a < b ∨ b < a.
26 notation "a # b" non associative with precedence 50 for
28 notation "a \approx b" non associative with precedence 50 for
30 interpretation "real apartness" 'apart a b =
31 (cic:/matita/lebesgue/rapart.con _ a b).
32 interpretation "real non apartness" 'napart a b =
33 (cic:/matita/constructive_connectives/Not.con
34 (cic:/matita/lebesgue/rapart.con _ a b)).
36 record is_semi_metric (R : real) (T : Type) (d : T → T → R): Prop ≝ {
37 sm_positive: ∀a,b:T. d a b ≤ 0;
38 sm_reflexive: ∀a. d a a ≈ 0;
39 sm_symmetric: ∀a,b. d a b ≈ d b a;
40 sm_tri_ineq: ∀a,b,c:T. d a b ≤ d a c + d c b
43 record semimetric_set (R : real) : Type ≝ {
45 ms_semimetric: ms_carr → ms_carr → R;
46 ms_properties:> is_semi_metric R ms_carr ms_semimetric
49 notation < "\nbsp \delta a \nbsp b" non associative with precedence 80 for @{ 'delta2 $a $b}.
50 interpretation "semimetric" 'delta2 a b = (cic:/matita/lebesgue/ms_semimetric.con _ _ a b).
51 notation "\delta" non associative with precedence 80 for @{ 'delta }.
52 interpretation "semimetric" 'delta = (cic:/matita/lebesgue/ms_semimetric.con _ _).
54 record pre_semimetric_lattice (R : real) : Type ≝ {
55 ml_lattice :> lattice;
56 ml_semimetric_set_ : semimetric_set R;
57 with_ml_lattice_eq_ml_semimetric_set_: ms_carr ? ml_semimetric_set_ = ml_lattice
60 lemma ml_semimetric_set : ∀R.∀ms:pre_semimetric_lattice R. semimetric_set R.
61 intros (R ml); constructor 1; [apply (ml : Type);]
62 cases ml (l ms with_); clear ml; simplify;
63 [1: rewrite < with_; apply (ms_semimetric ? ms);
64 |2: cases with_; simplify; apply (ms_properties ? ms);]
67 coercion cic:/matita/lebesgue/ml_semimetric_set.con.
69 record is_semimetric_lattice (R : real) (ml : pre_semimetric_lattice R) : Prop ≝ {
70 prop1a: ∀a : ml.δ (a ∧ a) a ≈ 0;
71 prop1b: ∀a : ml.δ (a ∨ a) a ≈ 0;
72 prop2a: ∀a,b: ml. δ (a ∨ b) (b ∨ a) ≈ 0;
73 prop2b: ∀a,b: ml. δ (a ∧ b) (b ∧ a) ≈ 0;
74 prop3a: ∀a,b,c: ml. δ (a ∨ (b ∨ c)) ((a ∨ b) ∨ c) ≈ 0;
75 prop3b: ∀a,b,c: ml. δ (a ∧ (b ∧ c)) ((a ∧ b) ∧ c) ≈ 0;
76 prop4a: ∀a,b: ml. δ (a ∨ (a ∧ b)) a ≈ 0;
77 prop4b: ∀a,b: ml. δ (a ∧ (a ∨ b)) a ≈ 0;
78 prop5: ∀a,b,c: ml. δ (a ∨ b) (a ∨ c) + δ (a ∧ b) (a ∧ c) ≤ δ b c
81 record semimetric_lattice (R : real) : Type ≝ {
82 ml_pre_semimetric_lattice:> pre_semimetric_lattice R;
83 ml_semimetric_lattice_properties: is_semimetric_lattice R ml_pre_semimetric_lattice
87 λR: real. λml: semimetric_lattice R. λa,b: ml.
88 (* 0 < δ a b *) ltT ? 0 (δ a b).
89 (* < scazzato, ma CSC dice che poi si cambia dopo *)
91 interpretation "semimetric lattice apartness" 'apart a b =
92 (cic:/matita/lebesgue/apart.con _ _ a b).
93 interpretation "semimetric lattice non apartness" 'napart a b =
94 (cic:/matita/constructive_connectives/Not.con
95 (cic:/matita/lebesgue/apart.con _ _ a b)).
98 lemma triangular_inequality : ∀R: real. ∀ms: semimetric_set R.∀a,b,c:ms.
99 δ a b ≤ δ a c + δ c b.
100 intros (R ms a b c); exact (sm_tri_ineq R ms (ms_semimetric R ms) ms a b c);
103 lemma foo : ∀R: real. ∀ml: semimetric_lattice R.∀a,b,a1,b1: ml.
104 a ≈ a1 → b ≈ b1 → (δ a b ≈ δ a1 b1).
106 lapply (triangular_inequality ? ? a b a1) as H1;
107 lapply (triangular_inequality ? ? a1 b b1) as H2;
108 lapply (og_ordered_abelian_group_properties R ? ? (δ a a1) H2) as H3;