From 0a443e57d4951768408fc7e1a4397dc67b31047b Mon Sep 17 00:00:00 2001 From: Claudio Sacerdoti Coen <claudio.sacerdoticoen@unibo.it> Date: Sun, 4 Jan 2009 10:30:12 +0000 Subject: [PATCH] More re-organization. --- .../formal_topology/overlap/o-algebra.ma | 5 +++ .../overlap/o-concrete_spaces.ma | 33 ++++++++----------- 2 files changed, 18 insertions(+), 20 deletions(-) diff --git a/helm/software/matita/contribs/formal_topology/overlap/o-algebra.ma b/helm/software/matita/contribs/formal_topology/overlap/o-algebra.ma index 533bff3dd..c6d91db94 100644 --- a/helm/software/matita/contribs/formal_topology/overlap/o-algebra.ma +++ b/helm/software/matita/contribs/formal_topology/overlap/o-algebra.ma @@ -345,3 +345,8 @@ lemma objs2_SET1_OF_OA: OA â objs2 SET1. intro; whd; apply (setoid1_of_OA t); qed. coercion objs2_SET1_OF_OA. + +lemma Type_OF_category2_OF_SET1_OF_OA: OA â Type_OF_category2 SET1. + intro; apply (oa_P t); +qed. +coercion Type_OF_category2_OF_SET1_OF_OA. \ No newline at end of file diff --git a/helm/software/matita/contribs/formal_topology/overlap/o-concrete_spaces.ma b/helm/software/matita/contribs/formal_topology/overlap/o-concrete_spaces.ma index 0bd8715a6..78ea321a5 100644 --- a/helm/software/matita/contribs/formal_topology/overlap/o-concrete_spaces.ma +++ b/helm/software/matita/contribs/formal_topology/overlap/o-concrete_spaces.ma @@ -17,51 +17,44 @@ include "o-saturations.ma". notation "â¡ \sub b" non associative with precedence 90 for @{'box $b}. notation > "â¡_term 90 b" non associative with precedence 90 for @{'box $b}. -interpretation "Universal image â©â»*" 'box x = (fun_1 _ _ (or_f_minus_star _ _) (rel x)). +interpretation "Universal image â©â»*" 'box x = (fun12 _ _ (or_f_minus_star _ _) (rel x)). notation "â \sub b" non associative with precedence 90 for @{'diamond $b}. notation > "â_term 90 b" non associative with precedence 90 for @{'diamond $b}. -interpretation "Existential image â©" 'diamond x = (fun_1 _ _ (or_f _ _) (rel x)). +interpretation "Existential image â©" 'diamond x = (fun12 _ _ (or_f _ _) (rel x)). notation "'Rest' \sub b" non associative with precedence 90 for @{'rest $b}. notation > "'Rest'â½term 90 b" non associative with precedence 90 for @{'rest $b}. -interpretation "Universal pre-image â©*" 'rest x = (fun_1 _ _ (or_f_star _ _) (rel x)). +interpretation "Universal pre-image â©*" 'rest x = (fun12 _ _ (or_f_star _ _) (rel x)). notation "'Ext' \sub b" non associative with precedence 90 for @{'ext $b}. notation > "'Ext'â½term 90 b" non associative with precedence 90 for @{'ext $b}. -interpretation "Existential pre-image â©â»" 'ext x = (fun_1 _ _ (or_f_minus _ _) (rel x)). +interpretation "Existential pre-image â©â»" 'ext x = (fun12 _ _ (or_f_minus _ _) (rel x)). -lemma hint : âp,q.arrows1 OA p q â ORelation_setoid p q. -intros; assumption; -qed. - -coercion hint nocomposites. - -definition A : âb:BP. unary_morphism (oa_P (form b)) (oa_P (form b)). +definition A : âb:BP. unary_morphism1 (form b) (form b). intros; constructor 1; [ apply (λx.â¡_b (Extâ½b x)); - | do 2 unfold uncurry_arrows; intros; apply (â (â H));] + | do 2 unfold FunClass_1_OF_Type_OF_setoid21; intros; apply (â (â e));] qed. +(* lemma xxx : âx.carr x â carr1 (setoid1_of_setoid x). intros; assumption; qed. coercion xxx nocomposites. +*) -lemma down_p : âS,I:SET.âu:SâS.âc:arrows1 SET I S.âa:I.âa':I.a=a'âu (c a)=u (c a'). -intros; unfold uncurry_arrows; change in c with (I â S); -apply (â (â H)); +lemma down_p : âS:SET1.âI:SET.âu:SâS.âc:arrows2 SET1 I S.âa:I.âa':I.a=a'âu (c a)=u (c a'). +intros; apply (â (â e)); qed. -alias symbol "eq" = "setoid eq". -alias symbol "and" = "o-algebra binary meet". -record concrete_space : Type â +record concrete_space : Type2 â { bp:> BP; (*distr : is_distributive (form bp);*) - downarrow: unary_morphism (oa_P (form bp)) (oa_P (form bp)); + downarrow: unary_morphism1 (form bp) (form bp); downarrow_is_sat: is_saturation ? downarrow; converges: âq1,q2. (Extâ½bp q1 ⧠(Extâ½bp q2)) = (Extâ½bp ((downarrow q1) ⧠(downarrow q2))); all_covered: Extâ½bp (oa_one (form bp)) = oa_one (concr bp); - il2: âI:SET.âp:arrows1 SET I (oa_P (form bp)). + il2: âI:SET.âp:arrows2 SET1 I (form bp). downarrow (⨠{ x â I | downarrow (p x) | down_p ???? }) = ⨠{ x â I | downarrow (p x) | down_p ???? }; il1: âq.downarrow (A ? q) = A ? q -- 2.39.2