src/HOL/Tools/datatype_abs_proofs.ML
changeset 20046 9c8909fc5865
parent 19233 77ca20b0ed77
child 20071 8f3e1ddb50e6
--- a/src/HOL/Tools/datatype_abs_proofs.ML	Sat Jul 08 12:54:32 2006 +0200
+++ b/src/HOL/Tools/datatype_abs_proofs.ML	Sat Jul 08 12:54:33 2006 +0200
@@ -74,12 +74,12 @@
           (split_conj_thm (cterm_instantiate insts' induct), i)) RSN (2, rev_mp))
 
       in
-        standard (Goal.prove thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
+        Goal.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
           (fn prems => EVERY
             [rtac induct' 1,
              REPEAT (rtac TrueI 1),
              REPEAT ((rtac impI 1) THEN (eresolve_tac prems 1)),
-             REPEAT (rtac TrueI 1)]))
+             REPEAT (rtac TrueI 1)])
       end;
 
     val casedist_thms = map prove_casedist_thm ((0 upto (length newTs - 1)) ~~
@@ -216,8 +216,8 @@
               THEN rewtac (mk_meta_eq choice_eq), rec_intrs),
             descr' ~~ rec_elims ~~ recTs ~~ rec_result_Ts);
 
-      in split_conj_thm (standard (Goal.prove thy1 [] []
-        (HOLogic.mk_Trueprop (mk_conj rec_unique_ts)) (K tac)))
+      in split_conj_thm (Goal.prove_global thy1 [] []
+        (HOLogic.mk_Trueprop (mk_conj rec_unique_ts)) (K tac))
       end;
 
     val rec_total_thms = map (fn r => r RS theI') rec_unique_thms;
@@ -250,13 +250,13 @@
 
     val _ = message "Proving characteristic theorems for primrec combinators ..."
 
-    val rec_thms = map (fn t => standard (Goal.prove thy2 [] [] t
+    val rec_thms = map (fn t => Goal.prove_global thy2 [] [] t
       (fn _ => EVERY
         [rewrite_goals_tac reccomb_defs,
          rtac the1_equality 1,
          resolve_tac rec_unique_thms 1,
          resolve_tac rec_intrs 1,
-         REPEAT (rtac allI 1 ORELSE resolve_tac rec_total_thms 1)])))
+         REPEAT (rtac allI 1 ORELSE resolve_tac rec_total_thms 1)]))
            (DatatypeProp.make_primrecs new_type_names descr sorts thy2)
 
   in
@@ -329,8 +329,8 @@
         end) (([], thy1), (hd descr) ~~ newTs ~~ case_names ~~
           (Library.take (length newTs, reccomb_names)));
 
-    val case_thms = map (map (fn t => standard (Goal.prove thy2 [] [] t
-      (fn _ => EVERY [rewrite_goals_tac (case_defs @ map mk_meta_eq primrec_thms), rtac refl 1]))))
+    val case_thms = map (map (fn t => Goal.prove_global thy2 [] [] t
+      (fn _ => EVERY [rewrite_goals_tac (case_defs @ map mk_meta_eq primrec_thms), rtac refl 1])))
           (DatatypeProp.make_cases new_type_names descr sorts thy2)
 
   in
@@ -362,8 +362,8 @@
         val tacf = K (EVERY [rtac exhaustion' 1, ALLGOALS (asm_simp_tac
           (HOL_ss addsimps (dist_rewrites' @ inject @ case_thms')))])
       in
-        (standard (Goal.prove thy [] [] t1 tacf),
-         standard (Goal.prove thy [] [] t2 tacf))
+        (Goal.prove_global thy [] [] t1 tacf,
+         Goal.prove_global thy [] [] t2 tacf)
       end;
 
     val split_thm_pairs = map prove_split_thms
@@ -432,8 +432,8 @@
 
     val rewrites = size_def_thms @ map mk_meta_eq primrec_thms;
 
-    val size_thms = map (fn t => standard (Goal.prove thy' [] [] t
-      (fn _ => EVERY [rewrite_goals_tac rewrites, rtac refl 1])))
+    val size_thms = map (fn t => Goal.prove_global thy' [] [] t
+      (fn _ => EVERY [rewrite_goals_tac rewrites, rtac refl 1]))
         (DatatypeProp.make_size descr sorts thy')
 
   in
@@ -447,8 +447,8 @@
 fun prove_weak_case_congs new_type_names descr sorts thy =
   let
     fun prove_weak_case_cong t =
-       standard (Goal.prove thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
-         (fn prems => EVERY [rtac ((hd prems) RS arg_cong) 1]))
+       Goal.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
+         (fn prems => EVERY [rtac ((hd prems) RS arg_cong) 1])
 
     val weak_case_congs = map prove_weak_case_cong (DatatypeProp.make_weak_case_congs
       new_type_names descr sorts thy)
@@ -468,10 +468,10 @@
               hyp_subst_tac i, REPEAT (rtac exI i), rtac refl i]
           | tac i n = rtac disjI2 i THEN tac i (n - 1)
       in 
-        standard (Goal.prove thy [] [] t (fn _ =>
+        Goal.prove_global thy [] [] t (fn _ =>
           EVERY [rtac allI 1,
            exh_tac (K exhaustion) 1,
-           ALLGOALS (fn i => tac i (i-1))]))
+           ALLGOALS (fn i => tac i (i-1))])
       end;
 
     val nchotomys =
@@ -490,14 +490,14 @@
         val nchotomy'' = cterm_instantiate
           [(cert (hd (add_term_vars (concl_of nchotomy', []))), cert Ma)] nchotomy'
       in
-        standard (Goal.prove thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
+        Goal.prove_global thy [] (Logic.strip_imp_prems t) (Logic.strip_imp_concl t)
           (fn prems => 
             let val simplify = asm_simp_tac (HOL_ss addsimps (prems @ case_rewrites))
             in EVERY [simp_tac (HOL_ss addsimps [hd prems]) 1,
                 cut_facts_tac [nchotomy''] 1,
                 REPEAT (etac disjE 1 THEN REPEAT (etac exE 1) THEN simplify 1),
                 REPEAT (etac exE 1) THEN simplify 1 (* Get last disjunct *)]
-            end))
+            end)
       end;
 
     val case_congs = map prove_case_cong (DatatypeProp.make_case_congs