From: Enrico Tassi Date: Fri, 30 Nov 2007 17:57:32 +0000 (+0000) Subject: carabinieri almost done X-Git-Tag: make_still_working~5747 X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=commitdiff_plain;h=f86ba51fa84583b5963fca03687cd23b838a99f5;p=helm.git carabinieri almost done --- diff --git a/helm/software/matita/dama/excedence.ma b/helm/software/matita/dama/excedence.ma index 42d53a5dc..ab7a53af0 100644 --- a/helm/software/matita/dama/excedence.ma +++ b/helm/software/matita/dama/excedence.ma @@ -199,6 +199,14 @@ whd in LE LEx APx; cases APx (EXx EXx); [cases (LEx EXx)] cases (exc_cotransitive ??? z EXx) (EXz EXz); [cases (LE EXz)] right; assumption; qed. + +lemma le_lt_transitive: ∀A:excedence.∀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] +cases LE; assumption; +qed. + lemma le_le_eq: ∀E:excedence.∀a,b:E. a ≤ b → b ≤ a → a ≈ b. intros (E x y L1 L2); intro H; cases H; [apply L1|apply L2] assumption; diff --git a/helm/software/matita/dama/metric_lattice.ma b/helm/software/matita/dama/metric_lattice.ma index a4ca3b4f4..0e92f0a89 100644 --- a/helm/software/matita/dama/metric_lattice.ma +++ b/helm/software/matita/dama/metric_lattice.ma @@ -100,6 +100,67 @@ apply (Eq≈ (δ(y∧x) (y∧z))); [apply (meq_l ????? (meet_comm ???));] apply eq_reflexive; qed. +include "sequence.ma". +include "nat/plus.ma". + +definition d2s ≝ + λR.λml:mlattice R.λs:sequence (pordered_set_of_excedence ml).λk.λn. δ (s n) k. +(* +notation "s ⇝ 'Zero'" non associative with precedence 50 for @{'tends0 $s }. + +interpretation "tends to" 'tends s x = + (cic:/matita/sequence/tends0.con _ s). +*) + +alias symbol "leq" = "ordered sets less or equal than". +alias symbol "and" = "constructive and". +alias symbol "exists" = "constructive exists (Type)". +lemma carabinieri: (* non trova la coercion .... *) + ∀R.∀ml:mlattice R.∀an,bn,xn:sequence (pordered_set_of_excedence ml). + (∀n. (an n ≤ xn n) ∧ (xn n ≤ bn n)) → + ∀x:ml. tends0 ? (d2s ? ml an x) → tends0 ? (d2s ? ml bn x) → + tends0 ? (d2s ? ml xn x). +intros (R ml an bn xn H x Ha Hb); unfold tends0 in Ha Hb ⊢ %. unfold d2s in Ha Hb ⊢ %. +intros (e He); +elim (Ha ? He) (n1 H1); clear Ha; elim (Hb e He) (n2 H2); clear Hb; +constructor 1; [apply (n1 + n2);] intros (n3 Hn3); +cut (n1 sym_plus in H3; assumption] + apply H3; rewrite > sym_plus in H4; simplify in H4; apply lt_S_to_lt; + rewrite > sym_plus in H4; assumption;] +elim (H1 ? Hcut) (H3 H4); clear Hcut; +cut (n2