interpretation "prop12" 'prop1 c = (prop12 _____ c).
interpretation "prop2" 'prop2 l r = (prop2 ________ l r).
interpretation "prop21" 'prop2 l r = (prop21 ________ l r).
+interpretation "prop22" 'prop2 l r = (prop22 ________ l r).
interpretation "refl" 'refl = (refl ___).
interpretation "refl1" 'refl = (refl1 ___).
interpretation "refl2" 'refl = (refl2 ___).
notation "'ASSOC1'" with precedence 90 for @{'assoc1}.
notation "'ASSOC2'" with precedence 90 for @{'assoc2}.
-interpretation "category1 composition" 'compose x y = (fun22 ___ (comp2 ____) y x).
-interpretation "category1 assoc" 'assoc1 = (comp_assoc2 ________).
+interpretation "category2 composition" 'compose x y = (fun22 ___ (comp2 ____) y x).
+interpretation "category2 assoc" 'assoc1 = (comp_assoc2 ________).
interpretation "category1 composition" 'compose x y = (fun21 ___ (comp1 ____) y x).
interpretation "category1 assoc" 'assoc1 = (comp_assoc1 ________).
interpretation "category composition" 'compose x y = (fun2 ___ (comp ____) y x).
intros; apply (prop11 A B w a b e);
qed.
-definition hint: Type_OF_category2 SET1 → Type1.
+definition setoid2_OF_category2: Type_OF_category2 SET1 → setoid2.
+ intro; apply (setoid2_of_setoid1 t); qed.
+coercion setoid2_OF_category2.
+
+definition objs2_OF_category1: Type_OF_category1 SET → objs2 SET1.
+ intro; apply (setoid1_of_setoid t); qed.
+coercion objs2_OF_category1.
+
+definition Type1_OF_SET1: Type_OF_category2 SET1 → Type1.
intro; whd in t; apply (carr1 t);
qed.
-coercion hint.
+coercion Type1_OF_SET1.
interpretation "SET dagger" 'prop1 h = (prop11_SET1 _ _ _ _ _ h).
interpretation "unary morphism1" 'Imply a b = (arrows2 SET1 a b).
interpretation "unary morphism1 comprehension with proof" 'comprehension_by s \eta.f p =
(mk_unary_morphism1 s _ f p).
-definition hint: Type_OF_category2 SET1 → setoid2.
- intro; apply (setoid2_of_setoid1 t); qed.
-coercion hint.
-
-definition hint2: Type_OF_category1 SET → objs2 SET1.
- intro; apply (setoid1_of_setoid t); qed.
-coercion hint2.
-
(* per il set-indexing vedere capitolo BPTools (foundational tools), Sect. 0.3.4 complete
lattices, Definizione 0.9 *)
(* USARE L'ESISTENZIALE DEBOLE *)
interpretation "o-algebra binary meet" 'oa_meet_bin a b =
(fun21 ___ (binary_meet _) a b).
+coercion Type1_OF_OAlgebra nocomposites.
+
lemma oa_overlap_preservers_meet: ∀O:OAlgebra.∀p,q:O.p >< q → p >< (p ∧ q).
+(* next change to avoid universe inconsistency *)
+change in ⊢ (?→%→%→?) with (Type1_OF_OAlgebra O);
intros; lapply (oa_overlap_preserves_meet_ O p q f);
lapply (prop21 O O CPROP (oa_overlap O) p p ? (p ∧ q) # ?);
[3: apply (if ?? (Hletin1)); apply Hletin;|skip] apply refl1;
coercion arrows1_OF_ORelation_setoid.
-lemma umorphism_OF_ORelation_setoid : ∀P,Q. ORelation_setoid P Q → P ⇒ Q.
+lemma umorphism_OF_ORelation_setoid : ∀P,Q. ORelation_setoid P Q → unary_morphism1 P Q.
intros; apply (or_f ?? t);
qed.
coercion umorphism_OF_ORelation_setoid.
+lemma umorphism_setoid_OF_ORelation_setoid : ∀P,Q. ORelation_setoid P Q → unary_morphism1_setoid1 P Q.
+intros; apply (or_f ?? t);
+qed.
+
+coercion umorphism_setoid_OF_ORelation_setoid.
-lemma uncurry_arrows : ∀B,C. arrows1 SET B C → B → C.
-intros; apply ((fun1 ?? t) t1);
+lemma uncurry_arrows : ∀B,C. ORelation_setoid B C → B → C.
+intros; apply ((fun11 ?? t) t1);
qed.
coercion uncurry_arrows 1.
-lemma hint6 : ∀P,Q. arrows1 SET P Q → P ⇒ Q. intros; apply t;qed.
+lemma hint6: ∀P,Q. Type_OF_setoid2 (hint5 P ⇒ hint5 Q) → unary_morphism1 P Q.
+ intros; apply t;
+qed.
coercion hint6.
-(*
-lemma hint2: OAlgebra → setoid. intros; apply (oa_P o). qed.
-coercion hint2 nocomposites.
-*)
-
-
notation "r \sup *" non associative with precedence 90 for @{'OR_f_star $r}.
notation > "r *" non associative with precedence 90 for @{'OR_f_star $r}.
qed.
definition ORelation_composition : ∀P,Q,R.
- binary_morphism1 (ORelation_setoid P Q) (ORelation_setoid Q R) (ORelation_setoid P R).
+ binary_morphism2 (ORelation_setoid P Q) (ORelation_setoid Q R) (ORelation_setoid P R).
intros;
constructor 1;
[ intros (F G);
constructor 1;
[ apply (G ∘ F);
- | apply (G⎻* ∘ F⎻* );
+ | apply rule (G⎻* ∘ F⎻* );
| apply (F* ∘ G* );
| apply (F⎻ ∘ G⎻);
| intros;
apply or_prop3;
]
| intros; split; simplify;
- [1,3: unfold arrows1_OF_ORelation_setoid; apply ((†H)‡(†H1));
- |2,4: apply ((†H1)‡(†H));]]
+ [1,3: unfold umorphism_setoid_OF_ORelation_setoid; unfold arrows1_OF_ORelation_setoid; apply ((†e)‡(†e1));
+ |2,4: apply ((†e1)‡(†e));]]
qed.
-definition OA : category1.
+definition OA : category2.
split;
[ apply (OAlgebra);
| intros; apply (ORelation_setoid o o1);
| intro O; split;
- [1,2,3,4: apply id1;
+ [1,2,3,4: apply id2;
|5,6,7:intros; apply refl1;]
| apply ORelation_composition;
| intros (P Q R S F G H); split;
[ change with (H⎻* ∘ G⎻* ∘ F⎻* = H⎻* ∘ (G⎻* ∘ F⎻* ));
- apply (comp_assoc1 ????? (F⎻* ) (G⎻* ) (H⎻* ));
- | apply ((comp_assoc1 ????? (H⎻) (G⎻) (F⎻))^-1);
- | apply ((comp_assoc1 ????? F G H)^-1);
- | apply ((comp_assoc1 ????? H* G* F* ));]
-| intros; split; unfold ORelation_composition; simplify; apply id_neutral_left1;
-| intros; split; unfold ORelation_composition; simplify; apply id_neutral_right1;]
-qed.
\ No newline at end of file
+ apply (comp_assoc2 ????? (F⎻* ) (G⎻* ) (H⎻* ));
+ | apply ((comp_assoc2 ????? (H⎻) (G⎻) (F⎻))^-1);
+ | apply ((comp_assoc2 ????? F G H)^-1);
+ | apply ((comp_assoc2 ????? H* G* F* ));]
+| intros; split; unfold ORelation_composition; simplify; apply id_neutral_left2;
+| intros; split; unfold ORelation_composition; simplify; apply id_neutral_right2;]
+qed.
+
+lemma setoid1_of_OA: OA → setoid1.
+ intro; apply (oa_P t);
+qed.
+coercion setoid1_of_OA.
alias symbol "eq" = "setoid1 eq".
definition is_saturation ≝
- λC:OA.λA:unary_morphism (oa_P C) (oa_P C).
+ λC:OA.λA:unary_morphism1 C C.
∀U,V. (U ≤ A V) = (A U ≤ A V).
definition is_reduction ≝
- λC:OA.λJ:unary_morphism (oa_P C) (oa_P C).
+ λC:OA.λJ:unary_morphism1 C C.
∀U,V. (J U ≤ V) = (J U ≤ J V).
theorem saturation_expansive: ∀C,A. is_saturation C A → ∀U. U ≤ A U.
- intros; apply (fi ?? (H ??)); apply (oa_leq_refl C).
+ intros; apply (fi ?? (i ??)); apply (oa_leq_refl C).
qed.
theorem saturation_monotone:
∀C,A. is_saturation C A →
∀U,V. U ≤ V → A U ≤ A V.
- intros; apply (if ?? (H ??)); apply (oa_leq_trans C);
+ intros; apply (if ?? (i ??)); apply (oa_leq_trans C);
[apply V|3: apply saturation_expansive ]
assumption.
qed.
theorem saturation_idempotent: ∀C,A. is_saturation C A → ∀U.
- eq (oa_P C) (A (A U)) (A U).
+ eq1 C (A (A U)) (A U).
intros; apply (oa_leq_antisym C);
- [ apply (if ?? (H (A U) U)); apply (oa_leq_refl C).
+ [ apply (if ?? (i (A U) U)); apply (oa_leq_refl C).
| apply saturation_expansive; assumption]
-qed.
+qed.
\ No newline at end of file