| author | haftmann | 
| Mon, 26 Oct 2009 09:03:57 +0100 | |
| changeset 33176 | d6936fd7cda8 | 
| parent 33040 | cffdb7b28498 | 
| child 35021 | c839a4c670c6 | 
| permissions | -rw-r--r-- | 
| 18120 | 1 | (* Title: Pure/old_goals.ML | 
| 2 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 3 | Copyright 1993 University of Cambridge | |
| 4 | ||
| 5 | Old-style goal stack package. The goal stack initially holds a dummy | |
| 6 | proof, and can never become empty. Each goal stack consists of a list | |
| 7 | of levels. The undo list is a list of goal stacks. Finally, there | |
| 8 | may be a stack of pending proofs. | |
| 9 | *) | |
| 10 | ||
| 32173 | 11 | signature OLD_GOALS = | 
| 18120 | 12 | sig | 
| 32231 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 13 | val strip_context: term -> (string * typ) list * term list * term | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 14 | val metahyps_thms: int -> thm -> thm list option | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 15 | val METAHYPS: (thm list -> tactic) -> int -> tactic | 
| 32173 | 16 | val simple_read_term: theory -> typ -> string -> term | 
| 17 | val read_term: theory -> string -> term | |
| 18 | val read_prop: theory -> string -> term | |
| 19 | type proof | |
| 18120 | 20 | val premises: unit -> thm list | 
| 32173 | 21 | val reset_goals: unit -> unit | 
| 32738 | 22 | val result_error_fn: (thm -> string -> thm) Unsynchronized.ref | 
| 32173 | 23 | val print_sign_exn: theory -> exn -> 'a | 
| 24 | val prove_goalw_cterm: thm list->cterm->(thm list->tactic list)->thm | |
| 25 | val prove_goalw_cterm_nocheck: thm list->cterm->(thm list->tactic list)->thm | |
| 26 | val prove_goalw: theory -> thm list -> string -> (thm list -> tactic list) -> thm | |
| 18120 | 27 | val prove_goal: theory -> string -> (thm list -> tactic list) -> thm | 
| 28 | val topthm: unit -> thm | |
| 29 | val result: unit -> thm | |
| 30 | val uresult: unit -> thm | |
| 31 | val getgoal: int -> term | |
| 32 | val gethyps: int -> thm list | |
| 32173 | 33 | val print_exn: exn -> 'a | 
| 34 | val filter_goal: (term*term->bool) -> thm list -> int -> thm list | |
| 18120 | 35 | val prlev: int -> unit | 
| 36 | val pr: unit -> unit | |
| 37 | val prlim: int -> unit | |
| 32173 | 38 | val goalw_cterm: thm list -> cterm -> thm list | 
| 39 | val goalw: theory -> thm list -> string -> thm list | |
| 18120 | 40 | val goal: theory -> string -> thm list | 
| 32173 | 41 | val Goalw: thm list -> string -> thm list | 
| 18120 | 42 | val Goal: string -> thm list | 
| 32173 | 43 | val simple_prove_goal_cterm: cterm->(thm list->tactic list)->thm | 
| 18120 | 44 | val by: tactic -> unit | 
| 32173 | 45 | val byev: tactic list -> unit | 
| 18120 | 46 | val back: unit -> unit | 
| 47 | val choplev: int -> unit | |
| 32173 | 48 | val chop: unit -> unit | 
| 18120 | 49 | val undo: unit -> unit | 
| 32173 | 50 | val save_proof: unit -> proof | 
| 51 | val restore_proof: proof -> thm list | |
| 52 | val push_proof: unit -> unit | |
| 53 | val pop_proof: unit -> thm list | |
| 54 | val rotate_proof: unit -> thm list | |
| 18120 | 55 | val qed: string -> unit | 
| 56 | val qed_goal: string -> theory -> string -> (thm list -> tactic list) -> unit | |
| 57 | val qed_goalw: string -> theory -> thm list -> string | |
| 58 | -> (thm list -> tactic list) -> unit | |
| 59 | val qed_spec_mp: string -> unit | |
| 60 | val qed_goal_spec_mp: string -> theory -> string -> (thm list -> tactic list) -> unit | |
| 61 | val qed_goalw_spec_mp: string -> theory -> thm list -> string | |
| 62 | -> (thm list -> tactic list) -> unit | |
| 63 | end; | |
| 64 | ||
| 65 | structure OldGoals: OLD_GOALS = | |
| 66 | struct | |
| 67 | ||
| 32231 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 68 | (**** METAHYPS -- tactical for using hypotheses as meta-level assumptions | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 69 | METAHYPS (fn prems => tac prems) i | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 70 | |
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 71 | converts subgoal i, of the form !!x1...xm. [| A1;...;An] ==> A into a new | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 72 | proof state A==>A, supplying A1,...,An as meta-level assumptions (in | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 73 | "prems"). The parameters x1,...,xm become free variables. If the | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 74 | resulting proof state is [| B1;...;Bk] ==> C (possibly assuming A1,...,An) | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 75 | then it is lifted back into the original context, yielding k subgoals. | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 76 | |
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 77 | Replaces unknowns in the context by Frees having the prefix METAHYP_ | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 78 | New unknowns in [| B1;...;Bk] ==> C are lifted over x1,...,xm. | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 79 | DOES NOT HANDLE TYPE UNKNOWNS. | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 80 | |
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 81 | |
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 82 | NOTE: This version does not observe the proof context, and thus cannot | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 83 | work reliably. See also Subgoal.SUBPROOF and Subgoal.FOCUS for | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 84 | properly localized variants of the same idea. | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 85 | ****) | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 86 | |
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 87 | (*Strips assumptions in goal yielding ( [x1,...,xm], [H1,...,Hn], B ) | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 88 | H1,...,Hn are the hypotheses; x1...xm are variants of the parameters. | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 89 | Main difference from strip_assums concerns parameters: | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 90 | it replaces the bound variables by free variables. *) | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 91 | fun strip_context_aux (params, Hs, Const ("==>", _) $ H $ B) =
 | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 92 | strip_context_aux (params, H :: Hs, B) | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 93 |   | strip_context_aux (params, Hs, Const ("all",_) $ Abs (a, T, t)) =
 | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 94 | let val (b, u) = Syntax.variant_abs (a, T, t) | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 95 | in strip_context_aux ((b, T) :: params, Hs, u) end | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 96 | | strip_context_aux (params, Hs, B) = (rev params, rev Hs, B); | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 97 | |
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 98 | fun strip_context A = strip_context_aux ([], [], A); | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 99 | |
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 100 | local | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 101 | |
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 102 | (*Left-to-right replacements: ctpairs = [...,(vi,ti),...]. | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 103 | Instantiates distinct free variables by terms of same type.*) | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 104 | fun free_instantiate ctpairs = | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 105 | forall_elim_list (map snd ctpairs) o forall_intr_list (map fst ctpairs); | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 106 | |
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 107 | fun free_of s ((a, i), T) = | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 108 | Free (s ^ (case i of 0 => a | _ => a ^ "_" ^ string_of_int i), T) | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 109 | |
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 110 | fun mk_inst v = (Var v, free_of "METAHYP1_" v) | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 111 | in | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 112 | |
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 113 | (*Common code for METAHYPS and metahyps_thms*) | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 114 | fun metahyps_split_prem prem = | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 115 | let (*find all vars in the hyps -- should find tvars also!*) | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 116 | val hyps_vars = fold Term.add_vars (Logic.strip_assums_hyp prem) [] | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 117 | val insts = map mk_inst hyps_vars | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 118 | (*replace the hyps_vars by Frees*) | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 119 | val prem' = subst_atomic insts prem | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 120 | val (params,hyps,concl) = strip_context prem' | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 121 | in (insts,params,hyps,concl) end; | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 122 | |
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 123 | fun metahyps_aux_tac tacf (prem,gno) state = | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 124 | let val (insts,params,hyps,concl) = metahyps_split_prem prem | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 125 | val maxidx = Thm.maxidx_of state | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 126 | val cterm = Thm.cterm_of (Thm.theory_of_thm state) | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 127 | val chyps = map cterm hyps | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 128 | val hypths = map assume chyps | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 129 | val subprems = map (Thm.forall_elim_vars 0) hypths | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 130 | val fparams = map Free params | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 131 | val cparams = map cterm fparams | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 132 | fun swap_ctpair (t,u) = (cterm u, cterm t) | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 133 | (*Subgoal variables: make Free; lift type over params*) | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 134 | fun mk_subgoal_inst concl_vars (v, T) = | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 135 | if member (op =) concl_vars (v, T) | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 136 | then ((v, T), true, free_of "METAHYP2_" (v, T)) | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 137 | else ((v, T), false, free_of "METAHYP2_" (v, map #2 params ---> T)) | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 138 | (*Instantiate subgoal vars by Free applied to params*) | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 139 | fun mk_ctpair (v, in_concl, u) = | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 140 | if in_concl then (cterm (Var v), cterm u) | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 141 | else (cterm (Var v), cterm (list_comb (u, fparams))) | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 142 | (*Restore Vars with higher type and index*) | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 143 | fun mk_subgoal_swap_ctpair (((a, i), T), in_concl, u as Free (_, U)) = | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 144 | if in_concl then (cterm u, cterm (Var ((a, i), T))) | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 145 | else (cterm u, cterm (Var ((a, i + maxidx), U))) | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 146 | (*Embed B in the original context of params and hyps*) | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 147 | fun embed B = list_all_free (params, Logic.list_implies (hyps, B)) | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 148 | (*Strip the context using elimination rules*) | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 149 | fun elim Bhyp = implies_elim_list (forall_elim_list cparams Bhyp) hypths | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 150 | (*A form of lifting that discharges assumptions.*) | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 151 | fun relift st = | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 152 | let val prop = Thm.prop_of st | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 153 | val subgoal_vars = (*Vars introduced in the subgoals*) | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 154 | fold Term.add_vars (Logic.strip_imp_prems prop) [] | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 155 | and concl_vars = Term.add_vars (Logic.strip_imp_concl prop) [] | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 156 | val subgoal_insts = map (mk_subgoal_inst concl_vars) subgoal_vars | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 157 | val st' = Thm.instantiate ([], map mk_ctpair subgoal_insts) st | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 158 | val emBs = map (cterm o embed) (prems_of st') | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 159 | val Cth = implies_elim_list st' (map (elim o assume) emBs) | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 160 | in (*restore the unknowns to the hypotheses*) | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 161 | free_instantiate (map swap_ctpair insts @ | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 162 | map mk_subgoal_swap_ctpair subgoal_insts) | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 163 | (*discharge assumptions from state in same order*) | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 164 | (implies_intr_list emBs | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 165 | (forall_intr_list cparams (implies_intr_list chyps Cth))) | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 166 | end | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 167 | (*function to replace the current subgoal*) | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 168 | fun next st = Thm.bicompose false (false, relift st, nprems_of st) gno state | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 169 | in Seq.maps next (tacf subprems (trivial (cterm concl))) end; | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 170 | |
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 171 | end; | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 172 | |
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 173 | (*Returns the theorem list that METAHYPS would supply to its tactic*) | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 174 | fun metahyps_thms i state = | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 175 | let val prem = Logic.nth_prem (i, Thm.prop_of state) | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 176 | and cterm = cterm_of (Thm.theory_of_thm state) | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 177 | val (_,_,hyps,_) = metahyps_split_prem prem | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 178 | in SOME (map (Thm.forall_elim_vars 0 o Thm.assume o cterm) hyps) end | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 179 |   handle TERM ("nth_prem", [A]) => NONE;
 | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 180 | |
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 181 | local | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 182 | |
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 183 | fun print_vars_terms thy (n,thm) = | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 184 | let | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 185 | fun typed ty = " has type: " ^ Syntax.string_of_typ_global thy ty; | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 186 | fun find_vars thy (Const (c, ty)) = | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 187 | if null (Term.add_tvarsT ty []) then I | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 188 | else insert (op =) (c ^ typed ty) | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 189 | | find_vars thy (Var (xi, ty)) = insert (op =) (Term.string_of_vname xi ^ typed ty) | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 190 | | find_vars _ (Free _) = I | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 191 | | find_vars _ (Bound _) = I | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 192 | | find_vars thy (Abs (_, _, t)) = find_vars thy t | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 193 | | find_vars thy (t1 $ t2) = | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 194 | find_vars thy t1 #> find_vars thy t1; | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 195 | val prem = Logic.nth_prem (n, Thm.prop_of thm) | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 196 | val tms = find_vars thy prem [] | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 197 | in | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 198 | (warning "Found schematic vars in assumptions:"; warning (cat_lines tms)) | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 199 | end; | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 200 | |
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 201 | in | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 202 | |
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 203 | fun METAHYPS tacf n thm = SUBGOAL (metahyps_aux_tac tacf) n thm | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 204 |   handle THM("assume: variables",_,_) => (print_vars_terms (theory_of_thm thm) (n,thm); Seq.empty)
 | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 205 | |
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 206 | end; | 
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 207 | |
| 
95b8afcbb0ed
moved METAHYPS to old_goals.ML (cf. SUBPROOF and FOCUS in subgoal.ML for properly localized versions of the same idea);
 wenzelm parents: 
32187diff
changeset | 208 | |
| 27256 
bcb071683184
added emulations for simple_read_term/read_term/read_prop (formerly in sign.ML);
 wenzelm parents: 
27242diff
changeset | 209 | (* old ways of reading terms *) | 
| 
bcb071683184
added emulations for simple_read_term/read_term/read_prop (formerly in sign.ML);
 wenzelm parents: 
27242diff
changeset | 210 | |
| 
bcb071683184
added emulations for simple_read_term/read_term/read_prop (formerly in sign.ML);
 wenzelm parents: 
27242diff
changeset | 211 | fun simple_read_term thy T s = | 
| 
bcb071683184
added emulations for simple_read_term/read_term/read_prop (formerly in sign.ML);
 wenzelm parents: 
27242diff
changeset | 212 | let | 
| 
bcb071683184
added emulations for simple_read_term/read_term/read_prop (formerly in sign.ML);
 wenzelm parents: 
27242diff
changeset | 213 | val ctxt = ProofContext.init thy | 
| 
bcb071683184
added emulations for simple_read_term/read_term/read_prop (formerly in sign.ML);
 wenzelm parents: 
27242diff
changeset | 214 | |> ProofContext.allow_dummies | 
| 
bcb071683184
added emulations for simple_read_term/read_term/read_prop (formerly in sign.ML);
 wenzelm parents: 
27242diff
changeset | 215 | |> ProofContext.set_mode ProofContext.mode_schematic; | 
| 
bcb071683184
added emulations for simple_read_term/read_term/read_prop (formerly in sign.ML);
 wenzelm parents: 
27242diff
changeset | 216 | val parse = if T = propT then Syntax.parse_prop else Syntax.parse_term; | 
| 
bcb071683184
added emulations for simple_read_term/read_term/read_prop (formerly in sign.ML);
 wenzelm parents: 
27242diff
changeset | 217 | in parse ctxt s |> TypeInfer.constrain T |> Syntax.check_term ctxt end; | 
| 
bcb071683184
added emulations for simple_read_term/read_term/read_prop (formerly in sign.ML);
 wenzelm parents: 
27242diff
changeset | 218 | |
| 
bcb071683184
added emulations for simple_read_term/read_term/read_prop (formerly in sign.ML);
 wenzelm parents: 
27242diff
changeset | 219 | fun read_term thy = simple_read_term thy dummyT; | 
| 
bcb071683184
added emulations for simple_read_term/read_term/read_prop (formerly in sign.ML);
 wenzelm parents: 
27242diff
changeset | 220 | fun read_prop thy = simple_read_term thy propT; | 
| 
bcb071683184
added emulations for simple_read_term/read_term/read_prop (formerly in sign.ML);
 wenzelm parents: 
27242diff
changeset | 221 | |
| 
bcb071683184
added emulations for simple_read_term/read_term/read_prop (formerly in sign.ML);
 wenzelm parents: 
27242diff
changeset | 222 | |
| 
bcb071683184
added emulations for simple_read_term/read_term/read_prop (formerly in sign.ML);
 wenzelm parents: 
27242diff
changeset | 223 | |
| 18120 | 224 | (*** Goal package ***) | 
| 225 | ||
| 226 | (*Each level of goal stack includes a proof state and alternative states, | |
| 227 | the output of the tactic applied to the preceeding level. *) | |
| 228 | type gstack = (thm * thm Seq.seq) list; | |
| 229 | ||
| 230 | datatype proof = Proof of gstack list * thm list * (bool*thm->thm); | |
| 231 | ||
| 232 | ||
| 233 | (*** References ***) | |
| 234 | ||
| 235 | (*Current assumption list -- set by "goal".*) | |
| 32738 | 236 | val curr_prems = Unsynchronized.ref([] : thm list); | 
| 18120 | 237 | |
| 238 | (*Return assumption list -- useful if you didn't save "goal"'s result. *) | |
| 239 | fun premises() = !curr_prems; | |
| 240 | ||
| 241 | (*Current result maker -- set by "goal", used by "result". *) | |
| 242 | fun init_mkresult _ = error "No goal has been supplied in subgoal module"; | |
| 32738 | 243 | val curr_mkresult = Unsynchronized.ref (init_mkresult: bool*thm->thm); | 
| 18120 | 244 | |
| 245 | (*List of previous goal stacks, for the undo operation. Set by setstate. | |
| 246 | A list of lists!*) | |
| 32738 | 247 | val undo_list = Unsynchronized.ref([[(asm_rl, Seq.empty)]] : gstack list); | 
| 18120 | 248 | |
| 249 | (* Stack of proof attempts *) | |
| 32738 | 250 | val proofstack = Unsynchronized.ref([]: proof list); | 
| 18120 | 251 | |
| 252 | (*reset all refs*) | |
| 253 | fun reset_goals () = | |
| 254 | (curr_prems := []; curr_mkresult := init_mkresult; | |
| 26429 | 255 | undo_list := [[(asm_rl, Seq.empty)]]); | 
| 18120 | 256 | |
| 257 | ||
| 258 | (*** Setting up goal-directed proof ***) | |
| 259 | ||
| 260 | (*Generates the list of new theories when the proof state's theory changes*) | |
| 261 | fun thy_error (thy,thy') = | |
| 33040 | 262 | let val names = subtract (op =) (Context.display_names thy) (Context.display_names thy') | 
| 18120 | 263 | in case names of | 
| 264 | [name] => "\nNew theory: " ^ name | |
| 265 | | _ => "\nNew theories: " ^ space_implode ", " names | |
| 266 | end; | |
| 267 | ||
| 268 | (*Default action is to print an error message; could be suppressed for | |
| 269 | special applications.*) | |
| 270 | fun result_error_default state msg : thm = | |
| 32145 
220c9e439d39
clarified pretty_goals, pretty_thm_aux: plain context;
 wenzelm parents: 
32091diff
changeset | 271 | Pretty.str "Bad final proof state:" :: | 
| 32187 | 272 | Goal_Display.pretty_goals_without_context (!goals_limit) state @ | 
| 18120 | 273 | [Pretty.str msg, Pretty.str "Proof failed!"] |> Pretty.chunks |> Pretty.string_of |> error; | 
| 274 | ||
| 32738 | 275 | val result_error_fn = Unsynchronized.ref result_error_default; | 
| 18120 | 276 | |
| 277 | ||
| 278 | (*Common treatment of "goal" and "prove_goal": | |
| 279 | Return assumptions, initial proof state, and function to make result. | |
| 280 | "atomic" indicates if the goal should be wrapped up in the function | |
| 281 | "Goal::prop=>prop" to avoid assumptions being returned separately. | |
| 282 | *) | |
| 283 | fun prepare_proof atomic rths chorn = | |
| 284 | let | |
| 26291 | 285 | val _ = legacy_feature "Old goal command"; | 
| 26626 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
26429diff
changeset | 286 | val thy = Thm.theory_of_cterm chorn; | 
| 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
26429diff
changeset | 287 | val horn = Thm.term_of chorn; | 
| 18120 | 288 | val _ = Term.no_dummy_patterns horn handle TERM (msg, _) => error msg; | 
| 289 | val (As, B) = Logic.strip_horn horn; | |
| 290 | val atoms = atomic andalso | |
| 27332 | 291 | forall (fn t => not (can Logic.dest_implies t orelse Logic.is_all t)) As; | 
| 18120 | 292 | val (As,B) = if atoms then ([],horn) else (As,B); | 
| 293 | val cAs = map (cterm_of thy) As; | |
| 26653 | 294 | val prems = map (rewrite_rule rths o Thm.forall_elim_vars 0 o Thm.assume) cAs; | 
| 18120 | 295 | val cB = cterm_of thy B; | 
| 296 | val st0 = let val st = Goal.init cB |> fold Thm.weaken cAs | |
| 297 | in rewrite_goals_rule rths st end | |
| 298 | (*discharges assumptions from state in the order they appear in goal; | |
| 299 | checks (if requested) that resulting theorem is equivalent to goal. *) | |
| 300 | fun mkresult (check,state) = | |
| 301 | let val state = Seq.hd (flexflex_rule state) | |
| 302 | handle THM _ => state (*smash flexflex pairs*) | |
| 303 | val ngoals = nprems_of state | |
| 304 | val ath = implies_intr_list cAs state | |
| 305 | val th = Goal.conclude ath | |
| 26626 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
26429diff
changeset | 306 | val thy' = Thm.theory_of_thm th | 
| 
c6231d64d264
rep_cterm/rep_thm: no longer dereference theory_ref;
 wenzelm parents: 
26429diff
changeset | 307 |             val {hyps, prop, ...} = Thm.rep_thm th
 | 
| 32957 | 308 | val final_th = Drule.standard th | 
| 18120 | 309 | in if not check then final_th | 
| 26665 | 310 | else if not (Theory.eq_thy(thy,thy')) then !result_error_fn state | 
| 18120 | 311 |                 ("Theory of proof state has changed!" ^
 | 
| 312 | thy_error (thy,thy')) | |
| 313 | else if ngoals>0 then !result_error_fn state | |
| 314 | (string_of_int ngoals ^ " unsolved goals!") | |
| 315 | else if not (null hyps) then !result_error_fn state | |
| 316 |                 ("Additional hypotheses:\n" ^
 | |
| 26939 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 wenzelm parents: 
26928diff
changeset | 317 | cat_lines (map (Syntax.string_of_term_global thy) hyps)) | 
| 18120 | 318 | else if Pattern.matches thy | 
| 319 | (Envir.beta_norm (term_of chorn), Envir.beta_norm prop) | |
| 320 | then final_th | |
| 321 | else !result_error_fn state "proved a different theorem" | |
| 322 | end | |
| 323 | in | |
| 26665 | 324 | if Theory.eq_thy(thy, Thm.theory_of_thm st0) | 
| 18120 | 325 | then (prems, st0, mkresult) | 
| 326 |      else error ("Definitions would change the proof state's theory" ^
 | |
| 327 | thy_error (thy, Thm.theory_of_thm st0)) | |
| 328 | end | |
| 329 |   handle THM(s,_,_) => error("prepare_proof: exception THM was raised!\n" ^ s);
 | |
| 330 | ||
| 331 | (*Prints exceptions readably to users*) | |
| 332 | fun print_sign_exn_unit thy e = | |
| 333 | case e of | |
| 334 | THM (msg,i,thms) => | |
| 335 |          (writeln ("Exception THM " ^ string_of_int i ^ " raised:\n" ^ msg);
 | |
| 32091 
30e2ffbba718
proper context for Display.pretty_thm etc. or old-style versions Display.pretty_thm_global, Display.pretty_thm_without_context etc.;
 wenzelm parents: 
32089diff
changeset | 336 | List.app (writeln o Display.string_of_thm_global thy) thms) | 
| 18120 | 337 | | THEORY (msg,thys) => | 
| 338 |          (writeln ("Exception THEORY raised:\n" ^ msg);
 | |
| 339 | List.app (writeln o Context.str_of_thy) thys) | |
| 340 | | TERM (msg,ts) => | |
| 341 |          (writeln ("Exception TERM raised:\n" ^ msg);
 | |
| 26939 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 wenzelm parents: 
26928diff
changeset | 342 | List.app (writeln o Syntax.string_of_term_global thy) ts) | 
| 18120 | 343 | | TYPE (msg,Ts,ts) => | 
| 344 |          (writeln ("Exception TYPE raised:\n" ^ msg);
 | |
| 26939 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 wenzelm parents: 
26928diff
changeset | 345 | List.app (writeln o Syntax.string_of_typ_global thy) Ts; | 
| 
1035c89b4c02
moved global pretty/string_of functions from Sign to Syntax;
 wenzelm parents: 
26928diff
changeset | 346 | List.app (writeln o Syntax.string_of_term_global thy) ts) | 
| 18120 | 347 | | e => raise e; | 
| 348 | ||
| 349 | (*Prints an exception, then fails*) | |
| 18678 | 350 | fun print_sign_exn thy e = (print_sign_exn_unit thy e; raise ERROR ""); | 
| 18120 | 351 | |
| 352 | (** the prove_goal.... commands | |
| 353 | Prove theorem using the listed tactics; check it has the specified form. | |
| 354 | Augment theory with all type assignments of goal. | |
| 355 | Syntax is similar to "goal" command for easy keyboard use. **) | |
| 356 | ||
| 357 | (*Version taking the goal as a cterm*) | |
| 358 | fun prove_goalw_cterm_general check rths chorn tacsf = | |
| 359 | let val (prems, st0, mkresult) = prepare_proof false rths chorn | |
| 360 | val tac = EVERY (tacsf prems) | |
| 361 | fun statef() = | |
| 362 | (case Seq.pull (tac st0) of | |
| 363 | SOME(st,_) => st | |
| 364 |              | _ => error ("prove_goal: tactic failed"))
 | |
| 32432 
64f30bdd3ba1
modernized messages -- eliminated ctyp/cterm operations;
 wenzelm parents: 
32231diff
changeset | 365 | in mkresult (check, cond_timeit (!Output.timing) "" statef) end; | 
| 18120 | 366 | |
| 367 | (*Two variants: one checking the result, one not. | |
| 368 | Neither prints runtime messages: they are for internal packages.*) | |
| 369 | fun prove_goalw_cterm rths chorn = | |
| 32966 
5b21661fe618
indicate CRITICAL nature of various setmp combinators;
 wenzelm parents: 
32957diff
changeset | 370 | setmp_CRITICAL Output.timing false (prove_goalw_cterm_general true rths chorn) | 
| 18120 | 371 | and prove_goalw_cterm_nocheck rths chorn = | 
| 32966 
5b21661fe618
indicate CRITICAL nature of various setmp combinators;
 wenzelm parents: 
32957diff
changeset | 372 | setmp_CRITICAL Output.timing false (prove_goalw_cterm_general false rths chorn); | 
| 18120 | 373 | |
| 374 | ||
| 375 | (*Version taking the goal as a string*) | |
| 376 | fun prove_goalw thy rths agoal tacsf = | |
| 27256 
bcb071683184
added emulations for simple_read_term/read_term/read_prop (formerly in sign.ML);
 wenzelm parents: 
27242diff
changeset | 377 | let val chorn = cterm_of thy (read_prop thy agoal) | 
| 18120 | 378 | in prove_goalw_cterm_general true rths chorn tacsf end | 
| 27256 
bcb071683184
added emulations for simple_read_term/read_term/read_prop (formerly in sign.ML);
 wenzelm parents: 
27242diff
changeset | 379 | handle ERROR msg => cat_error msg (*from read_prop?*) | 
| 18120 | 380 |                 ("The error(s) above occurred for " ^ quote agoal);
 | 
| 381 | ||
| 382 | (*String version with no meta-rewrite-rules*) | |
| 383 | fun prove_goal thy = prove_goalw thy []; | |
| 384 | ||
| 385 | ||
| 386 | ||
| 387 | (*** Commands etc ***) | |
| 388 | ||
| 389 | (*Return the current goal stack, if any, from undo_list*) | |
| 390 | fun getstate() : gstack = case !undo_list of | |
| 391 | [] => error"No current state in subgoal module" | |
| 392 | | x::_ => x; | |
| 393 | ||
| 394 | (*Pops the given goal stack*) | |
| 395 | fun pop [] = error"Cannot go back past the beginning of the proof!" | |
| 396 | | pop (pair::pairs) = (pair,pairs); | |
| 397 | ||
| 398 | ||
| 23635 | 399 | (* Print a level of the goal stack *) | 
| 18120 | 400 | |
| 401 | fun print_top ((th, _), pairs) = | |
| 23635 | 402 | let | 
| 403 | val n = length pairs; | |
| 404 | val m = (! goals_limit); | |
| 405 | val ngoals = nprems_of th; | |
| 406 | in | |
| 407 |     [Pretty.str ("Level " ^ string_of_int n ^
 | |
| 408 |       (if ngoals > 0 then " (" ^ string_of_int ngoals ^ " subgoal" ^
 | |
| 409 | (if ngoals <> 1 then "s" else "") ^ ")" | |
| 410 | else ""))] @ | |
| 32187 | 411 | Goal_Display.pretty_goals_without_context m th | 
| 23635 | 412 | end |> Pretty.chunks |> Pretty.writeln; | 
| 18120 | 413 | |
| 414 | (*Printing can raise exceptions, so the assignment occurs last. | |
| 415 | Can do setstate[(st,Seq.empty)] to set st as the state. *) | |
| 416 | fun setstate newgoals = | |
| 417 | (print_top (pop newgoals); undo_list := newgoals :: !undo_list); | |
| 418 | ||
| 419 | (*Given a proof state transformation, return a command that updates | |
| 420 | the goal stack*) | |
| 421 | fun make_command com = setstate (com (pop (getstate()))); | |
| 422 | ||
| 423 | (*Apply a function on proof states to the current goal stack*) | |
| 424 | fun apply_fun f = f (pop(getstate())); | |
| 425 | ||
| 426 | (*Return the top theorem, representing the proof state*) | |
| 427 | fun topthm () = apply_fun (fn ((th,_), _) => th); | |
| 428 | ||
| 429 | (*Return the final result. *) | |
| 430 | fun result () = !curr_mkresult (true, topthm()); | |
| 431 | ||
| 432 | (*Return the result UNCHECKED that it equals the goal -- for synthesis, | |
| 433 | answer extraction, or other instantiation of Vars *) | |
| 434 | fun uresult () = !curr_mkresult (false, topthm()); | |
| 435 | ||
| 436 | (*Get subgoal i from goal stack*) | |
| 437 | fun getgoal i = Logic.get_goal (prop_of (topthm())) i; | |
| 438 | ||
| 439 | (*Return subgoal i's hypotheses as meta-level assumptions. | |
| 440 | For debugging uses of METAHYPS*) | |
| 441 | local exception GETHYPS of thm list | |
| 442 | in | |
| 443 | fun gethyps i = | |
| 444 | (METAHYPS (fn hyps => raise (GETHYPS hyps)) i (topthm()); []) | |
| 445 | handle GETHYPS hyps => hyps | |
| 446 | end; | |
| 447 | ||
| 448 | (*Prints exceptions nicely at top level; | |
| 449 | raises the exception in order to have a polymorphic type!*) | |
| 450 | fun print_exn e = (print_sign_exn_unit (Thm.theory_of_thm (topthm())) e; raise e); | |
| 451 | ||
| 452 | (*Which thms could apply to goal i? (debugs tactics involving filter_thms) *) | |
| 453 | fun filter_goal could ths i = filter_thms could (999, getgoal i, ths); | |
| 454 | ||
| 455 | (*For inspecting earlier levels of the backward proof*) | |
| 456 | fun chop_level n (pair,pairs) = | |
| 457 | let val level = length pairs | |
| 458 | in if n<0 andalso ~n <= level | |
| 459 | then List.drop (pair::pairs, ~n) | |
| 460 | else if 0<=n andalso n<= level | |
| 461 | then List.drop (pair::pairs, level - n) | |
| 462 |       else  error ("Level number must lie between 0 and " ^
 | |
| 463 | string_of_int level) | |
| 464 | end; | |
| 465 | ||
| 466 | (*Print the given level of the proof; prlev ~1 prints previous level*) | |
| 23635 | 467 | fun prlev n = apply_fun (print_top o pop o (chop_level n)); | 
| 468 | fun pr () = apply_fun print_top; | |
| 18120 | 469 | |
| 470 | (*Set goals_limit and print again*) | |
| 471 | fun prlim n = (goals_limit:=n; pr()); | |
| 472 | ||
| 473 | (** the goal.... commands | |
| 474 | Read main goal. Set global variables curr_prems, curr_mkresult. | |
| 475 | Initial subgoal and premises are rewritten using rths. **) | |
| 476 | ||
| 477 | (*Version taking the goal as a cterm; if you have a term t and theory thy, use | |
| 478 | goalw_cterm rths (cterm_of thy t); *) | |
| 479 | fun agoalw_cterm atomic rths chorn = | |
| 480 | let val (prems, st0, mkresult) = prepare_proof atomic rths chorn | |
| 481 | in undo_list := []; | |
| 482 | setstate [ (st0, Seq.empty) ]; | |
| 483 | curr_prems := prems; | |
| 484 | curr_mkresult := mkresult; | |
| 485 | prems | |
| 486 | end; | |
| 487 | ||
| 488 | val goalw_cterm = agoalw_cterm false; | |
| 489 | ||
| 490 | (*Version taking the goal as a string*) | |
| 491 | fun agoalw atomic thy rths agoal = | |
| 27256 
bcb071683184
added emulations for simple_read_term/read_term/read_prop (formerly in sign.ML);
 wenzelm parents: 
27242diff
changeset | 492 | agoalw_cterm atomic rths (cterm_of thy (read_prop thy agoal)) | 
| 18678 | 493 | handle ERROR msg => cat_error msg (*from type_assign, etc via prepare_proof*) | 
| 18120 | 494 |         ("The error(s) above occurred for " ^ quote agoal);
 | 
| 495 | ||
| 496 | val goalw = agoalw false; | |
| 497 | fun goal thy = goalw thy []; | |
| 498 | ||
| 499 | (*now the versions that wrap the goal up in `Goal' to make it atomic*) | |
| 26429 | 500 | fun Goalw thms s = agoalw true (ML_Context.the_global_context ()) thms s; | 
| 18120 | 501 | val Goal = Goalw []; | 
| 502 | ||
| 503 | (*simple version with minimal amount of checking and postprocessing*) | |
| 504 | fun simple_prove_goal_cterm G f = | |
| 505 | let | |
| 26291 | 506 | val _ = legacy_feature "Old goal command"; | 
| 18120 | 507 | val As = Drule.strip_imp_prems G; | 
| 508 | val B = Drule.strip_imp_concl G; | |
| 20229 | 509 | val asms = map Assumption.assume As; | 
| 18120 | 510 | fun check NONE = error "prove_goal: tactic failed" | 
| 511 | | check (SOME (thm, _)) = (case nprems_of thm of | |
| 512 | 0 => thm | |
| 513 | | i => !result_error_fn thm (string_of_int i ^ " unsolved goals!")) | |
| 514 | in | |
| 32957 | 515 | Drule.standard (implies_intr_list As | 
| 18120 | 516 | (check (Seq.pull (EVERY (f asms) (trivial B))))) | 
| 517 | end; | |
| 518 | ||
| 519 | ||
| 520 | (*Proof step "by" the given tactic -- apply tactic to the proof state*) | |
| 521 | fun by_com tac ((th,ths), pairs) : gstack = | |
| 522 | (case Seq.pull(tac th) of | |
| 523 | NONE => error"by: tactic failed" | |
| 524 | | SOME(th2,ths2) => | |
| 22360 
26ead7ed4f4b
moved eq_thm etc. to structure Thm in Pure/more_thm.ML;
 wenzelm parents: 
22109diff
changeset | 525 | (if Thm.eq_thm(th,th2) | 
| 18120 | 526 | then warning "Warning: same as previous level" | 
| 22360 
26ead7ed4f4b
moved eq_thm etc. to structure Thm in Pure/more_thm.ML;
 wenzelm parents: 
22109diff
changeset | 527 | else if Thm.eq_thm_thy(th,th2) then () | 
| 18120 | 528 |           else warning ("Warning: theory of proof state has changed" ^
 | 
| 529 | thy_error (Thm.theory_of_thm th, Thm.theory_of_thm th2)); | |
| 530 | ((th2,ths2)::(th,ths)::pairs))); | |
| 531 | ||
| 25685 | 532 | fun by tac = cond_timeit (!Output.timing) "" | 
| 18120 | 533 | (fn() => make_command (by_com tac)); | 
| 534 | ||
| 535 | (* byev[tac1,...,tacn] applies tac1 THEN ... THEN tacn. | |
| 536 | Good for debugging proofs involving prove_goal.*) | |
| 537 | val byev = by o EVERY; | |
| 538 | ||
| 539 | ||
| 540 | (*Backtracking means find an alternative result from a tactic. | |
| 541 | If none at this level, try earlier levels*) | |
| 542 | fun backtrack [] = error"back: no alternatives" | |
| 543 | | backtrack ((th,thstr) :: pairs) = | |
| 544 | (case Seq.pull thstr of | |
| 545 | NONE => (writeln"Going back a level..."; backtrack pairs) | |
| 546 | | SOME(th2,thstr2) => | |
| 22360 
26ead7ed4f4b
moved eq_thm etc. to structure Thm in Pure/more_thm.ML;
 wenzelm parents: 
22109diff
changeset | 547 | (if Thm.eq_thm(th,th2) | 
| 18120 | 548 | then warning "Warning: same as previous choice at this level" | 
| 22360 
26ead7ed4f4b
moved eq_thm etc. to structure Thm in Pure/more_thm.ML;
 wenzelm parents: 
22109diff
changeset | 549 | else if Thm.eq_thm_thy(th,th2) then () | 
| 18120 | 550 | else warning "Warning: theory of proof state has changed"; | 
| 551 | (th2,thstr2)::pairs)); | |
| 552 | ||
| 553 | fun back() = setstate (backtrack (getstate())); | |
| 554 | ||
| 555 | (*Chop back to previous level of the proof*) | |
| 556 | fun choplev n = make_command (chop_level n); | |
| 557 | ||
| 558 | (*Chopping back the goal stack*) | |
| 559 | fun chop () = make_command (fn (_,pairs) => pairs); | |
| 560 | ||
| 561 | (*Restore the previous proof state; discard current state. *) | |
| 562 | fun undo() = case !undo_list of | |
| 563 | [] => error"No proof state" | |
| 564 | | [_] => error"Already at initial state" | |
| 565 | | _::newundo => (undo_list := newundo; pr()) ; | |
| 566 | ||
| 567 | ||
| 568 | (*** Managing the proof stack ***) | |
| 569 | ||
| 570 | fun save_proof() = Proof(!undo_list, !curr_prems, !curr_mkresult); | |
| 571 | ||
| 572 | fun restore_proof(Proof(ul,prems,mk)) = | |
| 573 | (undo_list:= ul; curr_prems:= prems; curr_mkresult := mk; prems); | |
| 574 | ||
| 575 | ||
| 576 | fun top_proof() = case !proofstack of | |
| 577 |         [] => error("Stack of proof attempts is empty!")
 | |
| 578 | | p::ps => (p,ps); | |
| 579 | ||
| 580 | (* push a copy of the current proof state on to the stack *) | |
| 581 | fun push_proof() = (proofstack := (save_proof() :: !proofstack)); | |
| 582 | ||
| 583 | (* discard the top proof state of the stack *) | |
| 584 | fun pop_proof() = | |
| 585 | let val (p,ps) = top_proof() | |
| 586 | val prems = restore_proof p | |
| 587 | in proofstack := ps; pr(); prems end; | |
| 588 | ||
| 589 | (* rotate the stack so that the top element goes to the bottom *) | |
| 32173 | 590 | fun rotate_proof() = | 
| 591 | let val (p,ps) = top_proof() | |
| 592 | in proofstack := ps@[save_proof()]; | |
| 593 | restore_proof p; | |
| 594 | pr(); | |
| 595 | !curr_prems | |
| 596 | end; | |
| 18120 | 597 | |
| 598 | ||
| 26414 | 599 | (** theorem bindings **) | 
| 18120 | 600 | |
| 26414 | 601 | fun qed name = ML_Context.ml_store_thm (name, result ()); | 
| 602 | fun qed_goal name thy goal tacsf = ML_Context.ml_store_thm (name, prove_goal thy goal tacsf); | |
| 18120 | 603 | fun qed_goalw name thy rews goal tacsf = | 
| 26414 | 604 | ML_Context.ml_store_thm (name, prove_goalw thy rews goal tacsf); | 
| 18120 | 605 | fun qed_spec_mp name = | 
| 26414 | 606 | ML_Context.ml_store_thm (name, ObjectLogic.rulify_no_asm (result ())); | 
| 18120 | 607 | fun qed_goal_spec_mp name thy s p = | 
| 608 | bind_thm (name, ObjectLogic.rulify_no_asm (prove_goal thy s p)); | |
| 609 | fun qed_goalw_spec_mp name thy defs s p = | |
| 610 | bind_thm (name, ObjectLogic.rulify_no_asm (prove_goalw thy defs s p)); | |
| 611 | ||
| 612 | end; | |
| 613 |