From befe31089d1d45360b5b7681556c8a762800b3a2 Mon Sep 17 00:00:00 2001 From: Claudio Sacerdoti Coen Date: Tue, 6 Jan 2009 13:13:23 +0000 Subject: [PATCH] orelation_of_relation preserves equality and identities. --- .../formal_topology/overlap/relations.ma | 62 ++++++++++++++++++- 1 file changed, 60 insertions(+), 2 deletions(-) diff --git a/helm/software/matita/contribs/formal_topology/overlap/relations.ma b/helm/software/matita/contribs/formal_topology/overlap/relations.ma index ed8b83bcb..738035be0 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. @@ -285,7 +292,7 @@ qed. include "o-algebra.ma". -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; @@ -322,3 +329,54 @@ definition orelation_of_relation: ∀o1,o2:REL. arrows1 ? o1 o2 → ORelation (S [ exists; [apply w] split; assumption; | assumption; ]]] 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. \ No newline at end of file -- 2.39.2