X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Flibrary%2Fdecidable_kit%2Feqtype.ma;h=6aa3ff59fd2bf57600a558665606f7aa20e8b626;hb=a88be1ca42c0969dbab9a5c76240f5931df876d9;hp=7b0f94ed49e2ef89ce94ad64cac4ac65d6938940;hpb=680039d60c1d69521f84580ee0069cb2d6ff56ba;p=helm.git diff --git a/helm/software/matita/library/decidable_kit/eqtype.ma b/helm/software/matita/library/decidable_kit/eqtype.ma index 7b0f94ed4..6aa3ff59f 100644 --- a/helm/software/matita/library/decidable_kit/eqtype.ma +++ b/helm/software/matita/library/decidable_kit/eqtype.ma @@ -12,8 +12,6 @@ (* *) (**************************************************************************) -set "baseuri" "cic:/matita/decidable_kit/eqtype/". - include "decidable_kit/decidable.ma". include "datatypes/constructors.ma". @@ -52,7 +50,7 @@ definition nat_eqType : eqType ≝ mk_eqType ? ? eqbP. (* XXX le coercion nel cast non vengono inserite *) definition nat_canonical_eqType : nat → nat_eqType := λn : nat.(n : sort nat_eqType). -coercion cic:/matita/decidable_kit/eqtype/nat_canonical_eqType.con. +coercion nat_canonical_eqType. definition bcmp ≝ λx,y:bool. match x with [ true ⇒ y | false => notb y ]. @@ -70,7 +68,7 @@ qed. definition bool_eqType : eqType ≝ mk_eqType ? ? bcmpP. definition bool_canonical_eqType : bool → bool_eqType := λb : bool.(b : sort bool_eqType). -coercion cic:/matita/decidable_kit/eqtype/bool_canonical_eqType.con. +coercion bool_canonical_eqType. (* ### subtype of an eqType ### *) @@ -83,30 +81,18 @@ notation "{x, h}" right associative with precedence 80 for @{'sig $x $h}. -interpretation "sub_eqType" 'sig x h = - (cic:/matita/decidable_kit/eqtype/sigma.ind#xpointer(1/1/1) _ _ x h). +interpretation "sub_eqType" 'sig x h = (mk_sigma ? ? x h). (* restricting an eqType gives an eqType *) lemma sigma_eq_dec : ∀d:eqType.∀p.∀x,y:sigma d p. eq_compatible ? x y (cmp ? (sval ? ? x) (sval ? ?y)). intros (d p x y); -generalize in match (refl_eq ? (cmp d (sval d p x) (sval d p y))); -generalize in match (cmp d (sval d p x) (sval d p y)) in ⊢ (? ? ? % → %); intros 1 (b); -cases b; cases x (s ps); cases y (t pt); simplify; intros (Hcmp); -[ constructor 1; - generalize in match (eqP d s t); intros (Est); - cases Est (H); clear Est; - [ generalize in match ps; - rewrite > H; intros (pt'); - rewrite < (pirrel bool (p t) true pt pt' (eqType_decidable bool_eqType)); - reflexivity; - | cases (H (b2pT ? ? (eqP d s t) Hcmp)) - ] -| constructor 2; unfold Not; intros (H); - (* XXX destruct H; *) - change in Hcmp with (cmp d (match {?,ps} with [(mk_sigma s _)⇒s]) t = false); - rewrite > H in Hcmp; simplify in Hcmp; rewrite > cmp_refl in Hcmp; destruct Hcmp; -] +cases (eqP d (sval ? ? x) (sval ? ? y)); generalize in match H; clear H; +cases x (s ps); cases y (t pt); simplify; intros (Est); +[1: constructor 1; generalize in match ps; rewrite > Est; intros (pt'); + rewrite < (pirrel ? ? ? pt pt' (eqType_decidable bool_eqType)); reflexivity; +|2: constructor 2; unfold Not; intros (H); destruct H; + cases (Est); reflexivity] qed. definition sub_eqType ≝ λd : eqType.λp. mk_eqType ? ? (sigma_eq_dec d p). @@ -138,33 +124,30 @@ generalize in match (refl_eq ? (ocmp ? a b)); generalize in match (ocmp ? a b) in ⊢ (? ? ? % → %); intros 1 (c); cases c; intros (H); [ apply reflect_true | apply reflect_false ]; generalize in match H; clear H; -cases a; cases b; -[1: intros; reflexivity; -|2,3: intros (H); destruct H; -|4: intros (H); simplify in H; rewrite > (b2pT ? ? (eqP d ? ?) H); reflexivity; - (* se faccio la rewrite diretta non tipa *) -|5: intros (H H1); destruct H -|6,7: unfold Not; intros (H H1); destruct H1 -|8: unfold Not; simplify; intros (H H1); - (* ancora injection non va *) - cut (s = s1); [ rewrite > Hcut in H; rewrite > cmp_refl in H; destruct H;]. - change with (match (Some d s) with - [ None ⇒ s | (Some s) ⇒ s] = s1); rewrite > H1; - simplify; reflexivity; -] qed. +cases a; cases b; simplify; intros (H); +[1: reflexivity; +|2,3,5: destruct H; +|4: rewrite > (b2pT ? ? (eqP d ? ?) H); reflexivity; +|6,7: unfold Not; intros (H1); destruct H1 +|8: unfold Not; intros (H1); destruct H1;rewrite > cmp_refl in H; destruct H;] +qed. definition option_eqType : eqType → eqType ≝ λd:eqType.mk_eqType ? ? (ocmpP d). definition option_canonical_eqType : ∀d:eqType.d → option_eqType d ≝ λd:eqType.λx:d.(Some ? x : sort (option_eqType d)). -coercion cic:/matita/decidable_kit/eqtype/option_canonical_eqType.con. +coercion option_canonical_eqType. (* belle le coercions! *) definition test_canonical_option_eqType ≝ (eq (option_eqType nat_eqType) O (S O)). -(* OUT OF PLACE *) -lemma eqbC : ∀x,y:nat. eqb x y = eqb y x. -intros (x y); apply bool_to_eq; split; intros (H); -rewrite > (b2pT ? ? (eqbP ? ?) H); rewrite > (cmp_refl nat_eqType); -reflexivity; +lemma cmpP : ∀d:eqType.∀x,y:d.∀P:bool → Prop. + (x=y → P true) → (cmp d x y = false → P false) → P (cmp d x y). +intros; cases (eqP ? x y); [2:apply H1; apply (p2bF ? ? (eqP d ? ?))] autobatch; +qed. + +lemma cmpC : ∀d:eqType.∀x,y:d. cmp d x y = cmp d y x. +intros; apply (cmpP ? x y); apply (cmpP ? y x); [1,4: intros; reflexivity] +[intros (H2 H1) | intros (H1 H2)] rewrite > H1 in H2; rewrite > cmp_refl in H2; +destruct H2. qed.