From: Claudio Sacerdoti Coen Date: Sun, 4 Jul 2010 11:24:14 +0000 (+0000) Subject: Some important proofs/definitions were (and are still) commented out and X-Git-Tag: make_still_working~2883 X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=commitdiff_plain;h=82b1a205fdf9bc2c8029296ebe94c5667798845b;p=helm.git Some important proofs/definitions were (and are still) commented out and do not compile. Added the notion of functor1 to state the main theorem. --- diff --git a/helm/software/matita/contribs/formal_topology/overlap/basic_pairs.ma b/helm/software/matita/contribs/formal_topology/overlap/basic_pairs.ma index 44f8e9f21..8e5421c1f 100644 --- a/helm/software/matita/contribs/formal_topology/overlap/basic_pairs.ma +++ b/helm/software/matita/contribs/formal_topology/overlap/basic_pairs.ma @@ -172,6 +172,7 @@ definition relation_pair_setoid_of_arrows1_BP : ∀P,Q. arrows1 BP P Q → relation_pair_setoid P Q ≝ λP,Q,x.x. coercion relation_pair_setoid_of_arrows1_BP. +(* definition BPext: ∀o: BP. (form o) ⇒_1 Ω^(concr o). intros; constructor 1; [ apply (ext ? ? (rel o)); @@ -219,3 +220,4 @@ qed. interpretation "basic pair relation for subsets" 'Vdash2 x y c = (fun21 (concr ?) ?? (relS c) x y). interpretation "basic pair relation for subsets (non applied)" 'Vdash c = (fun21 ??? (relS c)). +*) diff --git a/helm/software/matita/contribs/formal_topology/overlap/basic_pairs_to_basic_topologies.ma b/helm/software/matita/contribs/formal_topology/overlap/basic_pairs_to_basic_topologies.ma index 5a98136bc..b557bedae 100644 --- a/helm/software/matita/contribs/formal_topology/overlap/basic_pairs_to_basic_topologies.ma +++ b/helm/software/matita/contribs/formal_topology/overlap/basic_pairs_to_basic_topologies.ma @@ -40,4 +40,23 @@ definition continuous_relation_of_relation_pair: [ apply (rp \sub \f); | apply (Oreduced ?? ocr); | apply (Osaturated ?? ocr); ] +qed. + +alias symbol "compose" (instance 3) = "category1 composition". +alias symbol "compose" (instance 3) = "category1 composition". +record functor1 (C1: category1) (C2: category1) : Type2 ≝ + { map_objs1:1> C1 → C2; + map_arrows1: ∀S,T. unary_morphism1 (arrows1 ? S T) (arrows1 ? (map_objs1 S) (map_objs1 T)); + respects_id1: ∀o:C1. map_arrows1 ?? (id1 ? o) = id1 ? (map_objs1 o); + respects_comp1: + ∀o1,o2,o3.∀f1:arrows1 ? o1 o2.∀f2:arrows1 ? o2 o3. + map_arrows1 ?? (f2 ∘ f1) = map_arrows1 ?? f2 ∘ map_arrows1 ?? f1}. + +definition BTop_of_BP: functor1 BP BTop. + lapply OR as F; + constructor 1; + [ apply basic_topology_of_basic_pair + | intros; constructor 1 [ apply continuous_relation_of_relation_pair; ] + | simplify; intro; + ] qed. \ No newline at end of file diff --git a/helm/software/matita/contribs/formal_topology/overlap/basic_topologies.ma b/helm/software/matita/contribs/formal_topology/overlap/basic_topologies.ma index a48aae41c..b24a730cb 100644 --- a/helm/software/matita/contribs/formal_topology/overlap/basic_topologies.ma +++ b/helm/software/matita/contribs/formal_topology/overlap/basic_topologies.ma @@ -29,20 +29,25 @@ record continuous_relation (S,T: basic_topology) : Type1 ≝ reduced: ∀U. U = J ? U → image ?? cont_rel U = J ? (image ?? cont_rel U); saturated: ∀U. U = A ? U → minus_star_image ?? cont_rel U = A ? (minus_star_image ?? cont_rel U) }. -(* + definition continuous_relation_setoid: basic_topology → basic_topology → setoid1. intros (S T); constructor 1; [ apply (continuous_relation S T) | constructor 1; [ apply (λr,s:continuous_relation S T.∀b. A ? (ext ?? r b) = A ? (ext ?? s b)); | simplify; intros; apply refl1; - | simplify; intros; apply sym1; apply H - | simplify; intros; apply trans1; [2: apply H |3: apply H1; |1: skip]]] + | simplify; intros (x y H); apply sym1; apply H + | simplify; intros; apply trans1; [2: apply f |3: apply f1; |1: skip]]] qed. -theorem continuous_relation_eq': +definition continuos_relation_of_continuous_relation_setoid : + ∀P,Q. continuous_relation_setoid P Q → continuous_relation P Q ≝ λP,Q,x.x. +coercion continuos_relation_of_continuous_relation_setoid. + +axiom continuous_relation_eq': ∀o1,o2.∀a,a': continuous_relation_setoid o1 o2. a = a' → ∀X.minus_star_image ?? a (A o1 X) = minus_star_image ?? a' (A o1 X). +(* intros; split; intro; unfold minus_star_image; simplify; intros; [ cut (ext ?? a a1 ⊆ A ? X); [2: intros 2; apply (H1 a2); cases f1; assumption;] lapply (if ?? (A_is_saturation ???) Hcut); clear Hcut; @@ -56,12 +61,12 @@ theorem continuous_relation_eq': lapply (fi ?? (A_is_saturation ???) Hcut); apply (Hletin1 x); change with (x ∈ ext ?? a a1); split; simplify; [ apply I | assumption ]] -qed. +qed.*) -theorem continuous_relation_eq_inv': +axiom continuous_relation_eq_inv': ∀o1,o2.∀a,a': continuous_relation_setoid o1 o2. (∀X.minus_star_image ?? a (A o1 X) = minus_star_image ?? a' (A o1 X)) → a=a'. - intros 6; +(* intros 6; cut (∀a,a': continuous_relation_setoid o1 o2. (∀X.minus_star_image ?? a (A o1 X) = minus_star_image ?? a' (A o1 X)) → ∀V:o2. A ? (ext ?? a' V) ⊆ A ? (ext ?? a V)); @@ -86,6 +91,7 @@ theorem continuous_relation_eq_inv': assumption;] split; apply Hcut; [2: assumption | intro; apply sym1; apply H] qed. +*) definition continuous_relation_comp: ∀o1,o2,o3. @@ -131,15 +137,16 @@ definition BTop: category1. | intros; constructor 1; [ apply continuous_relation_comp; | intros; simplify; intro x; simplify; - lapply depth=0 (continuous_relation_eq' ???? H) as H'; - lapply depth=0 (continuous_relation_eq' ???? H1) as H1'; + lapply depth=0 (continuous_relation_eq' ???? e) as H'; + lapply depth=0 (continuous_relation_eq' ???? e1) as H1'; letin K ≝ (λX.H1' (minus_star_image ?? a (A ? X))); clearbody K; cut (∀X:Ω \sup o1. minus_star_image o2 o3 b (A o2 (minus_star_image o1 o2 a (A o1 X))) = minus_star_image o2 o3 b' (A o2 (minus_star_image o1 o2 a' (A o1 X)))); [2: intro; apply sym1; apply (.= #‡(†((H' ?)\sup -1))); apply sym1; apply (K X);] clear K H' H1'; - cut (∀X:Ω \sup o1. + alias symbol "compose" (instance 2) = "category1 composition". +cut (∀X:Ω \sup o1. minus_star_image o1 o3 (b ∘ a) (A o1 X) = minus_star_image o1 o3 (b'∘a') (A o1 X)); [2: intro; apply (.= (minus_star_image_comp ??????)); @@ -154,7 +161,12 @@ definition BTop: category1. apply (continuous_relation_eq_inv'); apply Hcut1;] | intros; simplify; intro; do 2 (unfold continuous_relation_comp); simplify; - apply (.= †(ASSOC1‡#)); + alias symbol "trans" (instance 1) = "trans1". +alias symbol "refl" (instance 5) = "refl1". +alias symbol "prop2" (instance 3) = "prop21". +alias symbol "prop1" (instance 2) = "prop11". +alias symbol "assoc" (instance 4) = "category1 assoc". +apply (.= †(ASSOC‡#)); apply refl1 | intros; simplify; intro; unfold continuous_relation_comp; simplify; apply (.= †((id_neutral_right1 ????)‡#)); @@ -190,4 +202,3 @@ theorem continuous_relation_eqS: apply Hcut2; assumption. qed. *) -*) \ No newline at end of file diff --git a/helm/software/matita/contribs/formal_topology/overlap/relations.ma b/helm/software/matita/contribs/formal_topology/overlap/relations.ma index f7827939d..b1589a827 100644 --- a/helm/software/matita/contribs/formal_topology/overlap/relations.ma +++ b/helm/software/matita/contribs/formal_topology/overlap/relations.ma @@ -240,25 +240,25 @@ definition minus_image: ∀U,V:REL. (U ⇒_\r1 V) × Ω^V ⇒_1 Ω^U. apply (if ?? (e ^ -1 ??)); assumption]] qed. -(* (* minus_image is the same as ext *) theorem image_id: ∀o,U. image o o (id1 REL o) U = U. intros; unfold image; simplify; split; simplify; intros; [ change with (a ∈ U); - cases H; cases x; change in f with (eq1 ? w a); apply (. f‡#); assumption + cases e; cases x; change in f with (eq1 ? w a); apply (. f^-1‡#); assumption | change in f with (a ∈ U); - exists; [apply a] split; [ change with (a = a); apply refl | assumption]] + exists; [apply a] split; [ change with (a = a); apply refl1 | assumption]] qed. theorem minus_star_image_id: ∀o,U. minus_star_image o o (id1 REL o) U = U. intros; unfold minus_star_image; simplify; split; simplify; intros; - [ change with (a ∈ U); apply H; change with (a=a); apply refl - | change in f1 with (eq1 ? x a); apply (. f1 \sup -1‡#); apply f] + [ change with (a ∈ U); apply f; change with (a=a); apply refl1 + | change in f1 with (eq1 ? x a); apply (. f1‡#); apply f] qed. +alias symbol "compose" (instance 2) = "category1 composition". theorem image_comp: ∀A,B,C,r,s,X. image A C (r ∘ s) X = image B C r (image A B s X). - intros; unfold image; simplify; split; simplify; intros; cases H; clear H; cases x; + intros; unfold image; simplify; split; simplify; intros; cases e; clear e; cases x; clear x; [ cases f; clear f; | cases f1; clear f1 ] exists; try assumption; cases x; clear x; split; try assumption; exists; try assumption; split; assumption. @@ -268,10 +268,11 @@ theorem minus_star_image_comp: ∀A,B,C,r,s,X. minus_star_image A C (r ∘ s) X = minus_star_image B C r (minus_star_image A B s X). intros; unfold minus_star_image; simplify; split; simplify; intros; whd; intros; - [ apply H; exists; try assumption; split; assumption - | change with (x ∈ X); cases f; cases x1; apply H; assumption] + [ apply f; exists; try assumption; split; assumption + | change with (x ∈ X); cases f1; cases x1; apply f; assumption] qed. +(* (*CSC: unused! *) theorem ext_comp: ∀o1,o2,o3: REL.