]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/library/algebra/CoRN/Setoids.ma
Nuova dimostrazione riflessiva di le_to_Bertrand.
[helm.git] / helm / software / matita / library / algebra / CoRN / Setoids.ma
index 982d61caf19297e62b99abb2ed35c1c98af465f3..ddbb807a85643681b923fc6ee91afb367b0d6c50 100644 (file)
@@ -215,19 +215,19 @@ qed.
 
 (* -----------------The product of setoids----------------------- *)
 
-definition prod_ap: \forall A,B : CSetoid.\forall c,d: ProdT A B. Prop \def
-\lambda A,B : CSetoid.\lambda c,d: ProdT A B.
-  ((cs_ap A (fstT A B c) (fstT A B d)) \or 
-   (cs_ap B (sndT A B c) (sndT A B d))).
-
-definition prod_eq: \forall A,B : CSetoid.\forall c,d: ProdT A B. Prop \def
-\lambda A,B : CSetoid.\lambda c,d: ProdT A B.
-  ((cs_eq A (fstT A B c) (fstT A B d)) \and 
-   (cs_eq B (sndT A B c) (sndT A B d))).
+definition prod_ap: \forall A,B : CSetoid.\forall c,d: Prod A B. Prop \def
+\lambda A,B : CSetoid.\lambda c,d: Prod A B.
+  ((cs_ap A (fst A B c) (fst A B d)) \or 
+   (cs_ap B (snd A B c) (snd A B d))).
+
+definition prod_eq: \forall A,B : CSetoid.\forall c,d: Prod A B. Prop \def
+\lambda A,B : CSetoid.\lambda c,d: Prod A B.
+  ((cs_eq A (fst A B c) (fst A B d)) \and 
+   (cs_eq B (snd A B c) (snd A B d))).
       
 
 lemma prodcsetoid_is_CSetoid: \forall A,B: CSetoid.
- is_CSetoid (ProdT A B) (prod_eq A B) (prod_ap A B).
+ is_CSetoid (Prod A B) (prod_eq A B) (prod_ap A B).
 intros.
 apply (mk_is_CSetoid ? (prod_eq A B) (prod_ap A B))
   [unfold.
@@ -301,7 +301,7 @@ qed.
 
 definition ProdCSetoid : \forall A,B: CSetoid. CSetoid \def
  \lambda A,B: CSetoid.
-  mk_CSetoid (ProdT A B) (prod_eq A B) (prod_ap A B) (prodcsetoid_is_CSetoid A B).
+  mk_CSetoid (Prod A B) (prod_eq A B) (prod_ap A B) (prodcsetoid_is_CSetoid A B).
 
 
 
@@ -1115,7 +1115,7 @@ intros.
 unfold.unfold.intros 2.elim x 2.elim y 2.
 simplify.
 intro.
-reduce in H2.
+normalize in H2.
 apply (un_op_wd_unfolded ? f ? ? H2).
 qed.
 
@@ -1123,7 +1123,7 @@ lemma restr_un_op_strext : \forall S:CSetoid. \forall P: S \to Prop.
 \forall f: CSetoid_un_op S. \forall pr: un_op_pres_pred S P f.  
 un_op_strext (mk_SubCSetoid S P) (restr_un_op S P f pr).
 intros.unfold.unfold. intros 2.elim y 2. elim x 2.
-intros.reduce in H2.
+intros.normalize in H2.
 apply (cs_un_op_strext ? f ? ? H2).
 qed.
 
@@ -1182,8 +1182,8 @@ intros.
 unfold.unfold.intros 2.elim x1 2. elim x2 2.intros 2. elim y1 2. elim y2 2.
 simplify.
 intros.
-reduce in H4.
-reduce in H5.
+normalize in H4.
+normalize in H5.
 apply (cs_bin_op_wd ? f ? ? ? ? H4 H5).
 qed.
 
@@ -1192,7 +1192,7 @@ lemma restr_bin_op_strext : \forall S:CSetoid. \forall P: S \to Prop.
 bin_op_strext (mk_SubCSetoid S P) (restr_bin_op S P f pr).
 intros.unfold.unfold. intros 2.elim x1 2. elim x2 2.intros 2. elim y1 2. elim y2 2.
 simplify.intros.
-reduce in H4.
+normalize in H4.
 apply (cs_bin_op_strext ? f ? ? ? ? H4).
 qed.