From f71d9d74c33e4ab3d86dc8244a7bdd3576bb0ac0 Mon Sep 17 00:00:00 2001 From: Enrico Tassi Date: Sat, 25 Oct 2008 17:11:56 +0000 Subject: [PATCH] ... --- .../contribs/dama/dama/ordered_uniform.ma | 79 ++++++++++++------- .../matita/contribs/dama/dama/supremum.ma | 10 +-- 2 files changed, 54 insertions(+), 35 deletions(-) diff --git a/helm/software/matita/contribs/dama/dama/ordered_uniform.ma b/helm/software/matita/contribs/dama/dama/ordered_uniform.ma index 701651cb1..63c966056 100644 --- a/helm/software/matita/contribs/dama/dama/ordered_uniform.ma +++ b/helm/software/matita/contribs/dama/dama/ordered_uniform.ma @@ -49,9 +49,17 @@ interpretation "relation invertion" 'invert a = (invert_os_relation _ a). interpretation "relation invertion" 'invert_symbol = (invert_os_relation _). interpretation "relation invertion" 'invert_appl a x = (invert_os_relation _ a x). +lemma hint_segment: ∀O. + segment (Type_of_ordered_set O) → + segment (hos_carr (os_l O)). +intros; assumption; +qed. + +coercion hint_segment nocomposites. + lemma segment_square_of_ordered_set_square: - ∀O:ordered_set.∀u,v:O.∀x:O squareO. - \fst x ∈ [u,v] → \snd x ∈ [u,v] → {[u,v]} squareO. + ∀O:ordered_set.∀s:‡O.∀x:O squareO. + \fst x ∈ s → \snd x ∈ s → {[s]} squareO. intros; split; exists; [1: apply (\fst x) |3: apply (\snd x)] assumption; qed. @@ -60,59 +68,70 @@ coercion segment_square_of_ordered_set_square with 0 2 nocomposites. alias symbol "pi1" (instance 4) = "exT \fst". alias symbol "pi1" (instance 2) = "exT \fst". lemma ordered_set_square_of_segment_square : - ∀O:ordered_set.∀u,v:O.{[u,v]} squareO → O squareO ≝ - λO:ordered_set.λu,v:O.λb:{[u,v]} squareO.〈\fst(\fst b),\fst(\snd b)〉. + ∀O:ordered_set.∀s:‡O.{[s]} squareO → O squareO ≝ + λO:ordered_set.λs:‡O.λb:{[s]} squareO.〈\fst(\fst b),\fst(\snd b)〉. coercion ordered_set_square_of_segment_square nocomposites. lemma restriction_agreement : - ∀O:ordered_uniform_space.∀l,r:O.∀P:{[l,r]} squareO → Prop.∀OP:O squareO → Prop.Prop. -apply(λO:ordered_uniform_space.λl,r:O. - λP:{[l,r]} squareO → Prop. λOP:O squareO → Prop. + ∀O:ordered_uniform_space.∀s:‡O.∀P:{[s]} squareO → Prop.∀OP:O squareO → Prop.Prop. +apply(λO:ordered_uniform_space.λs:‡O. + λP:{[s]} squareO → Prop. λOP:O squareO → Prop. ∀b:O squareO.∀H1,H2.(P b → OP b) ∧ (OP b → P b)); [5,7: apply H1|6,8:apply H2]skip; qed. -lemma unrestrict: ∀O:ordered_uniform_space.∀l,r:O.∀U,u.∀x:{[l,r]} squareO. - restriction_agreement ? l r U u → U x → u x. -intros 7; cases x (b b1); cases b (w1 H1); cases b1 (w2 H2); clear b b1 x; +lemma unrestrict: ∀O:ordered_uniform_space.∀s:‡O.∀U,u.∀x:{[s]} squareO. + restriction_agreement ? s U u → U x → u x. +intros 6; cases x (b b1); cases b (w1 H1); cases b1 (w2 H2); clear b b1 x; cases (H 〈w1,w2〉 H1 H2) (L _); intro Uw; apply L; apply Uw; qed. -lemma restrict: ∀O:ordered_uniform_space.∀l,r:O.∀U,u.∀x:{[l,r]} squareO. - restriction_agreement ? l r U u → u x → U x. -intros 6; cases x (b b1); cases b (w1 H1); cases b1 (w2 H2); clear b1 b x; +lemma restrict: ∀O:ordered_uniform_space.∀s:‡O.∀U,u.∀x:{[s]} squareO. + restriction_agreement ? s U u → u x → U x. +intros 5; cases x (b b1); cases b (w1 H1); cases b1 (w2 H2); clear b1 b x; intros (Ra uw); cases (Ra 〈w1,w2〉 H1 H2) (_ R); apply R; apply uw; qed. lemma invert_restriction_agreement: - ∀O:ordered_uniform_space.∀l,r:O. - ∀U:{[l,r]} squareO → Prop.∀u:O squareO → Prop. - restriction_agreement ? l r U u → - restriction_agreement ? l r (\inv U) (\inv u). -intros 9; split; intro; -[1: apply (unrestrict ????? (segment_square_of_ordered_set_square ??? 〈\snd b,\fst b〉 H2 H1) H H3); -|2: apply (restrict ????? (segment_square_of_ordered_set_square ??? 〈\snd b,\fst b〉 H2 H1) H H3);] + ∀O:ordered_uniform_space.∀s:‡O. + ∀U:{[s]} squareO → Prop.∀u:O squareO → Prop. + restriction_agreement ? s U u → + restriction_agreement ? s (\inv U) (\inv u). +intros 8; split; intro; +[1: apply (unrestrict ???? (segment_square_of_ordered_set_square ?? 〈\snd b,\fst b〉 H2 H1) H H3); +|2: apply (restrict ???? (segment_square_of_ordered_set_square ?? 〈\snd b,\fst b〉 H2 H1) H H3);] qed. lemma bs2_of_bss2: - ∀O:ordered_set.∀u,v:O.(bishop_set_of_ordered_set {[u,v]}) squareB → (bishop_set_of_ordered_set O) squareB ≝ - λO:ordered_set.λu,v:O.λb:{[u,v]} squareO.〈\fst(\fst b),\fst(\snd b)〉. + ∀O:ordered_set.∀s:‡O.(bishop_set_of_ordered_set {[s]}) squareB → (bishop_set_of_ordered_set O) squareB ≝ + λO:ordered_set.λs:‡O.λb:{[s]} squareO.〈\fst(\fst b),\fst(\snd b)〉. coercion bs2_of_bss2 nocomposites. +(* +lemma xxx : + ∀O,s,x.bs2_of_bss2 (ordered_set_OF_ordered_uniform_space O) s x + = + x. +intros; reflexivity; +*) + lemma segment_ordered_uniform_space: - ∀O:ordered_uniform_space.∀u,v:O.ordered_uniform_space. -intros (O l r); apply mk_ordered_uniform_space; -[1: apply (mk_ordered_uniform_space_ {[l,r]}); + ∀O:ordered_uniform_space.∀s:‡O.ordered_uniform_space. +intros (O s); apply mk_ordered_uniform_space; +[1: apply (mk_ordered_uniform_space_ {[s]}); [1: alias symbol "and" = "constructive and". - letin f ≝ (λP:{[l,r]} squareO → Prop. ∃OP:O squareO → Prop. - (us_unifbase O OP) ∧ restriction_agreement ??? P OP); - apply (mk_uniform_space (bishop_set_of_ordered_set {[l,r]}) f); + letin f ≝ (λP:{[s]} squareO → Prop. ∃OP:O squareO → Prop. + (us_unifbase O OP) ∧ restriction_agreement ?? P OP); + apply (mk_uniform_space (bishop_set_of_ordered_set {[s]}) f); [1: intros (U H); intro x; simplify; cases H (w Hw); cases Hw (Gw Hwp); clear H Hw; intro Hm; - lapply (us_phi1 O w Gw x Hm) as IH; - apply (restrict ? l r ??? Hwp IH); + lapply (us_phi1 O w Gw x) as IH;[2:intro;apply Hm;cases H; clear H; + [left;apply (x2sx ? s (\fst x) (\snd x) H1); + |right;apply (x2sx ? s ?? H1);] + + apply (restrict ? s ??? Hwp IH); |2: intros (U V HU HV); cases HU (u Hu); cases HV (v Hv); clear HU HV; cases Hu (Gu HuU); cases Hv (Gv HvV); clear Hu Hv; cases (us_phi2 O u v Gu Gv) (w HW); cases HW (Gw Hw); clear HW; diff --git a/helm/software/matita/contribs/dama/dama/supremum.ma b/helm/software/matita/contribs/dama/dama/supremum.ma index 7a52e5f06..54e9d98f0 100644 --- a/helm/software/matita/contribs/dama/dama/supremum.ma +++ b/helm/software/matita/contribs/dama/dama/supremum.ma @@ -350,15 +350,15 @@ notation "'segment_preserves_infimum'" non associative with precedence 90 for @{ interpretation "segment_preserves_supremum" 'segment_preserves_supremum = (h_segment_preserves_supremum (os_l _)). interpretation "segment_preserves_infimum" 'segment_preserves_infimum = (h_segment_preserves_supremum (os_r _)). -(* TEST, ma quanto godo! *) -lemma segment_preserves_infimum2: +(* +test segment_preserves_infimum2: ∀O:ordered_set.∀s:‡O.∀a:sequence {[s]}.∀x:{[s]}. ⌊n,\fst (a n)⌋ is_decreasing ∧ (\fst x) is_infimum ⌊n,\fst (a n)⌋ → a ↓ x. intros; apply (segment_preserves_infimum s a x H); qed. *) - + (* Definition 2.10 *) alias symbol "pi2" = "pair pi2". alias symbol "pi1" = "pair pi1". @@ -394,8 +394,8 @@ interpretation "Ordered set lower locatedness" 'lower_located s = lemma h_uparrow_upperlocated: ∀C:half_ordered_set.∀a:sequence C.∀u:C.uparrow ? a u → upper_located ? a. intros (C a u H); cases H (H2 H3); clear H; intros 3 (x y Hxy); -cases H3 (H4 H5); clear H3; cases (hos_cotransitive ??? u Hxy) (W W); -[2: cases (H5 ? W) (w Hw); left; exists [apply w] assumption; +cases H3 (H4 H5); clear H3; cases (hos_cotransitive C y x u Hxy) (W W); +[2: cases (H5 x W) (w Hw); left; exists [apply w] assumption; |1: right; exists [apply u]; split; [apply W|apply H4]] qed. -- 2.39.2