X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;ds=sidebyside;f=helm%2Fsoftware%2Fmatita%2Fcontribs%2Fformal_topology%2Foverlap%2Frelations.ma;h=c9685f2207f4cf0e93db0d23f2d89fde0fbd2bf7;hb=5be81fce195f2b45ec57c5422d35e4c03827891d;hp=e6f816156fe684296548c452f1d03a5583f2cd47;hpb=c78cbede35ed85575e274864e6b6b9c635c6956d;p=helm.git diff --git a/helm/software/matita/contribs/formal_topology/overlap/relations.ma b/helm/software/matita/contribs/formal_topology/overlap/relations.ma index e6f816156..c9685f220 100644 --- a/helm/software/matita/contribs/formal_topology/overlap/relations.ma +++ b/helm/software/matita/contribs/formal_topology/overlap/relations.ma @@ -103,11 +103,18 @@ definition full_subset: ∀s:REL. Ω \sup s. qed. coercion full_subset. +*) definition setoid1_of_REL: REL → setoid ≝ λS. S. - coercion setoid1_of_REL. +lemma Type_OF_setoid1_of_REL: ∀o1:Type_OF_category1 REL. Type_OF_objs1 o1 → Type_OF_setoid1 ?(*(setoid1_of_SET o1)*). + [ apply (setoid1_of_SET o1); + | intros; apply t;] +qed. +coercion Type_OF_setoid1_of_REL. + +(* definition comprehension: ∀b:REL. (b ⇒ CPROP) → Ω \sup b. apply (λb:REL. λP: b ⇒ CPROP. {x | x ∈ b ∧ P x}); intros; simplify; apply (.= (H‡#)‡(†H)); apply refl1. @@ -174,21 +181,11 @@ lemma extS_com: ∀o1,o2,o3,c1,c2,S. extS o1 o3 (c2 ∘ c1) S = extS o1 o2 c1 (e assumption] qed. *) -(* senza questo exT "fresco", universe inconsistency *) -inductive exT (A:Type0) (P:A→CProp0) : CProp0 ≝ - ex_introT: ∀w:A. P w → exT A P. - -lemma hint: ∀U. carr U → Type_OF_setoid1 ?(*(setoid1_of_SET U)*). - [ apply setoid1_of_SET; apply U - | intros; apply c;] -qed. -coercion hint. (* the same as ⋄ for a basic pair *) definition image: ∀U,V:REL. binary_morphism1 (arrows1 ? U V) (Ω \sup U) (Ω \sup V). intros; constructor 1; - [ apply (λr: arrows1 ? U V.λS: Ω \sup U. {y | (*∃x:U. x ♮r y ∧ x ∈ S*) - exT ? (λx:carr U.x ♮r y ∧ x ∈ S) }); + [ apply (λr: arrows1 ? U V.λS: Ω \sup U. {y | ∃x:carr U. x ♮r y ∧ x ∈ S }); intros; simplify; split; intro; cases e1; exists [1,3: apply w] [ apply (. (#‡e)‡#); assumption | apply (. (#‡e ^ -1)‡#); assumption] @@ -294,8 +291,8 @@ qed. *) include "o-algebra.ma". -axiom daemon: False. -definition orelation_of_relation: ∀o1,o2:REL. arrows1 ? o1 o2 → ORelation (SUBSETS o1) (SUBSETS o2). + +definition orelation_of_relation: ∀o1,o2:REL. arrows1 ? o1 o2 → arrows2 OA (SUBSETS o1) (SUBSETS o2). intros; constructor 1; [ constructor 1; @@ -331,4 +328,82 @@ definition orelation_of_relation: ∀o1,o2:REL. arrows1 ? o1 o2 → ORelation (S | cases x1; cases x2; clear x1 x2; exists; [apply w1] [ exists; [apply w] split; assumption; | assumption; ]]] -qed. sistemare anche l'hint da un'altra parte e capire l'exT (doppio!) \ No newline at end of file +qed. + +lemma orelation_of_relation_preserves_equality: + ∀o1,o2:REL.∀t,t': arrows1 ? o1 o2. eq1 ? t t' → orelation_of_relation ?? t = orelation_of_relation ?? t'. + intros; split; unfold orelation_of_relation; simplify; intro; split; intro; + simplify; whd in o1 o2; + [ change with (a1 ∈ minus_star_image ?? t a → a1 ∈ minus_star_image ?? t' a); + apply (. #‡(e‡#)); + | change with (a1 ∈ minus_star_image ?? t' a → a1 ∈ minus_star_image ?? t a); + apply (. #‡(e ^ -1‡#)); + | change with (a1 ∈ minus_image ?? t a → a1 ∈ minus_image ?? t' a); + apply (. #‡(e‡#)); + | change with (a1 ∈ minus_image ?? t' a → a1 ∈ minus_image ?? t a); + apply (. #‡(e ^ -1‡#)); + | change with (a1 ∈ image ?? t a → a1 ∈ image ?? t' a); + apply (. #‡(e‡#)); + | change with (a1 ∈ image ?? t' a → a1 ∈ image ?? t a); + apply (. #‡(e ^ -1‡#)); + | change with (a1 ∈ star_image ?? t a → a1 ∈ star_image ?? t' a); + apply (. #‡(e‡#)); + | change with (a1 ∈ star_image ?? t' a → a1 ∈ star_image ?? t a); + apply (. #‡(e ^ -1‡#)); ] +qed. + +lemma hint: ∀o1,o2:OA. Type_OF_setoid2 (arrows2 ? o1 o2) → carr2 (arrows2 OA o1 o2). + intros; apply t; +qed. +coercion hint. + +lemma orelation_of_relation_preserves_identity: + ∀o1:REL. orelation_of_relation ?? (id1 ? o1) = id2 OA (SUBSETS o1). + intros; split; intro; split; whd; intro; + [ change with ((∀x. x ♮(id1 REL o1) a1→x∈a) → a1 ∈ a); intros; + apply (f a1); change with (a1 = a1); apply refl1; + | change with (a1 ∈ a → ∀x. x ♮(id1 REL o1) a1→x∈a); intros; + change in f1 with (x = a1); apply (. f1 ^ -1‡#); apply f; + | alias symbol "and" = "and_morphism". + change with ((∃y: carr o1.a1 ♮(id1 REL o1) y ∧ y∈a) → a1 ∈ a); + intro; cases e; clear e; cases x; clear x; change in f with (a1=w); + apply (. f^-1‡#); apply f1; + | change with (a1 ∈ a → ∃y: carr o1.a1 ♮(id1 REL o1) y ∧ y∈a); + intro; exists; [apply a1]; split; [ change with (a1=a1); apply refl1; | apply f] + | change with ((∃x: carr o1.x ♮(id1 REL o1) a1∧x∈a) → a1 ∈ a); + intro; cases e; clear e; cases x; clear x; change in f with (w=a1); + apply (. f‡#); apply f1; + | change with (a1 ∈ a → ∃x: carr o1.x ♮(id1 REL o1) a1∧x∈a); + intro; exists; [apply a1]; split; [ change with (a1=a1); apply refl1; | apply f] + | change with ((∀y.a1 ♮(id1 REL o1) y→y∈a) → a1 ∈ a); intros; + apply (f a1); change with (a1 = a1); apply refl1; + | change with (a1 ∈ a → ∀y.a1 ♮(id1 REL o1) y→y∈a); intros; + change in f1 with (a1 = y); apply (. f1‡#); apply f;] +qed. + +lemma hint2: ∀S,T. carr2 (arrows2 OA S T) → Type_OF_setoid2 (arrows2 OA S T). + intros; apply c; +qed. +coercion hint2. + +(* CSC: ???? forse un uncertain mancato *) +lemma orelation_of_relation_preserves_composition: + ∀o1,o2,o3:REL.∀F: arrows1 ? o1 o2.∀G: arrows1 ? o2 o3. + orelation_of_relation ?? (G ∘ F) = + comp2 OA (SUBSETS o1) (SUBSETS o2) (SUBSETS o3) + ?? (*(orelation_of_relation ?? F) (orelation_of_relation ?? G)*). + [ apply (orelation_of_relation ?? F); | apply (orelation_of_relation ?? G); ] + intros; split; intro; split; whd; intro; whd in ⊢ (% → %); intros; + [ whd; intros; apply f; exists; [ apply x] split; assumption; + | cases f1; clear f1; cases x1; clear x1; apply (f w); assumption; + | cases e; cases x; cases f; cases x1; clear e x f x1; exists; [ apply w1 ] + split; [ assumption | exists; [apply w] split; assumption ] + | cases e; cases x; cases f1; cases x1; clear e x f1 x1; exists; [apply w1 ] + split; [ exists; [apply w] split; assumption | assumption ] + | cases e; cases x; cases f; cases x1; clear e x f x1; exists; [ apply w1 ] + split; [ assumption | exists; [apply w] split; assumption ] + | cases e; cases x; cases f1; cases x1; clear e x f1 x1; exists; [apply w1 ] + split; [ exists; [apply w] split; assumption | assumption ] + | whd; intros; apply f; exists; [ apply y] split; assumption; + | cases f1; clear f1; cases x; clear x; apply (f w); assumption;] +qed. \ No newline at end of file