X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fdama%2Fdama%2Fmodels%2Fnat_order_continuous.ma;h=0b4c92e3e9f67fe5bd8f04ea629ba50f43f19600;hb=f116dde7daadaffc015e06b3c38a57398bfe0ef4;hp=7949ba2d28714d077a1f5976b58af8b73f351276;hpb=7deb4b1f322850b8ff03d5626f7828736d074ec8;p=helm.git diff --git a/helm/software/matita/contribs/dama/dama/models/nat_order_continuous.ma b/helm/software/matita/contribs/dama/dama/models/nat_order_continuous.ma index 7949ba2d2..0b4c92e3e 100644 --- a/helm/software/matita/contribs/dama/dama/models/nat_order_continuous.ma +++ b/helm/software/matita/contribs/dama/dama/models/nat_order_continuous.ma @@ -15,20 +15,20 @@ include "models/nat_dedekind_sigma_complete.ma". include "models/nat_ordered_uniform.ma". -lemma nat_us_is_oc: ∀l,u:ℕ.order_continuity (segment_ordered_uniform_space ℕ l u). -intros 4; split; intros 3; cases H1; cases H2; clear H2 H1; cases H; clear H; -[1: cases (nat_dedekind_sigma_complete l u a H1 ? H2); +lemma nat_us_is_oc: ∀s:‡ℕ.order_continuity (segment_ordered_uniform_space ℕ s). +intros 3; split; intros 3; cases H1; cases H2; clear H2 H1; cases H; clear H; +[1: cases (nat_dedekind_sigma_complete s a H1 ? H2); exists [apply w1]; intros; - apply (restrict nat_ordered_uniform_space l u ??? H4); + apply (restrict nat_ordered_uniform_space s ??? H4); generalize in match (H ? H5); cases x; intro; generalize in match (refl_eq ? (a i):a i = a i); generalize in match (a i) in ⊢ (?? % ? → %); intro X; cases X; clear X; intro; rewrite < H9 in H7; simplify; rewrite < H7; simplify; apply (us_phi1 nat_uniform_space ? H3); simplify; apply (eq_reflexive (us_carr nat_uniform_space)); -|2: cases (nat_dedekind_sigma_complete_r l u a H1 ? H2); +|2: cases (nat_dedekind_sigma_complete_r s a H1 ? H2); exists [apply w1]; intros; - apply (restrict nat_ordered_uniform_space l u ??? H4); + apply (restrict nat_ordered_uniform_space s ??? H4); generalize in match (H ? H5); cases x; intro; generalize in match (refl_eq ? (a i):a i = a i);