]> matita.cs.unibo.it Git - helm.git/blob - matita/dama/constructive_pointfree/lebesgue.ma
xxx
[helm.git] / matita / dama / constructive_pointfree / lebesgue.ma
1 (**************************************************************************)
2 (*       ___                                                              *)
3 (*      ||M||                                                             *)
4 (*      ||A||       A project by Andrea Asperti                           *)
5 (*      ||T||                                                             *)
6 (*      ||I||       Developers:                                           *)
7 (*      ||T||         The HELM team.                                      *)
8 (*      ||A||         http://helm.cs.unibo.it                             *)
9 (*      \   /                                                             *)
10 (*       \ /        This file is distributed under the terms of the       *)
11 (*        v         GNU General Public License Version 2                  *)
12 (*                                                                        *)
13 (**************************************************************************)
14
15 set "baseuri" "cic:/matita/lebesgue/".
16
17 include "reals.ma".
18 include "lattices.ma".
19
20 record measured_set (R : real) : Type ≝ {
21   ms_carr :> Type;
22   ms_measure: ms_carr → ms_carr → R
23 }.
24
25 notation "\delta" non associative with precedence 90 for @{ 'delta }.
26 interpretation "measure" 'delta = (cic:/matita/lebesgue/ms_measure.con _ _).
27
28 record pre_measured_lattice (R : real) : Type ≝ {
29   ml_lattice :> lattice;
30   ml_measured_set_ : measured_set R;
31   with_ml_lattice_eq_ml_measured_set_: ms_carr ? ml_measured_set_ = ml_lattice
32 }.
33
34 lemma ml_measured_set : ∀R.∀ms:pre_measured_lattice R. measured_set R.
35 intros (R ml); constructor 1; [1:apply (ml : Type);] cases ml; 
36 rewrite < H; clear H; cases ml_measured_set_; simplify; exact f;
37 qed.  
38
39 coercion cic:/matita/lebesgue/ml_measured_set.con.
40
41 record is_measured_lattice (R : real) (ml : pre_measured_lattice R) : Prop ≝ {
42   prop1a: ∀a : ml.δ (a ∧ a) a = 0;
43   prop1b: ∀a : ml.δ (a ∨ a) a = 0;
44   prop2a: ∀a,b: ml. δ (a ∨ b) (b ∨ a) = 0;
45   prop2b: ∀a,b: ml. δ (a ∧ b) (b ∧ a) = 0;
46   prop3a: ∀a,b,c: ml. δ (a ∨ (b ∨ c)) ((a ∨ b) ∨ c) = 0;
47   prop3b: ∀a,b,c: ml. δ (a ∧ (b ∧ c)) ((a ∧ b) ∧ c) = 0;
48   prop4a: ∀a,b: ml. δ (a ∨ (a ∧ b)) a = 0;
49   prop4b: ∀a,b: ml. δ (a ∧ (a ∨ b)) a = 0;
50   prop5: ∀a,b,c: ml. δ (a ∨ b) (a ∨ c) + δ (a ∧ b) (a ∧ c) ≤ δ b c
51 }. 
52   
53 record measured_lattice (R : real) : Type ≝ {
54   ml_pre_measured_lattice:> pre_measured_lattice R;
55   ml_measured_lattice_properties: is_measured_lattice R ml_pre_measured_lattice
56 }.
57   
58 definition apart:= 
59  λR: real. λml: measured_lattice R. λa,b: ml. 0 < δ a b.
60  (* < scazzato, ma CSC dice che poi si cambia dopo  *)    
61
62 notation "a # b" non associative with precedence 50 for
63  @{ 'apart $a $b}.
64 interpretation "measured lattice apartness" 'apart a b =
65  (cic:/matita/lebesgue/apart.con _ _ a b). 
66 notation "a \approx b" non associative with precedence 50 for
67  @{ 'napart $a $b}.
68 interpretation "measured lattice non apartness" 'napart a b =
69  (cic:/matita/logic/connectives/Not.con
70    (cic:/matita/lebesgue/apart.con _ _ a b)). 
71
72 lemma foo : ∀R: real. ∀ml: measured_lattice R.∀a,b,a1,b1: ml. 
73   a ≈ a1 → b ≈ b1 → δ a b = δ a1 b1. 
74      (* =R scazzato *)  
75 intros;