]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/software/matita/library/decidable_kit/decidable.ma
...
[helm.git] / helm / software / matita / library / decidable_kit / decidable.ma
index dff89f8d35cc5d5d46a8539a9f8980b57fe2f70f..b2b21ccd26aa50e693b4b3e88cbe424076cc03a5 100644 (file)
@@ -12,8 +12,6 @@
 (*                                                                        *)
 (**************************************************************************)
 
-set "baseuri" "cic:/matita/decidable_kit/decidable/".
-
 (* classical connectives for decidable properties *)
 
 include "decidable_kit/streicher.ma".
@@ -58,7 +56,7 @@ qed.
 
 lemma prove_reflect : ∀P:Prop.∀b:bool.
   (b = true → P) → (b = false → ¬P) → reflect P b.
-intros 2 (P b); cases b; intros; [left|right] autobatch.
+intros 2 (P b); cases b; intros; [left|right] [autobatch.|autobatch;]
 qed.   
   
 (* ### standard connectives/relations with reflection predicate ### *)
@@ -139,7 +137,7 @@ qed.
 
 lemma leb_eqb : ∀n,m. orb (eqb n m) (leb (S n) m) = leb n m.
 intros (n m); apply bool_to_eq; split; intros (H);
-[1:cases (b2pT ? ? (orbP ? ?) H); [2: autobatch
+[1:cases (b2pT ? ? (orbP ? ?) H); [2: (*autobatch type;*) apply lebW; assumption; 
    rewrite > (eqb_true_to_eq ? ? H1); autobatch
 |2:cases (b2pT ? ? (lebP ? ?) H); 
    [ elim n; [reflexivity|assumption] 
@@ -170,7 +168,7 @@ intros (m n); apply bool_to_eq; split;
 [1: intros; cases (b2pT ? ? (orbP ? ?) H); [1: apply ltbW; assumption]
     rewrite > (eqb_true_to_eq ? ? H1); simplify; 
     rewrite > leb_refl; reflexivity  
-|2: generalize in match m; clear m; elim n 0;
+|2: elim n in m ⊢ % 0;
     [1: simplify; intros; cases n1; reflexivity;
     |2: intros 1 (m); elim m 0;
         [1: intros; apply (p2bT ? ? (orbP ? ?));