src/Pure/tactical.ML
changeset 32231 95b8afcbb0ed
parent 32187 cca43ca13f4f
child 32738 15bb09ca0378
--- a/src/Pure/tactical.ML	Mon Jul 27 17:36:30 2009 +0200
+++ b/src/Pure/tactical.ML	Mon Jul 27 20:45:40 2009 +0200
@@ -60,9 +60,6 @@
   val CHANGED_GOAL: (int -> tactic) -> int -> tactic
   val THEN_ALL_NEW: (int -> tactic) * (int -> tactic) -> int -> tactic
   val REPEAT_ALL_NEW: (int -> tactic) -> int -> tactic
-  val strip_context: term -> (string * typ) list * term list * term
-  val metahyps_thms: int -> thm -> thm list option
-  val METAHYPS: (thm list -> tactic) -> int -> tactic
   val PRIMSEQ: (thm -> thm Seq.seq) -> tactic
   val PRIMITIVE: (thm -> thm) -> tactic
   val SINGLE: tactic -> thm -> thm option
@@ -366,143 +363,6 @@
 fun REPEAT_ALL_NEW tac =
   tac THEN_ALL_NEW (TRY o (fn i => REPEAT_ALL_NEW tac i));
 
-
-(*Strips assumptions in goal yielding  ( [x1,...,xm], [H1,...,Hn], B )
-    H1,...,Hn are the hypotheses;  x1...xm are variants of the parameters.
-  Main difference from strip_assums concerns parameters:
-    it replaces the bound variables by free variables.  *)
-fun strip_context_aux (params, Hs, Const("==>", _) $ H $ B) =
-        strip_context_aux (params, H::Hs, B)
-  | strip_context_aux (params, Hs, Const("all",_)$Abs(a,T,t)) =
-        let val (b,u) = Syntax.variant_abs(a,T,t)
-        in  strip_context_aux ((b,T)::params, Hs, u)  end
-  | strip_context_aux (params, Hs, B) = (rev params, rev Hs, B);
-
-fun strip_context A = strip_context_aux ([],[],A);
-
-
-(**** METAHYPS -- tactical for using hypotheses as meta-level assumptions
-       METAHYPS (fn prems => tac prems) i
-
-converts subgoal i, of the form !!x1...xm. [| A1;...;An] ==> A into a new
-proof state A==>A, supplying A1,...,An as meta-level assumptions (in
-"prems").  The parameters x1,...,xm become free variables.  If the
-resulting proof state is [| B1;...;Bk] ==> C (possibly assuming A1,...,An)
-then it is lifted back into the original context, yielding k subgoals.
-
-Replaces unknowns in the context by Frees having the prefix METAHYP_
-New unknowns in [| B1;...;Bk] ==> C are lifted over x1,...,xm.
-DOES NOT HANDLE TYPE UNKNOWNS.
-****)
-
-local
-
-  (*Left-to-right replacements: ctpairs = [...,(vi,ti),...].
-    Instantiates distinct free variables by terms of same type.*)
-  fun free_instantiate ctpairs =
-    forall_elim_list (map snd ctpairs) o forall_intr_list (map fst ctpairs);
-
-  fun free_of s ((a, i), T) =
-    Free (s ^ (case i of 0 => a | _ => a ^ "_" ^ string_of_int i), T)
-
-  fun mk_inst v = (Var v, free_of "METAHYP1_" v)
-in
-
-(*Common code for METAHYPS and metahyps_thms*)
-fun metahyps_split_prem prem =
-  let (*find all vars in the hyps -- should find tvars also!*)
-      val hyps_vars = fold Term.add_vars (Logic.strip_assums_hyp prem) []
-      val insts = map mk_inst hyps_vars
-      (*replace the hyps_vars by Frees*)
-      val prem' = subst_atomic insts prem
-      val (params,hyps,concl) = strip_context prem'
-  in (insts,params,hyps,concl)  end;
-
-fun metahyps_aux_tac tacf (prem,gno) state =
-  let val (insts,params,hyps,concl) = metahyps_split_prem prem
-      val maxidx = Thm.maxidx_of state
-      val cterm = Thm.cterm_of (Thm.theory_of_thm state)
-      val chyps = map cterm hyps
-      val hypths = map assume chyps
-      val subprems = map (Thm.forall_elim_vars 0) hypths
-      val fparams = map Free params
-      val cparams = map cterm fparams
-      fun swap_ctpair (t,u) = (cterm u, cterm t)
-      (*Subgoal variables: make Free; lift type over params*)
-      fun mk_subgoal_inst concl_vars (v, T) =
-          if member (op =) concl_vars (v, T)
-          then ((v, T), true, free_of "METAHYP2_" (v, T))
-          else ((v, T), false, free_of "METAHYP2_" (v, map #2 params ---> T))
-      (*Instantiate subgoal vars by Free applied to params*)
-      fun mk_ctpair (v, in_concl, u) =
-          if in_concl then (cterm (Var v), cterm u)
-          else (cterm (Var v), cterm (list_comb (u, fparams)))
-      (*Restore Vars with higher type and index*)
-      fun mk_subgoal_swap_ctpair (((a, i), T), in_concl, u as Free (_, U)) =
-          if in_concl then (cterm u, cterm (Var ((a, i), T)))
-          else (cterm u, cterm (Var ((a, i + maxidx), U)))
-      (*Embed B in the original context of params and hyps*)
-      fun embed B = list_all_free (params, Logic.list_implies (hyps, B))
-      (*Strip the context using elimination rules*)
-      fun elim Bhyp = implies_elim_list (forall_elim_list cparams Bhyp) hypths
-      (*A form of lifting that discharges assumptions.*)
-      fun relift st =
-        let val prop = Thm.prop_of st
-            val subgoal_vars = (*Vars introduced in the subgoals*)
-              fold Term.add_vars (Logic.strip_imp_prems prop) []
-            and concl_vars = Term.add_vars (Logic.strip_imp_concl prop) []
-            val subgoal_insts = map (mk_subgoal_inst concl_vars) subgoal_vars
-            val st' = Thm.instantiate ([], map mk_ctpair subgoal_insts) st
-            val emBs = map (cterm o embed) (prems_of st')
-            val Cth  = implies_elim_list st' (map (elim o assume) emBs)
-        in  (*restore the unknowns to the hypotheses*)
-            free_instantiate (map swap_ctpair insts @
-                              map mk_subgoal_swap_ctpair subgoal_insts)
-                (*discharge assumptions from state in same order*)
-                (implies_intr_list emBs
-                  (forall_intr_list cparams (implies_intr_list chyps Cth)))
-        end
-      (*function to replace the current subgoal*)
-      fun next st = Thm.bicompose false (false, relift st, nprems_of st) gno state
-  in Seq.maps next (tacf subprems (trivial (cterm concl))) end;
-
-end;
-
-(*Returns the theorem list that METAHYPS would supply to its tactic*)
-fun metahyps_thms i state =
-  let val prem = Logic.nth_prem (i, Thm.prop_of state)
-      and cterm = cterm_of (Thm.theory_of_thm state)
-      val (_,_,hyps,_) = metahyps_split_prem prem
-  in SOME (map (Thm.forall_elim_vars 0 o Thm.assume o cterm) hyps) end
-  handle TERM ("nth_prem", [A]) => NONE;
-
-local
-
-fun print_vars_terms thy (n,thm) =
-  let
-    fun typed ty = " has type: " ^ Syntax.string_of_typ_global thy ty;
-    fun find_vars thy (Const (c, ty)) =
-          if null (Term.add_tvarsT ty []) then I
-          else insert (op =) (c ^ typed ty)
-      | find_vars thy (Var (xi, ty)) = insert (op =) (Term.string_of_vname xi ^ typed ty)
-      | find_vars _ (Free _) = I
-      | find_vars _ (Bound _) = I
-      | find_vars thy (Abs (_, _, t)) = find_vars thy t
-      | find_vars thy (t1 $ t2) =
-          find_vars thy t1 #> find_vars thy t1;
-    val prem = Logic.nth_prem (n, Thm.prop_of thm)
-    val tms = find_vars thy prem []
-  in
-    (warning "Found schematic vars in assumptions:"; warning (cat_lines tms))
-  end;
-
-in
-
-fun METAHYPS tacf n thm = SUBGOAL (metahyps_aux_tac tacf) n thm
-  handle THM("assume: variables",_,_) => (print_vars_terms (theory_of_thm thm) (n,thm); Seq.empty)
-
-end;
-
 (*Makes a tactic whose effect on a state is given by thmfun: thm->thm seq.*)
 fun PRIMSEQ thmfun st =  thmfun st handle THM _ => Seq.empty;