X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;ds=sidebyside;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fdama%2Fdama%2Fsupremum.ma;h=a4df778067905ea37101fd6cde92571158343462;hb=9156537d75378d7a254e93b5a2d036d687cd79ee;hp=b2411f65263a318cee22f7529a0dec08baa59a60;hpb=bb6f711dd0c8378d27118b73981515aff05f5750;p=helm.git diff --git a/helm/software/matita/contribs/dama/dama/supremum.ma b/helm/software/matita/contribs/dama/dama/supremum.ma index b2411f652..a4df77806 100644 --- a/helm/software/matita/contribs/dama/dama/supremum.ma +++ b/helm/software/matita/contribs/dama/dama/supremum.ma @@ -122,6 +122,17 @@ include "nat_ordered_set.ma". alias symbol "nleq" = "Ordered set excess". alias symbol "leq" = "Ordered set less or equal than". lemma trans_increasing: + ∀C:ordered_set.∀a:sequence C.a is_increasing → ∀n,m:nat_ordered_set. n ≤ m → a n ≤ a m. +intros 5 (C a Hs n m); elim m; [ + rewrite > (le_n_O_to_eq n (not_lt_to_le O n H)); + intro X; cases (os_coreflexive ?? X);] +cases (le_to_or_lt_eq ?? (not_lt_to_le (S n1) n H1)); clear H1; +[2: rewrite > H2; intro; cases (os_coreflexive ?? H1); +|1: apply (le_transitive ???? (H ?) (Hs ?)); + intro; whd in H1; apply (not_le_Sn_n n); apply (transitive_le ??? H2 H1);] +qed. + +lemma trans_increasing_exc: ∀C:ordered_set.∀a:sequence C.a is_increasing → ∀n,m:nat_ordered_set. m ≰ n → a n ≤ a m. intros 5 (C a Hs n m); elim m; [cases (not_le_Sn_O n H);] intro; apply H; @@ -153,12 +164,12 @@ lemma selection: ∀C:ordered_set.∀m:sequence nat_ordered_set.m is_strictly_increasing → ∀a:sequence C.∀u.a ↑ u → (λx.a (m x)) ↑ u. intros (C m Hm a u Ha); cases Ha (Ia Su); cases Su (Uu Hu); repeat split; -[1: intro n; simplify; apply trans_increasing; [assumption] apply (Hm n); +[1: intro n; simplify; apply trans_increasing_exc; [assumption] apply (Hm n); |2: intro n; simplify; apply Uu; |3: intros (y Hy); simplify; cases (Hu ? Hy); cases (strictly_increasing_reaches C ? Hm w); exists [apply w1]; cases (os_cotransitive ??? (a (m w1)) H); [2:assumption] - cases (trans_increasing C ? Ia ?? H1); assumption;] + cases (trans_increasing_exc C ? Ia ?? H1); assumption;] qed. (* Definition 2.7 *)