ndefinition proofs: CProp → setoid.
#P; napply (mk_setoid ?????);
##[ napply P;
-##| napply (λ_,_.True);
- #x; #y; napply True; (* DIVERGE *)
- intro;
- constructor 1;
- [ apply A
- | intros;
- apply True
- | intro;
- constructor 1
- | intros 3;
- constructor 1
- | intros 5;
- constructor 1]
-qed.
+##| #x; #y; napply True;
+##|##*: nwhd; nrepeat (#_); napply I;
+##]
+nqed.
+(*
definition reflexive1 ≝ λA:Type.λR:A→A→CProp.∀x:A.R x x.
definition symmetric1 ≝ λC:Type.λlt:C→C→CProp. ∀x,y:C.lt x y → lt y x.
definition transitive1 ≝ λA:Type.λR:A→A→CProp.∀x,y,z:A.R x y → R y z → R x z.
| intros 5;
constructor 1]
qed.
+*)
-definition CCProp: setoid1.
+(*
+ndefinition CCProp: setoid1.
constructor 1;
[ apply CProp
| apply iff
| apply (H3 (H5 H))]]
qed.
-record function_space (A,B: setoid): Type ≝
- { f:1> A → B;
+*)
+
+nrecord function_space (A,B: setoid): Type[1] ≝
+ { f:1> carr A → carr B}.;
f_ok: ∀a,a':A. proofs (eq ? a a') → proofs (eq ? (f a) (f a'))
}.
+
notation "hbox(a break ⇒ b)" right associative with precedence 20 for @{ 'Imply $a $b }.
+(*
record function_space1 (A: setoid1) (B: setoid1): Type ≝
{ f1:1> A → B;
f1_ok: ∀a,a':A. proofs1 (eq1 ? a a') → proofs1 (eq1 ? (f1 a) (f1 a'))
}.
+*)
definition function_space_setoid: setoid → setoid → setoid.
intros (A B);