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