| apply (. (#‡e1)‡(e‡#)); assumption]]
qed.
-interpretation "basic pair relation for subsets" 'Vdash2 x y = (fun21 (concr _) __ (relS _) x y).
-interpretation "basic pair relation for subsets (non applied)" 'Vdash = (fun21 ___ (relS _)).
+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)).
include "basic_pairs.ma".
-(* full_subset e' una coercion che non mette piu' *)
+(* carr1 e' necessario perche' ci sega via la coercion per gli oggetti di REL!
+ (confondendola con la coercion per gli oggetti di SET *)
record concrete_space : Type1 ≝
{ bp:> BP;
- converges: ∀a: concr bp.∀U,V: form bp. a ⊩ U → a ⊩ V → a ⊩ (U ↓ V);
- all_covered: ∀x: concr bp. x ⊩ full_subset (form bp)
+ converges: ∀a: carr1 (concr bp).∀U,V: carr1 (form bp). a ⊩ U → a ⊩ V → a ⊩_bp (U ↓ V);
+ all_covered: ∀x: carr1 (concr bp). x ⊩ form bp
}.
record convergent_relation_pair (CS1,CS2: concrete_space) : Type1 ≝
definition CProp0 : Type1 := Type0.
definition CProp1 : Type2 := Type1.
definition CProp2 : Type3 := Type2.
+definition CProp_of_CProp0: CProp0 → CProp ≝ λx.x.
+definition CProp_of_CProp1: CProp1 → CProp ≝ λx.x.
+definition CProp_of_CProp2: CProp2 → CProp ≝ λx.x.
+coercion CProp_of_CProp0.
+coercion CProp_of_CProp1.
+coercion CProp_of_CProp2.
inductive Or (A,B:CProp0) : CProp0 ≝
| Left : A → Or A B