X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fdama%2Fdama%2Fproperty_exhaustivity.ma;h=f7250f9e2e167a7a428b1d604c14638142580772;hb=59f65aaf6f8d23748e1294ecabffffaa903ae657;hp=44be295434eb62f3b72b9c391c396eb431126c83;hpb=99f153e43f18bc682339bed41c8230af2ac6fd2f;p=helm.git diff --git a/helm/software/matita/contribs/dama/dama/property_exhaustivity.ma b/helm/software/matita/contribs/dama/dama/property_exhaustivity.ma index 44be29543..f7250f9e2 100644 --- a/helm/software/matita/contribs/dama/dama/property_exhaustivity.ma +++ b/helm/software/matita/contribs/dama/dama/property_exhaustivity.ma @@ -23,39 +23,39 @@ definition exhaustive ≝ (b is_decreasing → b is_lower_located → b is_cauchy). lemma segment_upperbound: - ∀C:ordered_set.∀l,u:C.∀a:sequence {[l,u]}.u is_upper_bound (λn.fst (a n)). -intros 5; change with (fst (a n) ≤ u); cases (a n); cases H; assumption; + ∀C:ordered_set.∀l,u:C.∀a:sequence {[l,u]}.u is_upper_bound ⌊n,\fst (a n)⌋. +intros 5; change with (\fst (a n) ≤ u); cases (a n); cases H; assumption; qed. lemma segment_lowerbound: - ∀C:ordered_set.∀l,u:C.∀a:sequence {[l,u]}.l is_lower_bound (λn.fst (a n)). -intros 5; change with (l ≤ fst (a n)); cases (a n); cases H; assumption; + ∀C:ordered_set.∀l,u:C.∀a:sequence {[l,u]}.l is_lower_bound ⌊n,\fst (a n)⌋. +intros 5; change with (l ≤ \fst (a n)); cases (a n); cases H; assumption; qed. lemma segment_preserves_uparrow: ∀C:ordered_set.∀l,u:C.∀a:sequence {[l,u]}.∀x,h. - (λn.fst (a n)) ↑ x → a ↑ (sig_in ?? x h). + ⌊n,\fst (a n)⌋ ↑ x → a ↑ ≪x,h≫. intros; cases H (Ha Hx); split [apply Ha] cases Hx; split; [apply H1] intros; -cases (H2 (fst y)); [2: apply H3;] exists [apply w] assumption; +cases (H2 (\fst y)); [2: apply H3;] exists [apply w] assumption; qed. lemma segment_preserves_downarrow: ∀C:ordered_set.∀l,u:C.∀a:sequence {[l,u]}.∀x,h. - (λn.fst (a n)) ↓ x → a ↓ (sig_in ?? x h). + ⌊n,\fst (a n)⌋ ↓ x → a ↓ ≪x,h≫. intros; cases H (Ha Hx); split [apply Ha] cases Hx; split; [apply H1] intros; -cases (H2 (fst y));[2:apply H3]; exists [apply w] assumption; +cases (H2 (\fst y));[2:apply H3]; exists [apply w] assumption; qed. (* Fact 2.18 *) lemma segment_cauchy: ∀C:ordered_uniform_space.∀l,u:C.∀a:sequence {[l,u]}. - a is_cauchy → (λn:nat.fst (a n)) is_cauchy. + a is_cauchy → ⌊n,\fst (a n)⌋ is_cauchy. intros 7; alias symbol "pi1" (instance 3) = "pair pi1". alias symbol "pi2" = "pair pi2". -apply (H (λx:{[l,u]} square.U 〈fst (fst x),fst (snd x)〉)); +apply (H (λx:{[l,u]} square.U 〈\fst (\fst x),\fst (\snd x)〉)); (unfold segment_ordered_uniform_space; simplify); exists [apply U] split; [assumption;] intro; cases b; intros; simplify; split; intros; assumption; @@ -65,18 +65,18 @@ qed. lemma restrict_uniform_convergence_uparrow: ∀C:ordered_uniform_space.property_sigma C → ∀l,u:C.exhaustive {[l,u]} → - ∀a:sequence {[l,u]}.∀x:C. (λn.fst (a n)) ↑ x → - x∈[l,u] ∧ ∀h:x ∈ [l,u].a uniform_converges (sig_in ?? x h). + ∀a:sequence {[l,u]}.∀x:C. ⌊n,\fst (a n)⌋ ↑ x → + x∈[l,u] ∧ ∀h:x ∈ [l,u].a uniform_converges ≪x,h≫. intros; cases H2 (Ha Hx); clear H2; cases Hx; split; [1: split; [1: apply (supremum_is_upper_bound C ?? Hx u); apply (segment_upperbound ? l); - |2: apply (le_transitive ?? (fst (a 0))); [2: apply H2;] + |2: apply (le_transitive ? ??? ? (H2 O)); apply (segment_lowerbound ?l u);] |2: intros; - lapply (uparrow_upperlocated ? a (sig_in ?? x h)) as Ha1; + lapply (uparrow_upperlocated ? a ≪x,h≫) as Ha1; [2: apply segment_preserves_uparrow;split; assumption;] - lapply (segment_preserves_supremum ?l u a (sig_in ??? h)) as Ha2; + lapply (segment_preserves_supremum ? l u a ≪?,h≫) as Ha2; [2:split; assumption]; cases Ha2; clear Ha2; cases (H1 a a); lapply (H6 H4 Ha1) as HaC; lapply (segment_cauchy ? l u ? HaC) as Ha; @@ -87,18 +87,18 @@ qed. lemma restrict_uniform_convergence_downarrow: ∀C:ordered_uniform_space.property_sigma C → ∀l,u:C.exhaustive {[l,u]} → - ∀a:sequence {[l,u]}.∀x:C. (λn.fst (a n)) ↓ x → - x∈[l,u] ∧ ∀h:x ∈ [l,u].a uniform_converges (sig_in ?? x h). + ∀a:sequence {[l,u]}.∀x:C. ⌊n,\fst (a n)⌋ ↓ x → + x∈[l,u] ∧ ∀h:x ∈ [l,u].a uniform_converges ≪x,h≫. intros; cases H2 (Ha Hx); clear H2; cases Hx; split; [1: split; [2: apply (infimum_is_lower_bound C ?? Hx l); apply (segment_lowerbound ? l u); - |1: apply (le_transitive ?? (fst (a 0))); [apply H2;] + |1: apply (le_transitive ???? (H2 O)); apply (segment_upperbound ? l u);] |2: intros; - lapply (downarrow_lowerlocated ? a (sig_in ?? x h)) as Ha1; + lapply (downarrow_lowerlocated ? a ≪x,h≫) as Ha1; [2: apply segment_preserves_downarrow;split; assumption;] - lapply (segment_preserves_infimum ?l u a (sig_in ??? h)) as Ha2; + lapply (segment_preserves_infimum ?l u a ≪?,h≫) as Ha2; [2:split; assumption]; cases Ha2; clear Ha2; cases (H1 a a); lapply (H7 H4 Ha1) as HaC; lapply (segment_cauchy ? l u ? HaC) as Ha;