| author | huffman | 
| Mon, 07 Mar 2005 23:54:01 +0100 | |
| changeset 15587 | f363e6e080e7 | 
| parent 15574 | b1d1b5bfc464 | 
| child 15596 | 8665d08085df | 
| permissions | -rw-r--r-- | 
| 5820 | 1 | (* Title: Pure/Isar/proof.ML | 
| 2 | ID: $Id$ | |
| 3 | Author: Markus Wenzel, TU Muenchen | |
| 4 | ||
| 5 | Proof states and methods. | |
| 6 | *) | |
| 7 | ||
| 15194 | 8 | |
| 9 | (*Jia Meng: added in atp_hook and modified enter_forward. *) | |
| 10 | ||
| 8152 | 11 | signature BASIC_PROOF = | 
| 12 | sig | |
| 8374 | 13 | val FINDGOAL: (int -> thm -> 'a Seq.seq) -> thm -> 'a Seq.seq | 
| 14 | val HEADGOAL: (int -> thm -> 'a Seq.seq) -> thm -> 'a Seq.seq | |
| 8152 | 15 | end; | 
| 16 | ||
| 5820 | 17 | signature PROOF = | 
| 18 | sig | |
| 8152 | 19 | include BASIC_PROOF | 
| 5820 | 20 | type context | 
| 21 | type state | |
| 14129 | 22 | datatype mode = Forward | ForwardChain | Backward | 
| 5820 | 23 | exception STATE of string * state | 
| 6871 | 24 | val check_result: string -> state -> 'a Seq.seq -> 'a Seq.seq | 
| 12045 | 25 | val init_state: theory -> state | 
| 5820 | 26 | val context_of: state -> context | 
| 27 | val theory_of: state -> theory | |
| 28 | val sign_of: state -> Sign.sg | |
| 13377 | 29 | val map_context: (context -> context) -> state -> state | 
| 7924 | 30 | val warn_extra_tfrees: state -> state -> state | 
| 14554 | 31 | val put_thms: string * thm list -> state -> state | 
| 7605 | 32 | val reset_thms: string -> state -> state | 
| 6091 | 33 | val the_facts: state -> thm list | 
| 9469 | 34 | val the_fact: state -> thm | 
| 11079 
a378479996f7
val get_goal: state -> context * (thm list * thm);
 wenzelm parents: 
10938diff
changeset | 35 | val get_goal: state -> context * (thm list * thm) | 
| 6091 | 36 | val goal_facts: (state -> thm list) -> state -> state | 
| 14129 | 37 | val get_mode: state -> mode | 
| 9469 | 38 | val is_chain: state -> bool | 
| 6891 | 39 | val assert_forward: state -> state | 
| 9469 | 40 | val assert_forward_or_chain: state -> state | 
| 5820 | 41 | val assert_backward: state -> state | 
| 8206 | 42 | val assert_no_chain: state -> state | 
| 15452 | 43 | val atp_hook: (ProofContext.context * Thm.thm -> unit) ref | 
| 5820 | 44 | val enter_forward: state -> state | 
| 6529 | 45 | val verbose: bool ref | 
| 13698 | 46 | val show_main_goal: bool ref | 
| 12423 | 47 | val pretty_state: int -> state -> Pretty.T list | 
| 10360 | 48 | val pretty_goals: bool -> state -> Pretty.T list | 
| 6776 | 49 | val level: state -> int | 
| 5820 | 50 | type method | 
| 6848 | 51 | val method: (thm list -> tactic) -> method | 
| 8374 | 52 | val method_cases: (thm list -> thm -> (thm * (string * RuleCases.T) list) Seq.seq) -> method | 
| 5820 | 53 | val refine: (context -> method) -> state -> state Seq.seq | 
| 8234 | 54 | val refine_end: (context -> method) -> state -> state Seq.seq | 
| 5936 | 55 | val match_bind: (string list * string) list -> state -> state | 
| 56 | val match_bind_i: (term list * term) list -> state -> state | |
| 8617 
33e2bd53aec3
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8582diff
changeset | 57 | val let_bind: (string list * string) list -> state -> state | 
| 
33e2bd53aec3
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8582diff
changeset | 58 | val let_bind_i: (term list * term) list -> state -> state | 
| 14564 | 59 | val simple_note_thms: string -> thm list -> state -> state | 
| 60 | val note_thmss: ((bstring * context attribute list) * | |
| 15456 
956d6acacf89
Specific theorems in a named list of theorems can now be referred to
 berghofe parents: 
15452diff
changeset | 61 | (thmref * context attribute list) list) list -> state -> state | 
| 14564 | 62 | val note_thmss_i: ((bstring * context attribute list) * | 
| 9196 | 63 | (thm list * context attribute list) list) list -> state -> state | 
| 12714 | 64 | val with_thmss: ((bstring * context attribute list) * | 
| 15456 
956d6acacf89
Specific theorems in a named list of theorems can now be referred to
 berghofe parents: 
15452diff
changeset | 65 | (thmref * context attribute list) list) list -> state -> state | 
| 12714 | 66 | val with_thmss_i: ((bstring * context attribute list) * | 
| 9196 | 67 | (thm list * context attribute list) list) list -> state -> state | 
| 15456 
956d6acacf89
Specific theorems in a named list of theorems can now be referred to
 berghofe parents: 
15452diff
changeset | 68 | val using_thmss: ((thmref * context attribute list) list) list -> state -> state | 
| 12930 | 69 | val using_thmss_i: ((thm list * context attribute list) list) list -> state -> state | 
| 14528 
1457584110ac
Locale instantiation: label parameter optional, new attribute paramter.
 ballarin parents: 
14508diff
changeset | 70 | val instantiate_locale: string * (string * context attribute list) -> state | 
| 
1457584110ac
Locale instantiation: label parameter optional, new attribute paramter.
 ballarin parents: 
14508diff
changeset | 71 | -> state | 
| 7412 | 72 | val fix: (string list * string option) list -> state -> state | 
| 7665 | 73 | val fix_i: (string list * typ option) list -> state -> state | 
| 9469 | 74 | val assm: ProofContext.exporter | 
| 10464 | 75 | -> ((string * context attribute list) * (string * (string list * string list)) list) list | 
| 7271 | 76 | -> state -> state | 
| 9469 | 77 | val assm_i: ProofContext.exporter | 
| 10464 | 78 | -> ((string * context attribute list) * (term * (term list * term list)) list) list | 
| 6932 | 79 | -> state -> state | 
| 10464 | 80 | val assume: | 
| 81 | ((string * context attribute list) * (string * (string list * string list)) list) list | |
| 7271 | 82 | -> state -> state | 
| 10464 | 83 | val assume_i: ((string * context attribute list) * (term * (term list * term list)) list) list | 
| 6932 | 84 | -> state -> state | 
| 10464 | 85 | val presume: | 
| 86 | ((string * context attribute list) * (string * (string list * string list)) list) list | |
| 6932 | 87 | -> state -> state | 
| 10464 | 88 | val presume_i: ((string * context attribute list) * (term * (term list * term list)) list) list | 
| 6932 | 89 | -> state -> state | 
| 11891 | 90 | val def: string -> context attribute list -> string * (string * string list) -> state -> state | 
| 91 | val def_i: string -> context attribute list -> string * (term * term list) -> state -> state | |
| 11793 
5f0ab6f5c280
support impromptu terminology of cases parameters;
 wenzelm parents: 
11742diff
changeset | 92 | val invoke_case: string * string option list * context attribute list -> state -> state | 
| 12959 | 93 | val multi_theorem: string | 
| 94 | -> (xstring * (context attribute list * context attribute list list)) option | |
| 15206 
09d78ec709c7
Modified locales: improved implementation of "includes".
 ballarin parents: 
15194diff
changeset | 95 | -> bstring * theory attribute list -> context attribute Locale.element Locale.elem_expr list | 
| 12146 | 96 | -> ((bstring * theory attribute list) * (string * (string list * string list)) list) list | 
| 97 | -> theory -> state | |
| 12959 | 98 | val multi_theorem_i: string | 
| 99 | -> (string * (context attribute list * context attribute list list)) option | |
| 100 | -> bstring * theory attribute list | |
| 15206 
09d78ec709c7
Modified locales: improved implementation of "includes".
 ballarin parents: 
15194diff
changeset | 101 | -> context attribute Locale.element_i Locale.elem_expr list | 
| 12146 | 102 | -> ((bstring * theory attribute list) * (term * (term list * term list)) list) list | 
| 103 | -> theory -> state | |
| 5820 | 104 | val chain: state -> state | 
| 6091 | 105 | val from_facts: thm list -> state -> state | 
| 12971 | 106 | val show: (bool -> state -> state) -> (state -> state Seq.seq) -> bool | 
| 12146 | 107 | -> ((string * context attribute list) * (string * (string list * string list)) list) list | 
| 108 | -> bool -> state -> state | |
| 12971 | 109 | val show_i: (bool -> state -> state) -> (state -> state Seq.seq) -> bool | 
| 12146 | 110 | -> ((string * context attribute list) * (term * (term list * term list)) list) list | 
| 111 | -> bool -> state -> state | |
| 12971 | 112 | val have: (state -> state Seq.seq) -> bool | 
| 12146 | 113 | -> ((string * context attribute list) * (string * (string list * string list)) list) list | 
| 114 | -> state -> state | |
| 12971 | 115 | val have_i: (state -> state Seq.seq) -> bool | 
| 12146 | 116 | -> ((string * context attribute list) * (term * (term list * term list)) list) list | 
| 117 | -> state -> state | |
| 6404 | 118 | val at_bottom: state -> bool | 
| 6982 | 119 | val local_qed: (state -> state Seq.seq) | 
| 12146 | 120 | -> (context -> string * (string * thm list) list -> unit) * (context -> thm -> unit) | 
| 121 | -> state -> state Seq.seq | |
| 6950 | 122 | val global_qed: (state -> state Seq.seq) -> state | 
| 12244 | 123 | -> (theory * ((string * string) * (string * thm list) list)) Seq.seq | 
| 6896 | 124 | val begin_block: state -> state | 
| 6932 | 125 | val end_block: state -> state Seq.seq | 
| 6896 | 126 | val next_block: state -> state | 
| 14548 | 127 | val thisN: string | 
| 15194 | 128 | val call_atp: bool ref; | 
| 129 | ||
| 130 | ||
| 5820 | 131 | end; | 
| 132 | ||
| 13377 | 133 | structure Proof: PROOF = | 
| 5820 | 134 | struct | 
| 135 | ||
| 15194 | 136 | (*Jia Meng: by default, no ATP is called. *) | 
| 137 | val call_atp = ref false; | |
| 5820 | 138 | |
| 139 | (** proof state **) | |
| 140 | ||
| 141 | type context = ProofContext.context; | |
| 142 | ||
| 143 | ||
| 144 | (* type goal *) | |
| 145 | ||
| 15099 | 146 | (* CB: three kinds of Isar goals are distinguished: | 
| 147 | - Theorem: global goal in a theory, corresponds to Isar commands theorem, | |
| 148 | lemma and corollary, | |
| 149 | - Have: local goal in a proof, Isar command have | |
| 150 | - Show: local goal in a proof, Isar command show | |
| 151 | *) | |
| 152 | ||
| 5820 | 153 | datatype kind = | 
| 12959 | 154 |   Theorem of {kind: string,
 | 
| 155 | theory_spec: (bstring * theory attribute list) * theory attribute list list, | |
| 13415 | 156 | locale_spec: (string * (context attribute list * context attribute list list)) option, | 
| 15206 
09d78ec709c7
Modified locales: improved implementation of "includes".
 ballarin parents: 
15194diff
changeset | 157 | view: cterm list * cterm list} | | 
| 
09d78ec709c7
Modified locales: improved implementation of "includes".
 ballarin parents: 
15194diff
changeset | 158 | (* target view and includes view *) | 
| 12959 | 159 | Show of context attribute list list | | 
| 160 | Have of context attribute list list; | |
| 5820 | 161 | |
| 15099 | 162 | (* CB: this function prints the goal kind (Isar command), name and target in | 
| 163 | the proof state *) | |
| 164 | ||
| 13415 | 165 | fun kind_name _ (Theorem {kind = s, theory_spec = ((a, _), _),
 | 
| 15531 | 166 | locale_spec = NONE, view = _}) = s ^ (if a = "" then "" else " " ^ a) | 
| 13399 
c136276dc847
support locale ``views'' (for cumulative predicates);
 wenzelm parents: 
13377diff
changeset | 167 |   | kind_name sg (Theorem {kind = s, theory_spec = ((a, _), _),
 | 
| 15531 | 168 | locale_spec = SOME (name, _), view = _}) = | 
| 13415 | 169 | s ^ " (in " ^ Locale.cond_extern sg name ^ ")" ^ (if a = "" then "" else " " ^ a) | 
| 12015 | 170 | | kind_name _ (Show _) = "show" | 
| 171 | | kind_name _ (Have _) = "have"; | |
| 5820 | 172 | |
| 173 | type goal = | |
| 12146 | 174 | (kind * (*result kind*) | 
| 175 | string list * (*result names*) | |
| 176 | term list list) * (*result statements*) | |
| 177 | (thm list * (*use facts*) | |
| 178 | thm); (*goal: subgoals ==> statement*) | |
| 5820 | 179 | |
| 180 | ||
| 181 | (* type mode *) | |
| 182 | ||
| 183 | datatype mode = Forward | ForwardChain | Backward; | |
| 7201 | 184 | val mode_name = (fn Forward => "state" | ForwardChain => "chain" | Backward => "prove"); | 
| 5820 | 185 | |
| 186 | ||
| 187 | (* datatype state *) | |
| 188 | ||
| 7176 | 189 | datatype node = | 
| 190 | Node of | |
| 191 |    {context: context,
 | |
| 192 | facts: thm list option, | |
| 193 | mode: mode, | |
| 12971 | 194 | goal: (goal * ((state -> state Seq.seq) * bool)) option} | 
| 7176 | 195 | and state = | 
| 5820 | 196 | State of | 
| 197 | node * (*current*) | |
| 198 | node list; (*parents wrt. block structure*) | |
| 199 | ||
| 7176 | 200 | fun make_node (context, facts, mode, goal) = | 
| 201 |   Node {context = context, facts = facts, mode = mode, goal = goal};
 | |
| 202 | ||
| 203 | ||
| 5820 | 204 | exception STATE of string * state; | 
| 205 | ||
| 206 | fun err_malformed name state = | |
| 207 | raise STATE (name ^ ": internal error -- malformed proof state", state); | |
| 208 | ||
| 6871 | 209 | fun check_result msg state sq = | 
| 210 | (case Seq.pull sq of | |
| 15531 | 211 | NONE => raise STATE (msg, state) | 
| 212 | | SOME s_sq => Seq.cons s_sq); | |
| 6871 | 213 | |
| 5820 | 214 | |
| 7176 | 215 | fun map_current f (State (Node {context, facts, mode, goal}, nodes)) =
 | 
| 5820 | 216 | State (make_node (f (context, facts, mode, goal)), nodes); | 
| 217 | ||
| 12045 | 218 | fun init_state thy = | 
| 15531 | 219 | State (make_node (ProofContext.init thy, NONE, Forward, NONE), []); | 
| 12045 | 220 | |
| 5820 | 221 | |
| 222 | ||
| 223 | (** basic proof state operations **) | |
| 224 | ||
| 225 | (* context *) | |
| 226 | ||
| 7176 | 227 | fun context_of (State (Node {context, ...}, _)) = context;
 | 
| 5820 | 228 | val theory_of = ProofContext.theory_of o context_of; | 
| 229 | val sign_of = ProofContext.sign_of o context_of; | |
| 230 | ||
| 231 | fun map_context f = map_current (fn (ctxt, facts, mode, goal) => (f ctxt, facts, mode, goal)); | |
| 232 | ||
| 7176 | 233 | fun map_context_result f (state as State (Node {context, facts, mode, goal}, nodes)) =
 | 
| 5820 | 234 | let val (context', result) = f context | 
| 235 | in (State (make_node (context', facts, mode, goal), nodes), result) end; | |
| 236 | ||
| 237 | ||
| 7924 | 238 | val warn_extra_tfrees = map_context o ProofContext.warn_extra_tfrees o context_of; | 
| 10809 | 239 | val add_binds_i = map_context o ProofContext.add_binds_i; | 
| 6790 | 240 | val auto_bind_goal = map_context o ProofContext.auto_bind_goal; | 
| 12146 | 241 | val auto_bind_facts = map_context o ProofContext.auto_bind_facts; | 
| 6091 | 242 | val put_thms = map_context o ProofContext.put_thms; | 
| 243 | val put_thmss = map_context o ProofContext.put_thmss; | |
| 7605 | 244 | val reset_thms = map_context o ProofContext.reset_thms; | 
| 8374 | 245 | val get_case = ProofContext.get_case o context_of; | 
| 5820 | 246 | |
| 247 | ||
| 248 | (* facts *) | |
| 249 | ||
| 15531 | 250 | fun the_facts (State (Node {facts = SOME facts, ...}, _)) = facts
 | 
| 5820 | 251 |   | the_facts state = raise STATE ("No current facts available", state);
 | 
| 252 | ||
| 9469 | 253 | fun the_fact state = | 
| 254 | (case the_facts state of | |
| 255 | [thm] => thm | |
| 256 |   | _ => raise STATE ("Single fact expected", state));
 | |
| 257 | ||
| 6848 | 258 | fun assert_facts state = (the_facts state; state); | 
| 7176 | 259 | fun get_facts (State (Node {facts, ...}, _)) = facts;
 | 
| 6848 | 260 | |
| 7605 | 261 | |
| 262 | val thisN = "this"; | |
| 263 | ||
| 5820 | 264 | fun put_facts facts state = | 
| 265 | state | |
| 266 | |> map_current (fn (ctxt, _, mode, goal) => (ctxt, facts, mode, goal)) | |
| 15531 | 267 | |> (case facts of NONE => reset_thms thisN | SOME ths => put_thms (thisN, ths)); | 
| 5820 | 268 | |
| 15531 | 269 | val reset_facts = put_facts NONE; | 
| 5820 | 270 | |
| 9196 | 271 | fun these_factss (state, named_factss) = | 
| 15570 | 272 | state |> put_facts (SOME (List.concat (map snd named_factss))); | 
| 5820 | 273 | |
| 274 | ||
| 275 | (* goal *) | |
| 276 | ||
| 10553 | 277 | local | 
| 15531 | 278 |   fun find i (state as State (Node {goal = SOME goal, ...}, _)) = (context_of state, (i, goal))
 | 
| 279 |     | find i (State (Node {goal = NONE, ...}, node :: nodes)) = find (i + 1) (State (node, nodes))
 | |
| 10553 | 280 | | find _ (state as State (_, [])) = err_malformed "find_goal" state; | 
| 281 | in val find_goal = find 0 end; | |
| 7176 | 282 | |
| 11079 
a378479996f7
val get_goal: state -> context * (thm list * thm);
 wenzelm parents: 
10938diff
changeset | 283 | fun get_goal state = | 
| 
a378479996f7
val get_goal: state -> context * (thm list * thm);
 wenzelm parents: 
10938diff
changeset | 284 | let val (ctxt, (_, ((_, x), _))) = find_goal state | 
| 
a378479996f7
val get_goal: state -> context * (thm list * thm);
 wenzelm parents: 
10938diff
changeset | 285 | in (ctxt, x) end; | 
| 
a378479996f7
val get_goal: state -> context * (thm list * thm);
 wenzelm parents: 
10938diff
changeset | 286 | |
| 15452 | 287 | |
| 288 | (* | |
| 289 | (* Jia: added here: get all goals from the state 10/01/05 *) | |
| 15531 | 290 | fun find_all i (state as State (Node {goal = SOME goal, ...}, node::nodes)) = 
 | 
| 15452 | 291 | let val others = find_all (i+1) (State (node, nodes)) | 
| 292 | in | |
| 293 | (context_of state, (i, goal)) :: others | |
| 294 | end | |
| 15531 | 295 |   | find_all i (State (Node {goal = NONE, ...}, node :: nodes)) = find_all (i + 1) (State (node, nodes))
 | 
| 15452 | 296 | | find_all _ (state as State (_, [])) = []; | 
| 297 | ||
| 298 | val find_all_goals = find_all 0; | |
| 299 | ||
| 300 | fun find_all_goals_ctxts state = | |
| 301 | let val all_goals = find_all_goals state | |
| 302 | fun find_ctxt_g [] = [] | |
| 303 | | find_ctxt_g ((ctxt, (_, ((_, (_, thm)), _)))::others) = (ctxt,thm) :: (find_ctxt_g others) | |
| 304 | in | |
| 305 | find_ctxt_g all_goals | |
| 306 | end; | |
| 307 | *) | |
| 308 | ||
| 5820 | 309 | fun put_goal goal = map_current (fn (ctxt, facts, mode, _) => (ctxt, facts, mode, goal)); | 
| 310 | ||
| 15531 | 311 | fun map_goal f g (State (Node {context, facts, mode, goal = SOME goal}, nodes)) =
 | 
| 312 | State (make_node (f context, facts, mode, SOME (g goal)), nodes) | |
| 8374 | 313 | | map_goal f g (State (nd, node :: nodes)) = | 
| 314 | let val State (node', nodes') = map_goal f g (State (node, nodes)) | |
| 315 | in map_context f (State (nd, node' :: nodes')) end | |
| 316 | | map_goal _ _ state = state; | |
| 5820 | 317 | |
| 318 | fun goal_facts get state = | |
| 319 | state | |
| 8374 | 320 | |> map_goal I (fn ((result, (_, thm)), f) => ((result, (get state, thm)), f)); | 
| 5820 | 321 | |
| 322 | fun use_facts state = | |
| 323 | state | |
| 324 | |> goal_facts the_facts | |
| 325 | |> reset_facts; | |
| 326 | ||
| 327 | ||
| 328 | (* mode *) | |
| 329 | ||
| 7176 | 330 | fun get_mode (State (Node {mode, ...}, _)) = mode;
 | 
| 5820 | 331 | fun put_mode mode = map_current (fn (ctxt, facts, _, goal) => (ctxt, facts, mode, goal)); | 
| 332 | ||
| 15194 | 333 | val enter_forward_aux = put_mode Forward; | 
| 5820 | 334 | val enter_forward_chain = put_mode ForwardChain; | 
| 335 | val enter_backward = put_mode Backward; | |
| 336 | ||
| 337 | fun assert_mode pred state = | |
| 338 | let val mode = get_mode state in | |
| 339 | if pred mode then state | |
| 12010 | 340 |     else raise STATE ("Illegal application of proof command in "
 | 
| 341 | ^ quote (mode_name mode) ^ " mode", state) | |
| 5820 | 342 | end; | 
| 343 | ||
| 344 | fun is_chain state = get_mode state = ForwardChain; | |
| 345 | val assert_forward = assert_mode (equal Forward); | |
| 346 | val assert_forward_or_chain = assert_mode (equal Forward orf equal ForwardChain); | |
| 347 | val assert_backward = assert_mode (equal Backward); | |
| 8206 | 348 | val assert_no_chain = assert_mode (not_equal ForwardChain); | 
| 5820 | 349 | |
| 350 | ||
| 15194 | 351 | (*Jia Meng: a hook to be used for calling ATPs. *) | 
| 15452 | 352 | (* send the entire proof context *) | 
| 353 | val atp_hook = ref (fn ctxt_state: ProofContext.context * Thm.thm => ()); | |
| 15194 | 354 | |
| 355 | (*Jia Meng: the modified version of enter_forward. *) | |
| 356 | (*Jia Meng: atp: Proof.state -> unit *) | |
| 357 | local | |
| 358 | fun atp state = | |
| 15452 | 359 | let val (ctxt, (_, ((_, (_, thm)), _))) = find_goal state | 
| 360 | in | |
| 361 | (!atp_hook) (ctxt,thm) | |
| 362 | end; | |
| 363 | ||
| 15194 | 364 | in | 
| 365 | ||
| 366 | fun enter_forward state = | |
| 367 | if (!call_atp) then | |
| 368 | ((atp state; enter_forward_aux state) | |
| 369 | handle (STATE _) => enter_forward_aux state) | |
| 370 | else (enter_forward_aux state); | |
| 371 | ||
| 372 | end; | |
| 373 | ||
| 374 | ||
| 5820 | 375 | (* blocks *) | 
| 376 | ||
| 6776 | 377 | fun level (State (_, nodes)) = length nodes; | 
| 378 | ||
| 5820 | 379 | fun open_block (State (node, nodes)) = State (node, node :: nodes); | 
| 380 | ||
| 381 | fun new_block state = | |
| 382 | state | |
| 383 | |> open_block | |
| 15531 | 384 | |> put_goal NONE; | 
| 5820 | 385 | |
| 7487 
c0f9b956a3e7
close_block: removed ProofContext.transfer_used_names;
 wenzelm parents: 
7478diff
changeset | 386 | fun close_block (state as State (_, node :: nodes)) = State (node, nodes) | 
| 5820 | 387 |   | close_block state = raise STATE ("Unbalanced block parentheses", state);
 | 
| 388 | ||
| 389 | ||
| 390 | ||
| 12423 | 391 | (** pretty_state **) | 
| 5820 | 392 | |
| 13869 | 393 | val show_main_goal = ref false; | 
| 13698 | 394 | |
| 6529 | 395 | val verbose = ProofContext.verbose; | 
| 396 | ||
| 14876 | 397 | val pretty_goals_local = Display.pretty_goals_aux o ProofContext.pp; | 
| 12085 | 398 | |
| 15531 | 399 | fun pretty_facts _ _ NONE = [] | 
| 400 | | pretty_facts s ctxt (SOME ths) = | |
| 12055 | 401 | [Pretty.big_list (s ^ "this:") (map (ProofContext.pretty_thm ctxt) ths), Pretty.str ""]; | 
| 6756 | 402 | |
| 12423 | 403 | fun pretty_state nr (state as State (Node {context = ctxt, facts, mode, goal = _}, nodes)) =
 | 
| 5820 | 404 | let | 
| 11891 | 405 | val ref (_, _, begin_goal) = Display.current_goals_markers; | 
| 5820 | 406 | |
| 407 | fun levels_up 0 = "" | |
| 7575 | 408 | | levels_up 1 = ", 1 level up" | 
| 409 | | levels_up i = ", " ^ string_of_int i ^ " levels up"; | |
| 5820 | 410 | |
| 11556 | 411 | fun subgoals 0 = "" | 
| 412 | | subgoals 1 = ", 1 subgoal" | |
| 413 | | subgoals n = ", " ^ string_of_int n ^ " subgoals"; | |
| 12146 | 414 | |
| 415 | fun prt_names names = | |
| 12242 
282a92bc5655
multi_theorem: common statement header (covers *all* results);
 wenzelm parents: 
12223diff
changeset | 416 | (case filter_out (equal "") names of [] => "" | 
| 15570 | 417 |       | nms => " " ^ enclose "(" ")" (space_implode " and " (Library.take (7, nms) @
 | 
| 12308 | 418 | (if length nms > 7 then ["..."] else [])))); | 
| 12146 | 419 | |
| 420 | fun prt_goal (_, (i, (((kind, names, _), (goal_facts, thm)), _))) = | |
| 12085 | 421 | pretty_facts "using " ctxt | 
| 15531 | 422 | (if mode <> Backward orelse null goal_facts then NONE else SOME goal_facts) @ | 
| 12146 | 423 |       [Pretty.str ("goal (" ^ kind_name (sign_of state) kind ^ prt_names names ^
 | 
| 424 | levels_up (i div 2) ^ subgoals (Thm.nprems_of thm) ^ "):")] @ | |
| 13698 | 425 | pretty_goals_local ctxt begin_goal (true, !show_main_goal) (! Display.goals_limit) thm; | 
| 6848 | 426 | |
| 8462 | 427 | val ctxt_prts = | 
| 12085 | 428 | if ! verbose orelse mode = Forward then ProofContext.pretty_context ctxt | 
| 429 | else if mode = Backward then ProofContext.pretty_asms ctxt | |
| 7575 | 430 | else []; | 
| 8462 | 431 | |
| 432 | val prts = | |
| 8582 | 433 |      [Pretty.str ("proof (" ^ mode_name mode ^ "): step " ^ string_of_int nr ^
 | 
| 12010 | 434 | (if ! verbose then ", depth " ^ string_of_int (length nodes div 2 - 1) | 
| 8561 | 435 | else "")), Pretty.str ""] @ | 
| 8462 | 436 | (if null ctxt_prts then [] else ctxt_prts @ [Pretty.str ""]) @ | 
| 437 | (if ! verbose orelse mode = Forward then | |
| 12146 | 438 | (pretty_facts "" ctxt facts @ prt_goal (find_goal state)) | 
| 12085 | 439 | else if mode = ForwardChain then pretty_facts "picking " ctxt facts | 
| 12146 | 440 | else prt_goal (find_goal state)) | 
| 12423 | 441 | in prts end; | 
| 5820 | 442 | |
| 12085 | 443 | fun pretty_goals main state = | 
| 444 | let val (ctxt, (_, ((_, (_, thm)), _))) = find_goal state | |
| 445 | in pretty_goals_local ctxt "" (false, main) (! Display.goals_limit) thm end; | |
| 10320 | 446 | |
| 5820 | 447 | |
| 448 | ||
| 449 | (** proof steps **) | |
| 450 | ||
| 451 | (* datatype method *) | |
| 452 | ||
| 8374 | 453 | datatype method = | 
| 454 | Method of thm list -> thm -> (thm * (string * RuleCases.T) list) Seq.seq; | |
| 455 | ||
| 456 | fun method tac = Method (fn facts => fn st => Seq.map (rpair []) (tac facts st)); | |
| 457 | val method_cases = Method; | |
| 5820 | 458 | |
| 459 | ||
| 460 | (* refine goal *) | |
| 461 | ||
| 8234 | 462 | local | 
| 463 | ||
| 5820 | 464 | fun check_sign sg state = | 
| 465 | if Sign.subsig (sg, sign_of state) then state | |
| 466 |   else raise STATE ("Bad signature of result: " ^ Sign.str_of_sg sg, state);
 | |
| 467 | ||
| 15533 
accd51fdae3c
refine now provides specific cases "goal1" ... "goaln" for addressing
 berghofe parents: 
15531diff
changeset | 468 | fun mk_goals_cases st = | 
| 
accd51fdae3c
refine now provides specific cases "goal1" ... "goaln" for addressing
 berghofe parents: 
15531diff
changeset | 469 | RuleCases.make true NONE (sign_of_thm st, prop_of st) | 
| 
accd51fdae3c
refine now provides specific cases "goal1" ... "goaln" for addressing
 berghofe parents: 
15531diff
changeset | 470 | (map (fn i => "goal" ^ string_of_int i) (1 upto nprems_of st)); | 
| 
accd51fdae3c
refine now provides specific cases "goal1" ... "goaln" for addressing
 berghofe parents: 
15531diff
changeset | 471 | |
| 8234 | 472 | fun gen_refine current_context meth_fun state = | 
| 6848 | 473 | let | 
| 12971 | 474 | val (goal_ctxt, (_, ((result, (facts, thm)), x))) = find_goal state; | 
| 8234 | 475 | val Method meth = meth_fun (if current_context then context_of state else goal_ctxt); | 
| 5820 | 476 | |
| 8374 | 477 | fun refn (thm', cases) = | 
| 6848 | 478 | state | 
| 479 | |> check_sign (Thm.sign_of_thm thm') | |
| 15533 
accd51fdae3c
refine now provides specific cases "goal1" ... "goaln" for addressing
 berghofe parents: 
15531diff
changeset | 480 | |> map_goal (ProofContext.add_cases (cases @ mk_goals_cases thm')) | 
| 
accd51fdae3c
refine now provides specific cases "goal1" ... "goaln" for addressing
 berghofe parents: 
15531diff
changeset | 481 | (K ((result, (facts, thm')), x)); | 
| 6848 | 482 | in Seq.map refn (meth facts thm) end; | 
| 5820 | 483 | |
| 8234 | 484 | in | 
| 485 | ||
| 486 | val refine = gen_refine true; | |
| 487 | val refine_end = gen_refine false; | |
| 488 | ||
| 489 | end; | |
| 490 | ||
| 5820 | 491 | |
| 11917 | 492 | (* goal addressing *) | 
| 6932 | 493 | |
| 8152 | 494 | fun FINDGOAL tac st = | 
| 8374 | 495 | let fun find (i, n) = if i > n then Seq.fail else Seq.APPEND (tac i, find (i + 1, n)) | 
| 496 | in find (1, Thm.nprems_of st) st end; | |
| 8152 | 497 | |
| 8166 | 498 | fun HEADGOAL tac = tac 1; | 
| 499 | ||
| 9469 | 500 | |
| 501 | (* export results *) | |
| 502 | ||
| 11816 | 503 | fun refine_tac rule = | 
| 12703 
af5fec8a418f
refine_tac: Tactic.norm_hhf_tac before trying rule;
 wenzelm parents: 
12698diff
changeset | 504 | Tactic.norm_hhf_tac THEN' Tactic.rtac rule THEN_ALL_NEW (SUBGOAL (fn (goal, i) => | 
| 11816 | 505 | if can Logic.dest_goal (Logic.strip_assums_concl goal) then | 
| 506 | Tactic.etac Drule.triv_goal i | |
| 507 | else all_tac)); | |
| 508 | ||
| 12146 | 509 | fun export_goal inner print_rule raw_rule state = | 
| 6932 | 510 | let | 
| 12971 | 511 | val (outer, (_, ((result, (facts, thm)), x))) = find_goal state; | 
| 12010 | 512 | val exp_tac = ProofContext.export true inner outer; | 
| 9469 | 513 | fun exp_rule rule = | 
| 514 | let | |
| 12055 | 515 | val _ = print_rule outer rule; | 
| 11816 | 516 | val thmq = FINDGOAL (refine_tac rule) thm; | 
| 12971 | 517 | in Seq.map (fn th => map_goal I (K ((result, (facts, th)), x)) state) thmq end; | 
| 9469 | 518 | in raw_rule |> exp_tac |> (Seq.flat o Seq.map exp_rule) end; | 
| 6932 | 519 | |
| 12146 | 520 | fun export_goals inner print_rule raw_rules = | 
| 521 | Seq.EVERY (map (export_goal inner print_rule) raw_rules); | |
| 522 | ||
| 6932 | 523 | fun transfer_facts inner_state state = | 
| 524 | (case get_facts inner_state of | |
| 15531 | 525 | NONE => Seq.single (reset_facts state) | 
| 526 | | SOME thms => | |
| 8617 
33e2bd53aec3
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8582diff
changeset | 527 | let | 
| 12010 | 528 | val exp_tac = ProofContext.export false (context_of inner_state) (context_of state); | 
| 11816 | 529 | val thmqs = map exp_tac thms; | 
| 15531 | 530 | in Seq.map (fn ths => put_facts (SOME ths) state) (Seq.commute thmqs) end); | 
| 6932 | 531 | |
| 532 | ||
| 533 | (* prepare result *) | |
| 534 | ||
| 12146 | 535 | fun prep_result state ts raw_th = | 
| 5820 | 536 | let | 
| 537 | val ctxt = context_of state; | |
| 538 | fun err msg = raise STATE (msg, state); | |
| 539 | ||
| 12146 | 540 | val ngoals = Thm.nprems_of raw_th; | 
| 5820 | 541 | val _ = | 
| 542 | if ngoals = 0 then () | |
| 13698 | 543 | else (err (Pretty.string_of (Pretty.chunks (pretty_goals_local ctxt "" (true, !show_main_goal) | 
| 12146 | 544 | (! Display.goals_limit) raw_th @ | 
| 12085 | 545 | [Pretty.str (string_of_int ngoals ^ " unsolved goal(s)!")])))); | 
| 5820 | 546 | |
| 12146 | 547 |     val {hyps, sign, ...} = Thm.rep_thm raw_th;
 | 
| 5820 | 548 | val tsig = Sign.tsig_of sign; | 
| 15570 | 549 | val casms = List.concat (map #1 (ProofContext.assumptions_of (context_of state))); | 
| 11816 | 550 | val bad_hyps = Library.gen_rems Term.aconv (hyps, map Thm.term_of casms); | 
| 12146 | 551 | |
| 552 | val th = raw_th RS Drule.rev_triv_goal; | |
| 553 | val ths = Drule.conj_elim_precise (length ts) th | |
| 554 | handle THM _ => err "Main goal structure corrupted"; | |
| 5820 | 555 | in | 
| 12146 | 556 |     conditional (not (null bad_hyps)) (fn () => err ("Additional hypotheses:\n" ^
 | 
| 557 | cat_lines (map (ProofContext.string_of_term ctxt) bad_hyps))); | |
| 12808 
a529b4b91888
RuleCases.make interface based on term instead of thm;
 wenzelm parents: 
12760diff
changeset | 558 | conditional (exists (not o Pattern.matches tsig) (ts ~~ map Thm.prop_of ths)) (fn () => | 
| 12146 | 559 |       err ("Proved a different theorem: " ^ Pretty.string_of (ProofContext.pretty_thm ctxt th)));
 | 
| 560 | ths | |
| 5820 | 561 | end; | 
| 562 | ||
| 563 | ||
| 564 | ||
| 565 | (*** structured proof commands ***) | |
| 566 | ||
| 567 | (** context **) | |
| 568 | ||
| 569 | (* bind *) | |
| 570 | ||
| 571 | fun gen_bind f x state = | |
| 572 | state | |
| 573 | |> assert_forward | |
| 574 | |> map_context (f x) | |
| 575 | |> reset_facts; | |
| 576 | ||
| 8617 
33e2bd53aec3
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8582diff
changeset | 577 | val match_bind = gen_bind (ProofContext.match_bind false); | 
| 
33e2bd53aec3
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8582diff
changeset | 578 | val match_bind_i = gen_bind (ProofContext.match_bind_i false); | 
| 
33e2bd53aec3
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8582diff
changeset | 579 | val let_bind = gen_bind (ProofContext.match_bind true); | 
| 
33e2bd53aec3
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8582diff
changeset | 580 | val let_bind_i = gen_bind (ProofContext.match_bind_i true); | 
| 5820 | 581 | |
| 582 | ||
| 14564 | 583 | (* note_thmss *) | 
| 5820 | 584 | |
| 12714 | 585 | local | 
| 586 | ||
| 14564 | 587 | fun gen_note_thmss f args state = | 
| 5820 | 588 | state | 
| 589 | |> assert_forward | |
| 12714 | 590 | |> map_context_result (f args) | 
| 9196 | 591 | |> these_factss; | 
| 5820 | 592 | |
| 12714 | 593 | in | 
| 594 | ||
| 14564 | 595 | val note_thmss = gen_note_thmss ProofContext.note_thmss; | 
| 596 | val note_thmss_i = gen_note_thmss ProofContext.note_thmss_i; | |
| 12714 | 597 | |
| 14564 | 598 | fun simple_note_thms name thms = note_thmss_i [((name, []), [(thms, [])])]; | 
| 12714 | 599 | |
| 600 | end; | |
| 601 | ||
| 602 | ||
| 603 | (* with_thmss *) | |
| 9196 | 604 | |
| 12714 | 605 | local | 
| 606 | ||
| 607 | fun gen_with_thmss f args state = | |
| 608 | let val state' = state |> f args | |
| 14564 | 609 | in state' |> simple_note_thms "" (the_facts state' @ the_facts state) end; | 
| 6876 | 610 | |
| 12714 | 611 | in | 
| 612 | ||
| 14564 | 613 | val with_thmss = gen_with_thmss note_thmss; | 
| 614 | val with_thmss_i = gen_with_thmss note_thmss_i; | |
| 12714 | 615 | |
| 616 | end; | |
| 617 | ||
| 5820 | 618 | |
| 12930 | 619 | (* using_thmss *) | 
| 620 | ||
| 621 | local | |
| 622 | ||
| 623 | fun gen_using_thmss f args state = | |
| 624 | state | |
| 625 | |> assert_backward | |
| 626 |   |> map_context_result (f (map (pair ("", [])) args))
 | |
| 627 | |> (fn (st, named_facts) => | |
| 12971 | 628 | let val (_, (_, ((result, (facts, thm)), x))) = find_goal st; | 
| 15570 | 629 | in st |> map_goal I (K ((result, (facts @ List.concat (map snd named_facts), thm)), x)) end); | 
| 12930 | 630 | |
| 631 | in | |
| 632 | ||
| 14564 | 633 | val using_thmss = gen_using_thmss ProofContext.note_thmss; | 
| 634 | val using_thmss_i = gen_using_thmss ProofContext.note_thmss_i; | |
| 12930 | 635 | |
| 636 | end; | |
| 637 | ||
| 638 | ||
| 14508 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 ballarin parents: 
14404diff
changeset | 639 | (* locale instantiation *) | 
| 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 ballarin parents: 
14404diff
changeset | 640 | |
| 14528 
1457584110ac
Locale instantiation: label parameter optional, new attribute paramter.
 ballarin parents: 
14508diff
changeset | 641 | fun instantiate_locale (loc_name, prfx_attribs) state = let | 
| 15531 | 642 | val facts = if is_chain state then get_facts state else NONE; | 
| 14508 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 ballarin parents: 
14404diff
changeset | 643 | in | 
| 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 ballarin parents: 
14404diff
changeset | 644 | state | 
| 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 ballarin parents: 
14404diff
changeset | 645 | |> assert_forward_or_chain | 
| 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 ballarin parents: 
14404diff
changeset | 646 | |> enter_forward | 
| 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 ballarin parents: 
14404diff
changeset | 647 | |> reset_facts | 
| 14528 
1457584110ac
Locale instantiation: label parameter optional, new attribute paramter.
 ballarin parents: 
14508diff
changeset | 648 | |> map_context (Locale.instantiate loc_name prfx_attribs facts) | 
| 14508 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 ballarin parents: 
14404diff
changeset | 649 | end; | 
| 
859b11514537
Experimental command for instantiation of locales in proof contexts:
 ballarin parents: 
14404diff
changeset | 650 | |
| 5820 | 651 | (* fix *) | 
| 652 | ||
| 653 | fun gen_fix f xs state = | |
| 654 | state | |
| 655 | |> assert_forward | |
| 656 | |> map_context (f xs) | |
| 657 | |> reset_facts; | |
| 658 | ||
| 659 | val fix = gen_fix ProofContext.fix; | |
| 660 | val fix_i = gen_fix ProofContext.fix_i; | |
| 661 | ||
| 662 | ||
| 11891 | 663 | (* assume and presume *) | 
| 5820 | 664 | |
| 9469 | 665 | fun gen_assume f exp args state = | 
| 5820 | 666 | state | 
| 667 | |> assert_forward | |
| 9469 | 668 | |> map_context_result (f exp args) | 
| 11924 | 669 | |> these_factss; | 
| 6932 | 670 | |
| 7271 | 671 | val assm = gen_assume ProofContext.assume; | 
| 672 | val assm_i = gen_assume ProofContext.assume_i; | |
| 11917 | 673 | val assume = assm ProofContext.export_assume; | 
| 674 | val assume_i = assm_i ProofContext.export_assume; | |
| 675 | val presume = assm ProofContext.export_presume; | |
| 676 | val presume_i = assm_i ProofContext.export_presume; | |
| 5820 | 677 | |
| 7271 | 678 | |
| 5820 | 679 | |
| 11891 | 680 | (** local definitions **) | 
| 681 | ||
| 682 | fun gen_def fix prep_term prep_pats raw_name atts (x, (raw_rhs, raw_pats)) state = | |
| 683 | let | |
| 684 | val _ = assert_forward state; | |
| 685 | val ctxt = context_of state; | |
| 686 | ||
| 687 | (*rhs*) | |
| 688 | val name = if raw_name = "" then Thm.def_name x else raw_name; | |
| 689 | val rhs = prep_term ctxt raw_rhs; | |
| 690 | val T = Term.fastype_of rhs; | |
| 691 | val pats = prep_pats T (ProofContext.declare_term rhs ctxt) raw_pats; | |
| 692 | ||
| 693 | (*lhs*) | |
| 694 | val lhs = ProofContext.bind_skolem ctxt [x] (Free (x, T)); | |
| 695 | in | |
| 696 | state | |
| 15531 | 697 | |> fix [([x], NONE)] | 
| 11891 | 698 | |> match_bind_i [(pats, rhs)] | 
| 11917 | 699 | |> assm_i ProofContext.export_def [((name, atts), [(Logic.mk_equals (lhs, rhs), ([], []))])] | 
| 11891 | 700 | end; | 
| 701 | ||
| 702 | val def = gen_def fix ProofContext.read_term ProofContext.read_term_pats; | |
| 703 | val def_i = gen_def fix_i ProofContext.cert_term ProofContext.cert_term_pats; | |
| 704 | ||
| 705 | ||
| 8374 | 706 | (* invoke_case *) | 
| 707 | ||
| 11793 
5f0ab6f5c280
support impromptu terminology of cases parameters;
 wenzelm parents: 
11742diff
changeset | 708 | fun invoke_case (name, xs, atts) state = | 
| 9292 | 709 | let | 
| 10830 | 710 | val (state', (lets, asms)) = | 
| 11793 
5f0ab6f5c280
support impromptu terminology of cases parameters;
 wenzelm parents: 
11742diff
changeset | 711 | map_context_result (ProofContext.apply_case (get_case state name xs)) state; | 
| 13425 
119ae829ad9b
support for split assumptions in cases (hyps vs. prems);
 wenzelm parents: 
13415diff
changeset | 712 | val assumptions = asms |> map (fn (a, ts) => | 
| 
119ae829ad9b
support for split assumptions in cases (hyps vs. prems);
 wenzelm parents: 
13415diff
changeset | 713 | ((if a = "" then "" else NameSpace.append name a, atts), map (rpair ([], [])) ts)); | 
| 
119ae829ad9b
support for split assumptions in cases (hyps vs. prems);
 wenzelm parents: 
13415diff
changeset | 714 | val qnames = filter_out (equal "") (map (#1 o #1) assumptions); | 
| 9292 | 715 | in | 
| 10830 | 716 | state' | 
| 10809 | 717 | |> add_binds_i lets | 
| 13425 
119ae829ad9b
support for split assumptions in cases (hyps vs. prems);
 wenzelm parents: 
13415diff
changeset | 718 | |> map_context (ProofContext.qualified true) | 
| 
119ae829ad9b
support for split assumptions in cases (hyps vs. prems);
 wenzelm parents: 
13415diff
changeset | 719 | |> assume_i assumptions | 
| 
119ae829ad9b
support for split assumptions in cases (hyps vs. prems);
 wenzelm parents: 
13415diff
changeset | 720 | |> map_context (ProofContext.hide_thms false qnames) | 
| 14564 | 721 | |> (fn st => simple_note_thms name (the_facts st) st) | 
| 13425 
119ae829ad9b
support for split assumptions in cases (hyps vs. prems);
 wenzelm parents: 
13415diff
changeset | 722 | |> map_context (ProofContext.restore_qualified (context_of state)) | 
| 8374 | 723 | end; | 
| 724 | ||
| 725 | ||
| 5820 | 726 | |
| 727 | (** goals **) | |
| 728 | ||
| 729 | (* forward chaining *) | |
| 730 | ||
| 731 | fun chain state = | |
| 732 | state | |
| 733 | |> assert_forward | |
| 6848 | 734 | |> assert_facts | 
| 5820 | 735 | |> enter_forward_chain; | 
| 736 | ||
| 737 | fun from_facts facts state = | |
| 738 | state | |
| 15531 | 739 | |> put_facts (SOME facts) | 
| 5820 | 740 | |> chain; | 
| 741 | ||
| 742 | ||
| 743 | (* setup goals *) | |
| 744 | ||
| 11549 | 745 | val rule_contextN = "rule_context"; | 
| 8991 | 746 | |
| 12971 | 747 | fun setup_goal opt_block prepp autofix kind after_qed pr names raw_propp state = | 
| 5820 | 748 | let | 
| 12146 | 749 | val (state', (propss, gen_binds)) = | 
| 5936 | 750 | state | 
| 751 | |> assert_forward_or_chain | |
| 752 | |> enter_forward | |
| 8617 
33e2bd53aec3
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8582diff
changeset | 753 | |> opt_block | 
| 12534 | 754 | |> map_context_result (fn ctxt => prepp (ctxt, raw_propp)); | 
| 12808 
a529b4b91888
RuleCases.make interface based on term instead of thm;
 wenzelm parents: 
12760diff
changeset | 755 | val sign' = sign_of state'; | 
| 12146 | 756 | |
| 15570 | 757 | val props = List.concat propss; | 
| 12808 
a529b4b91888
RuleCases.make interface based on term instead of thm;
 wenzelm parents: 
12760diff
changeset | 758 | val cprop = Thm.cterm_of sign' (Logic.mk_conjunction_list props); | 
| 7556 | 759 | val goal = Drule.mk_triv_goal cprop; | 
| 10553 | 760 | |
| 15574 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 761 | val tvars = foldr Term.add_term_tvars [] props; | 
| 
b1d1b5bfc464
Removed practically all references to Library.foldr.
 skalberg parents: 
15570diff
changeset | 762 | val vars = foldr Term.add_term_vars [] props; | 
| 5820 | 763 | in | 
| 10553 | 764 | if null tvars then () | 
| 765 |     else raise STATE ("Goal statement contains illegal schematic type variable(s): " ^
 | |
| 766 | commas (map (Syntax.string_of_vname o #1) tvars), state); | |
| 767 | if null vars then () | |
| 768 |     else warning ("Goal statement contains unbound schematic variable(s): " ^
 | |
| 12055 | 769 | commas (map (ProofContext.string_of_term (context_of state')) vars)); | 
| 5936 | 770 | state' | 
| 12146 | 771 | |> map_context (autofix props) | 
| 15531 | 772 | |> put_goal (SOME (((kind, names, propss), ([], goal)), | 
| 12971 | 773 | (after_qed o map_context gen_binds, pr))) | 
| 12167 | 774 | |> map_context (ProofContext.add_cases | 
| 12808 
a529b4b91888
RuleCases.make interface based on term instead of thm;
 wenzelm parents: 
12760diff
changeset | 775 | (if length props = 1 then | 
| 15531 | 776 | RuleCases.make true NONE (Thm.sign_of_thm goal, Thm.prop_of goal) [rule_contextN] | 
| 12167 | 777 | else [(rule_contextN, RuleCases.empty)])) | 
| 12146 | 778 | |> auto_bind_goal props | 
| 5820 | 779 | |> (if is_chain state then use_facts else reset_facts) | 
| 780 | |> new_block | |
| 781 | |> enter_backward | |
| 782 | end; | |
| 783 | ||
| 784 | (*global goals*) | |
| 12959 | 785 | fun global_goal prep kind raw_locale a elems concl thy = | 
| 12503 | 786 | let | 
| 12549 | 787 | val init = init_state thy; | 
| 13399 
c136276dc847
support locale ``views'' (for cumulative predicates);
 wenzelm parents: 
13377diff
changeset | 788 | val (opt_name, view, locale_ctxt, elems_ctxt, propp) = | 
| 15570 | 789 | prep (Option.map fst raw_locale) elems (map snd concl) (context_of init); | 
| 12503 | 790 | in | 
| 12549 | 791 | init | 
| 792 | |> open_block | |
| 12511 | 793 | |> map_context (K locale_ctxt) | 
| 12065 | 794 | |> open_block | 
| 12511 | 795 | |> map_context (K elems_ctxt) | 
| 15206 
09d78ec709c7
Modified locales: improved implementation of "includes".
 ballarin parents: 
15194diff
changeset | 796 | (* CB: elems_ctxt is an extension of locale_ctxt, see prep_context_statement | 
| 
09d78ec709c7
Modified locales: improved implementation of "includes".
 ballarin parents: 
15194diff
changeset | 797 | in locale.ML, naming there: ctxt and import_ctxt. *) | 
| 12959 | 798 | |> setup_goal I ProofContext.bind_propp_schematic_i ProofContext.fix_frees | 
| 13415 | 799 |       (Theorem {kind = kind,
 | 
| 800 | theory_spec = (a, map (snd o fst) concl), | |
| 15570 | 801 | locale_spec = (case raw_locale of NONE => NONE | SOME (_, x) => SOME (valOf opt_name, x)), | 
| 13415 | 802 | view = view}) | 
| 12971 | 803 | Seq.single true (map (fst o fst) concl) propp | 
| 12065 | 804 | end; | 
| 5820 | 805 | |
| 12534 | 806 | val multi_theorem = global_goal Locale.read_context_statement; | 
| 807 | val multi_theorem_i = global_goal Locale.cert_context_statement; | |
| 12146 | 808 | |
| 5820 | 809 | |
| 810 | (*local goals*) | |
| 12971 | 811 | fun local_goal' prepp kind (check: bool -> state -> state) f pr args int state = | 
| 7928 | 812 | state | 
| 12146 | 813 | |> setup_goal open_block prepp (K I) (kind (map (snd o fst) args)) | 
| 12971 | 814 | f pr (map (fst o fst) args) (map snd args) | 
| 10936 | 815 | |> warn_extra_tfrees state | 
| 816 | |> check int; | |
| 5820 | 817 | |
| 12971 | 818 | fun local_goal prepp kind f pr args = local_goal' prepp kind (K I) f pr args false; | 
| 10936 | 819 | |
| 820 | val show = local_goal' ProofContext.bind_propp Show; | |
| 821 | val show_i = local_goal' ProofContext.bind_propp_i Show; | |
| 822 | val have = local_goal ProofContext.bind_propp Have; | |
| 10380 | 823 | val have_i = local_goal ProofContext.bind_propp_i Have; | 
| 5820 | 824 | |
| 825 | ||
| 826 | ||
| 827 | (** conclusions **) | |
| 828 | ||
| 829 | (* current goal *) | |
| 830 | ||
| 15531 | 831 | fun current_goal (State (Node {context, goal = SOME goal, ...}, _)) = (context, goal)
 | 
| 5820 | 832 |   | current_goal state = raise STATE ("No current goal!", state);
 | 
| 833 | ||
| 15531 | 834 | fun assert_current_goal true (state as State (Node {goal = NONE, ...}, _)) =
 | 
| 6404 | 835 |       raise STATE ("No goal in this block!", state)
 | 
| 15531 | 836 |   | assert_current_goal false (state as State (Node {goal = SOME _, ...}, _)) =
 | 
| 6404 | 837 |       raise STATE ("Goal present in this block!", state)
 | 
| 838 | | assert_current_goal _ state = state; | |
| 5820 | 839 | |
| 12010 | 840 | fun assert_bottom b (state as State (_, nodes)) = | 
| 841 | let val bot = (length nodes <= 2) in | |
| 842 |     if b andalso not bot then raise STATE ("Not at bottom of proof!", state)
 | |
| 843 |     else if not b andalso bot then raise STATE ("Already at bottom of proof!", state)
 | |
| 844 | else state | |
| 845 | end; | |
| 5820 | 846 | |
| 6404 | 847 | val at_bottom = can (assert_bottom true o close_block); | 
| 848 | ||
| 6932 | 849 | fun end_proof bot state = | 
| 5820 | 850 | state | 
| 851 | |> assert_forward | |
| 852 | |> close_block | |
| 853 | |> assert_bottom bot | |
| 7011 | 854 | |> assert_current_goal true | 
| 855 | |> goal_facts (K []); | |
| 5820 | 856 | |
| 857 | ||
| 6404 | 858 | (* local_qed *) | 
| 5820 | 859 | |
| 6982 | 860 | fun finish_local (print_result, print_rule) state = | 
| 5820 | 861 | let | 
| 12971 | 862 | val (goal_ctxt, (((kind, names, tss), (_, raw_thm)), (after_qed, pr))) = current_goal state; | 
| 8617 
33e2bd53aec3
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8582diff
changeset | 863 | val outer_state = state |> close_block; | 
| 
33e2bd53aec3
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8582diff
changeset | 864 | val outer_ctxt = context_of outer_state; | 
| 
33e2bd53aec3
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8582diff
changeset | 865 | |
| 15570 | 866 | val ts = List.concat tss; | 
| 12146 | 867 | val results = prep_result state ts raw_thm; | 
| 868 | val resultq = | |
| 869 | results |> map (ProofContext.export false goal_ctxt outer_ctxt) | |
| 870 | |> Seq.commute |> Seq.map (Library.unflat tss); | |
| 9469 | 871 | |
| 12146 | 872 | val (attss, opt_solve) = | 
| 5820 | 873 | (case kind of | 
| 12971 | 874 | Show attss => (attss, | 
| 875 | export_goals goal_ctxt (if pr then print_rule else (K (K ()))) results) | |
| 12146 | 876 | | Have attss => (attss, Seq.single) | 
| 6932 | 877 | | _ => err_malformed "finish_local" state); | 
| 5820 | 878 | in | 
| 12971 | 879 | conditional pr (fn () => print_result goal_ctxt | 
| 880 | (kind_name (sign_of state) kind, names ~~ Library.unflat tss results)); | |
| 8617 
33e2bd53aec3
support Hindley-Milner polymorphisms in binds and facts;
 wenzelm parents: 
8582diff
changeset | 881 | outer_state | 
| 12549 | 882 | |> auto_bind_facts (ProofContext.generalize goal_ctxt outer_ctxt ts) | 
| 9469 | 883 | |> (fn st => Seq.map (fn res => | 
| 14564 | 884 | note_thmss_i ((names ~~ attss) ~~ map (single o Thm.no_attributes) res) st) resultq) | 
| 9469 | 885 | |> (Seq.flat o Seq.map opt_solve) | 
| 7176 | 886 | |> (Seq.flat o Seq.map after_qed) | 
| 5820 | 887 | end; | 
| 888 | ||
| 6982 | 889 | fun local_qed finalize print state = | 
| 6404 | 890 | state | 
| 6932 | 891 | |> end_proof false | 
| 6871 | 892 | |> finalize | 
| 6982 | 893 | |> (Seq.flat o Seq.map (finish_local print)); | 
| 5820 | 894 | |
| 895 | ||
| 12703 
af5fec8a418f
refine_tac: Tactic.norm_hhf_tac before trying rule;
 wenzelm parents: 
12698diff
changeset | 896 | (* global_qed *) | 
| 12065 | 897 | |
| 6950 | 898 | fun finish_global state = | 
| 5820 | 899 | let | 
| 12146 | 900 | val (goal_ctxt, (((kind, names, tss), (_, raw_thm)), _)) = current_goal state; | 
| 12065 | 901 | val locale_ctxt = context_of (state |> close_block); | 
| 902 | val theory_ctxt = context_of (state |> close_block |> close_block); | |
| 15206 
09d78ec709c7
Modified locales: improved implementation of "includes".
 ballarin parents: 
15194diff
changeset | 903 |     val {kind = k, theory_spec = ((name, atts), attss), locale_spec, view = (target_view, elems_view)} =
 | 
| 13399 
c136276dc847
support locale ``views'' (for cumulative predicates);
 wenzelm parents: 
13377diff
changeset | 904 | (case kind of Theorem x => x | _ => err_malformed "finish_global" state); | 
| 5820 | 905 | |
| 15570 | 906 | val ts = List.concat tss; | 
| 15206 
09d78ec709c7
Modified locales: improved implementation of "includes".
 ballarin parents: 
15194diff
changeset | 907 | val locale_results = map (ProofContext.export_standard elems_view goal_ctxt locale_ctxt) | 
| 13377 | 908 | (prep_result state ts raw_thm); | 
| 14215 
ebf291f3b449
Improvements to Isar/Locales:  premises generated by "includes" elements
 ballarin parents: 
14129diff
changeset | 909 | |
| 13377 | 910 | val results = map (Drule.strip_shyps_warning o | 
| 15206 
09d78ec709c7
Modified locales: improved implementation of "includes".
 ballarin parents: 
15194diff
changeset | 911 | ProofContext.export_standard target_view locale_ctxt theory_ctxt) locale_results; | 
| 12703 
af5fec8a418f
refine_tac: Tactic.norm_hhf_tac before trying rule;
 wenzelm parents: 
12698diff
changeset | 912 | |
| 12959 | 913 | val (theory', results') = | 
| 914 | theory_of state | |
| 15531 | 915 | |> (case locale_spec of NONE => I | SOME (loc, (loc_atts, loc_attss)) => fn thy => | 
| 12959 | 916 | if length names <> length loc_attss then | 
| 917 |           raise THEORY ("Bad number of locale attributes", [thy])
 | |
| 13335 | 918 | else (thy, locale_ctxt) | 
| 15206 
09d78ec709c7
Modified locales: improved implementation of "includes".
 ballarin parents: 
15194diff
changeset | 919 | |> Locale.add_thmss loc ((names ~~ Library.unflat tss locale_results) ~~ loc_attss) | 
| 13335 | 920 | |> (fn ((thy', ctxt'), res) => | 
| 12959 | 921 | if name = "" andalso null loc_atts then thy' | 
| 13335 | 922 | else (thy', ctxt') | 
| 15570 | 923 | |> (#1 o #1 o Locale.add_thmss loc [((name, List.concat (map #2 res)), loc_atts)]))) | 
| 14564 | 924 | |> Locale.smart_note_thmss k locale_spec | 
| 12703 
af5fec8a418f
refine_tac: Tactic.norm_hhf_tac before trying rule;
 wenzelm parents: 
12698diff
changeset | 925 | ((names ~~ attss) ~~ map (single o Thm.no_attributes) (Library.unflat tss results)) | 
| 12959 | 926 | |> (fn (thy, res) => (thy, res) | 
| 14564 | 927 | |>> (#1 o Locale.smart_note_thmss k locale_spec | 
| 15570 | 928 | [((name, atts), [(List.concat (map #2 res), [])])])); | 
| 12959 | 929 | in (theory', ((k, name), results')) end; | 
| 930 | ||
| 5820 | 931 | |
| 10553 | 932 | (*note: clients should inspect first result only, as backtracking may destroy theory*) | 
| 6950 | 933 | fun global_qed finalize state = | 
| 5820 | 934 | state | 
| 6932 | 935 | |> end_proof true | 
| 6871 | 936 | |> finalize | 
| 12065 | 937 | |> Seq.map finish_global; | 
| 5820 | 938 | |
| 939 | ||
| 6896 | 940 | |
| 941 | (** blocks **) | |
| 942 | ||
| 943 | (* begin_block *) | |
| 944 | ||
| 945 | fun begin_block state = | |
| 946 | state | |
| 947 | |> assert_forward | |
| 948 | |> new_block | |
| 949 | |> open_block; | |
| 950 | ||
| 951 | ||
| 952 | (* end_block *) | |
| 953 | ||
| 954 | fun end_block state = | |
| 955 | state | |
| 956 | |> assert_forward | |
| 957 | |> close_block | |
| 958 | |> assert_current_goal false | |
| 959 | |> close_block | |
| 6932 | 960 | |> transfer_facts state; | 
| 6896 | 961 | |
| 962 | ||
| 963 | (* next_block *) | |
| 964 | ||
| 965 | fun next_block state = | |
| 966 | state | |
| 967 | |> assert_forward | |
| 968 | |> close_block | |
| 8718 | 969 | |> new_block | 
| 970 | |> reset_facts; | |
| 6896 | 971 | |
| 972 | ||
| 5820 | 973 | end; | 
| 8152 | 974 | |
| 975 | structure BasicProof: BASIC_PROOF = Proof; | |
| 976 | open BasicProof; |