X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;ds=sidebyside;f=helm%2Fsoftware%2Fmatita%2Flibrary%2Falgebra%2FCoRN%2FSetoids.ma;h=982d61caf19297e62b99abb2ed35c1c98af465f3;hb=5c1b44dfefa085fbb56e23047652d3650be9d855;hp=cfc6f524e63b48492c88c0ec3330b56f46d3cd4c;hpb=0070f04a500ec6e12dc063ff5acb8a7622b51283;p=helm.git diff --git a/helm/software/matita/library/algebra/CoRN/Setoids.ma b/helm/software/matita/library/algebra/CoRN/Setoids.ma index cfc6f524e..982d61caf 100644 --- a/helm/software/matita/library/algebra/CoRN/Setoids.ma +++ b/helm/software/matita/library/algebra/CoRN/Setoids.ma @@ -48,10 +48,10 @@ record CSetoid : Type \def cs_proof : is_CSetoid cs_crr cs_eq cs_ap}. interpretation "setoid equality" - 'eq x y = (cic:/matita/algebra/CoRN/Setoid/cs_eq.con _ x y). + 'eq x y = (cic:/matita/algebra/CoRN/Setoids/cs_eq.con _ x y). interpretation "setoid apart" - 'neq x y = (cic:/matita/algebra/CoRN/Setoid/cs_ap.con _ x y). + 'neq x y = (cic:/matita/algebra/CoRN/Setoids/cs_ap.con _ x y). (* visto che sia "ap" che "eq" vanno in Prop e data la "tight-apartness", "cs_neq" e "ap" non sono la stessa cosa? *) @@ -101,7 +101,7 @@ generalize in match (ap_tight S y x). intro. generalize in match (ap_symmetric S y x). intro. elim H1. clear H1. elim H2. clear H2. -apply H1. unfold. intro. auto. +apply H1. unfold. intro. autobatch. qed. *) lemma eq_transitive : \forall S : CSetoid. transitive S (cs_eq S). @@ -135,7 +135,7 @@ qed. lemma eq_wdl : \forall S:CSetoid. \forall x,y,z:S. x = y \to x = z \to z = y. intros. -(* perche' auto non arriva in fondo ??? *) +(* perche' autobatch non arriva in fondo ??? *) apply (eq_transitive_unfolded ? ? x). apply eq_symmetric_unfolded. exact H1. @@ -361,7 +361,7 @@ intro P. intro x; intros y H H0. elim (csp'_strext P x y H). -auto. +autobatch. intro H1. elimtype False. @@ -503,12 +503,12 @@ generalize (H x1 x2 y y). intro H1. elim (H1 H0). -auto. +autobatch. intro H3. elim H3; intro H4. -auto. +autobatch. elim (ap_irreflexive _ _ H4). Qed. @@ -517,22 +517,22 @@ unfold Crel_strext, Crel_strext_rht in |- *; intros H x y1 y2 H0. generalize (H x x y1 y2 H0); intro H1. elim H1; intro H2. -auto. +autobatch. elim H2; intro H3. elim (ap_irreflexive _ _ H3). -auto. +autobatch. Qed. Lemma Crel_strextarg_imp_strext : Crel_strext_rht -> Crel_strext_lft -> Crel_strext. unfold Crel_strext, Crel_strext_lft, Crel_strext_rht in |- *; intros H H0 x1 x2 y1 y2 H1. -elim (H x1 y1 y2 H1); auto. +elim (H x1 y1 y2 H1); autobatch. intro H2. -elim (H0 x1 x2 y2 H2); auto. +elim (H0 x1 x2 y2 H2); autobatch. Qed. *) @@ -553,7 +553,7 @@ intro R. red in |- *; intros x y z H H0. elim (Ccsr_strext R x x y z H). -auto. +autobatch. intro H1; elimtype False. elim H1; intro H2. @@ -569,7 +569,7 @@ intro R. red in |- *; intros x y z H H0. elim (Ccsr_strext R x z y y H). -auto. +autobatch. intro H1; elimtype False. elim H1; intro H2. @@ -614,11 +614,11 @@ Lemma ap_strext : Crel_strext (cs_ap (c:=S)). red in |- *; intros x1 x2 y1 y2 H. case (ap_cotransitive_unfolded _ _ _ H x2); intro H0. -auto. +autobatch. case (ap_cotransitive_unfolded _ _ _ H0 y2); intro H1. -auto. +autobatch. right; right. apply ap_symmetric_unfolded. @@ -809,7 +809,7 @@ definition id_un_op : \forall S:CSetoid. CSetoid_un_op S definition un_op_fun: \forall S:CSetoid. CSetoid_un_op S \to CSetoid_fun S S \def \lambda S.\lambda f.f. -coercion cic:/matita/algebra/CoRN/Setoid/un_op_fun.con. +coercion cic:/matita/algebra/CoRN/Setoids/un_op_fun.con. definition cs_un_op_strext : \forall S:CSetoid. \forall f: CSetoid_fun S S. fun_strext S S (csf_fun S S f) \def \lambda S:CSetoid. \lambda f : CSetoid_fun S S. csf_strext S S f. @@ -862,7 +862,7 @@ definition cs_bin_op_strext : \forall S:CSetoid. \forall f: CSetoid_bin_fun S S definition bin_op_bin_fun: \forall S:CSetoid. CSetoid_bin_op S \to CSetoid_bin_fun S S S \def \lambda S.\lambda f.f. -coercion cic:/matita/algebra/CoRN/Setoid/bin_op_bin_fun.con. +coercion cic:/matita/algebra/CoRN/Setoids/bin_op_bin_fun.con. @@ -982,7 +982,7 @@ definition outer_op_bin_fun: \forall S:CSetoid. CSetoid_outer_op S S \to CSetoid_bin_fun S S S \def \lambda S.\lambda f.f. -coercion cic:/matita/algebra/CoRN/Setoid/outer_op_bin_fun.con. +coercion cic:/matita/algebra/CoRN/Setoids/outer_op_bin_fun.con. (* begin hide Identity Coercion outer_op_bin_fun : CSetoid_outer_op >-> CSetoid_bin_fun. end hide *) @@ -1263,7 +1263,7 @@ apply f. apply n1. apply m1. apply eq_symmetric_unfolded.assumption. apply eq_symmetric_unfolded.assumption. apply H. -auto new. +autobatch new. qed. (*