src/Pure/Isar/proof.ML
author wenzelm
Sun, 18 Mar 2012 13:04:22 +0100
changeset 47876 421760a1efe7
parent 47602 85f8e3932712
child 47937 e2741ec9ae36
permissions -rw-r--r--
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
more explicit Context.generic for Name_Space.declare/define and derivatives (NB: naming changed after Proof_Context.init_global);
prefer Context.pretty in low-level operations of structure Sorts/Type (defer full Syntax.init_pretty until error output);
simplified signatures;
     1 (*  Title:      Pure/Isar/proof.ML
     2     Author:     Markus Wenzel, TU Muenchen
     3 
     4 The Isar/VM proof language interpreter: maintains a structured flow of
     5 context elements, goals, refinements, and facts.
     6 *)
     7 
     8 signature PROOF =
     9 sig
    10   type context = Proof.context
    11   type method = Method.method
    12   type state
    13   val init: context -> state
    14   val level: state -> int
    15   val assert_bottom: bool -> state -> state
    16   val context_of: state -> context
    17   val theory_of: state -> theory
    18   val map_context: (context -> context) -> state -> state
    19   val map_context_result : (context -> 'a * context) -> state -> 'a * state
    20   val map_contexts: (context -> context) -> state -> state
    21   val propagate_ml_env: state -> state
    22   val bind_terms: (indexname * term option) list -> state -> state
    23   val put_thms: bool -> string * thm list option -> state -> state
    24   val the_facts: state -> thm list
    25   val the_fact: state -> thm
    26   val put_facts: thm list option -> state -> state
    27   val assert_forward: state -> state
    28   val assert_chain: state -> state
    29   val assert_forward_or_chain: state -> state
    30   val assert_backward: state -> state
    31   val assert_no_chain: state -> state
    32   val enter_forward: state -> state
    33   val goal_message: (unit -> Pretty.T) -> state -> state
    34   val pretty_state: int -> state -> Pretty.T list
    35   val pretty_goals: bool -> state -> Pretty.T list
    36   val refine: Method.text -> state -> state Seq.seq
    37   val refine_end: Method.text -> state -> state Seq.seq
    38   val refine_insert: thm list -> state -> state
    39   val refine_goals: (context -> thm -> unit) -> context -> thm list -> state -> state Seq.seq
    40   val raw_goal: state -> {context: context, facts: thm list, goal: thm}
    41   val goal: state -> {context: context, facts: thm list, goal: thm}
    42   val simple_goal: state -> {context: context, goal: thm}
    43   val status_markup: state -> Markup.T
    44   val let_bind: (term list * term) list -> state -> state
    45   val let_bind_cmd: (string list * string) list -> state -> state
    46   val write: Syntax.mode -> (term * mixfix) list -> state -> state
    47   val write_cmd: Syntax.mode -> (string * mixfix) list -> state -> state
    48   val fix: (binding * typ option * mixfix) list -> state -> state
    49   val fix_cmd: (binding * string option * mixfix) list -> state -> state
    50   val assm: Assumption.export ->
    51     (Thm.binding * (term * term list) list) list -> state -> state
    52   val assm_cmd: Assumption.export ->
    53     (Attrib.binding * (string * string list) list) list -> state -> state
    54   val assume: (Thm.binding * (term * term list) list) list -> state -> state
    55   val assume_cmd: (Attrib.binding * (string * string list) list) list -> state -> state
    56   val presume: (Thm.binding * (term * term list) list) list -> state -> state
    57   val presume_cmd: (Attrib.binding * (string * string list) list) list -> state -> state
    58   val def: (Thm.binding * ((binding * mixfix) * (term * term list))) list -> state -> state
    59   val def_cmd: (Attrib.binding * ((binding * mixfix) * (string * string list))) list -> state -> state
    60   val chain: state -> state
    61   val chain_facts: thm list -> state -> state
    62   val note_thmss: (Thm.binding * (thm list * attribute list) list) list -> state -> state
    63   val note_thmss_cmd: (Attrib.binding * (Facts.ref * Attrib.src list) list) list -> state -> state
    64   val from_thmss: ((thm list * attribute list) list) list -> state -> state
    65   val from_thmss_cmd: ((Facts.ref * Attrib.src list) list) list -> state -> state
    66   val with_thmss: ((thm list * attribute list) list) list -> state -> state
    67   val with_thmss_cmd: ((Facts.ref * Attrib.src list) list) list -> state -> state
    68   val using: ((thm list * attribute list) list) list -> state -> state
    69   val using_cmd: ((Facts.ref * Attrib.src list) list) list -> state -> state
    70   val unfolding: ((thm list * attribute list) list) list -> state -> state
    71   val unfolding_cmd: ((Facts.ref * Attrib.src list) list) list -> state -> state
    72   val invoke_case: string * binding option list * attribute list -> state -> state
    73   val invoke_case_cmd: string * binding option list * Attrib.src list -> state -> state
    74   val begin_block: state -> state
    75   val next_block: state -> state
    76   val end_block: state -> state
    77   val begin_notepad: Proof.context -> state
    78   val end_notepad: state -> Proof.context
    79   val proof: Method.text option -> state -> state Seq.seq
    80   val defer: int option -> state -> state Seq.seq
    81   val prefer: int -> state -> state Seq.seq
    82   val apply: Method.text -> state -> state Seq.seq
    83   val apply_end: Method.text -> state -> state Seq.seq
    84   val local_goal: (context -> ((string * string) * (string * thm list) list) -> unit) ->
    85     (theory -> 'a -> attribute) ->
    86     ('b list -> context -> (term list list * (context -> context)) * context) ->
    87     string -> Method.text option -> (thm list list -> state -> state) ->
    88     ((binding * 'a list) * 'b) list -> state -> state
    89   val local_qed: Method.text option * bool -> state -> state
    90   val theorem: Method.text option -> (thm list list -> context -> context) ->
    91     (term * term list) list list -> context -> state
    92   val theorem_cmd: Method.text option -> (thm list list -> context -> context) ->
    93     (string * string list) list list -> context -> state
    94   val global_qed: Method.text option * bool -> state -> context
    95   val local_terminal_proof: Method.text * Method.text option -> state -> state
    96   val local_default_proof: state -> state
    97   val local_immediate_proof: state -> state
    98   val local_skip_proof: bool -> state -> state
    99   val local_done_proof: state -> state
   100   val global_terminal_proof: Method.text * Method.text option -> state -> context
   101   val global_default_proof: state -> context
   102   val global_immediate_proof: state -> context
   103   val global_skip_proof: bool -> state -> context
   104   val global_done_proof: state -> context
   105   val have: Method.text option -> (thm list list -> state -> state) ->
   106     (Thm.binding * (term * term list) list) list -> bool -> state -> state
   107   val have_cmd: Method.text option -> (thm list list -> state -> state) ->
   108     (Attrib.binding * (string * string list) list) list -> bool -> state -> state
   109   val show: Method.text option -> (thm list list -> state -> state) ->
   110     (Thm.binding * (term * term list) list) list -> bool -> state -> state
   111   val show_cmd: Method.text option -> (thm list list -> state -> state) ->
   112     (Attrib.binding * (string * string list) list) list -> bool -> state -> state
   113   val schematic_goal: state -> bool
   114   val is_relevant: state -> bool
   115   val local_future_proof: (state -> ('a * state) Future.future) ->
   116     state -> 'a Future.future * state
   117   val global_future_proof: (state -> ('a * Proof.context) Future.future) ->
   118     state -> 'a Future.future * context
   119   val local_future_terminal_proof: Method.text * Method.text option -> bool -> state -> state
   120   val global_future_terminal_proof: Method.text * Method.text option -> bool -> state -> context
   121 end;
   122 
   123 structure Proof: PROOF =
   124 struct
   125 
   126 type context = Proof.context;
   127 type method = Method.method;
   128 
   129 
   130 (** proof state **)
   131 
   132 (* datatype state *)
   133 
   134 datatype mode = Forward | Chain | Backward;
   135 
   136 datatype state =
   137   State of node Stack.T
   138 and node =
   139   Node of
   140    {context: context,
   141     facts: thm list option,
   142     mode: mode,
   143     goal: goal option}
   144 and goal =
   145   Goal of
   146    {statement: (string * Position.T) * term list list * term,
   147       (*goal kind and statement (starting with vars), initial proposition*)
   148     messages: (unit -> Pretty.T) list,    (*persistent messages (hints etc.)*)
   149     using: thm list,                      (*goal facts*)
   150     goal: thm,                            (*subgoals ==> statement*)
   151     before_qed: Method.text option,
   152     after_qed:
   153       (thm list list -> state -> state) *
   154       (thm list list -> context -> context)};
   155 
   156 fun make_goal (statement, messages, using, goal, before_qed, after_qed) =
   157   Goal {statement = statement, messages = messages, using = using, goal = goal,
   158     before_qed = before_qed, after_qed = after_qed};
   159 
   160 fun make_node (context, facts, mode, goal) =
   161   Node {context = context, facts = facts, mode = mode, goal = goal};
   162 
   163 fun map_node f (Node {context, facts, mode, goal}) =
   164   make_node (f (context, facts, mode, goal));
   165 
   166 val init_context =
   167   Proof_Context.set_stmt true #>
   168   Proof_Context.map_naming (K Name_Space.local_naming);
   169 
   170 fun init ctxt =
   171   State (Stack.init (make_node (init_context ctxt, NONE, Forward, NONE)));
   172 
   173 fun current (State st) = Stack.top st |> (fn Node node => node);
   174 fun map_current f (State st) = State (Stack.map_top (map_node f) st);
   175 fun map_all f (State st) = State (Stack.map_all (map_node f) st);
   176 
   177 
   178 
   179 (** basic proof state operations **)
   180 
   181 (* block structure *)
   182 
   183 fun open_block (State st) = State (Stack.push st);
   184 
   185 fun close_block (State st) = State (Stack.pop st)
   186   handle Empty => error "Unbalanced block parentheses";
   187 
   188 fun level (State st) = Stack.level st;
   189 
   190 fun assert_bottom b state =
   191   let val b' = (level state <= 2) in
   192     if b andalso not b' then error "Not at bottom of proof!"
   193     else if not b andalso b' then error "Already at bottom of proof!"
   194     else state
   195   end;
   196 
   197 
   198 (* context *)
   199 
   200 val context_of = #context o current;
   201 val theory_of = Proof_Context.theory_of o context_of;
   202 
   203 fun map_context f =
   204   map_current (fn (ctxt, facts, mode, goal) => (f ctxt, facts, mode, goal));
   205 
   206 fun map_context_result f state =
   207   f (context_of state) ||> (fn ctxt => map_context (K ctxt) state);
   208 
   209 fun map_contexts f = map_all (fn (ctxt, facts, mode, goal) => (f ctxt, facts, mode, goal));
   210 
   211 fun propagate_ml_env state = map_contexts
   212   (Context.proof_map (ML_Env.inherit (Context.Proof (context_of state)))) state;
   213 
   214 val bind_terms = map_context o Proof_Context.bind_terms;
   215 val put_thms = map_context oo Proof_Context.put_thms;
   216 
   217 
   218 (* facts *)
   219 
   220 val get_facts = #facts o current;
   221 
   222 fun the_facts state =
   223   (case get_facts state of SOME facts => facts
   224   | NONE => error "No current facts available");
   225 
   226 fun the_fact state =
   227   (case the_facts state of [thm] => thm
   228   | _ => error "Single theorem expected");
   229 
   230 fun put_facts facts =
   231   map_current (fn (ctxt, _, mode, goal) => (ctxt, facts, mode, goal)) #>
   232   put_thms true (Auto_Bind.thisN, facts);
   233 
   234 fun these_factss more_facts (named_factss, state) =
   235   (named_factss, state |> put_facts (SOME (maps snd named_factss @ more_facts)));
   236 
   237 fun export_facts inner outer =
   238   (case get_facts inner of
   239     NONE => put_facts NONE outer
   240   | SOME thms =>
   241       thms
   242       |> Proof_Context.export (context_of inner) (context_of outer)
   243       |> (fn ths => put_facts (SOME ths) outer));
   244 
   245 
   246 (* mode *)
   247 
   248 val get_mode = #mode o current;
   249 fun put_mode mode = map_current (fn (ctxt, facts, _, goal) => (ctxt, facts, mode, goal));
   250 
   251 val mode_name = (fn Forward => "state" | Chain => "chain" | Backward => "prove");
   252 
   253 fun assert_mode pred state =
   254   let val mode = get_mode state in
   255     if pred mode then state
   256     else error ("Illegal application of proof command in " ^ quote (mode_name mode) ^ " mode")
   257   end;
   258 
   259 val assert_forward = assert_mode (fn mode => mode = Forward);
   260 val assert_chain = assert_mode (fn mode => mode = Chain);
   261 val assert_forward_or_chain = assert_mode (fn mode => mode = Forward orelse mode = Chain);
   262 val assert_backward = assert_mode (fn mode => mode = Backward);
   263 val assert_no_chain = assert_mode (fn mode => mode <> Chain);
   264 
   265 val enter_forward = put_mode Forward;
   266 val enter_chain = put_mode Chain;
   267 val enter_backward = put_mode Backward;
   268 
   269 
   270 (* current goal *)
   271 
   272 fun current_goal state =
   273   (case current state of
   274     {context, goal = SOME (Goal goal), ...} => (context, goal)
   275   | _ => error "No current goal!");
   276 
   277 fun assert_current_goal g state =
   278   let val g' = can current_goal state in
   279     if g andalso not g' then error "No goal in this block!"
   280     else if not g andalso g' then error "Goal present in this block!"
   281     else state
   282   end;
   283 
   284 fun put_goal goal = map_current (fn (ctxt, using, mode, _) => (ctxt, using, mode, goal));
   285 
   286 val before_qed = #before_qed o #2 o current_goal;
   287 
   288 
   289 (* nested goal *)
   290 
   291 fun map_goal f g (State (Node {context, facts, mode, goal = SOME goal}, nodes)) =
   292       let
   293         val Goal {statement, messages, using, goal, before_qed, after_qed} = goal;
   294         val goal' = make_goal (g (statement, messages, using, goal, before_qed, after_qed));
   295       in State (make_node (f context, facts, mode, SOME goal'), nodes) end
   296   | map_goal f g (State (nd, node :: nodes)) =
   297       let val State (node', nodes') = map_goal f g (State (node, nodes))
   298       in map_context f (State (nd, node' :: nodes')) end
   299   | map_goal _ _ state = state;
   300 
   301 fun set_goal goal = map_goal I (fn (statement, _, using, _, before_qed, after_qed) =>
   302   (statement, [], using, goal, before_qed, after_qed));
   303 
   304 fun goal_message msg = map_goal I (fn (statement, messages, using, goal, before_qed, after_qed) =>
   305   (statement, msg :: messages, using, goal, before_qed, after_qed));
   306 
   307 fun using_facts using = map_goal I (fn (statement, _, _, goal, before_qed, after_qed) =>
   308   (statement, [], using, goal, before_qed, after_qed));
   309 
   310 local
   311   fun find i state =
   312     (case try current_goal state of
   313       SOME (ctxt, goal) => (ctxt, (i, goal))
   314     | NONE => find (i + 1) (close_block state handle ERROR _ => error "No goal present"));
   315 in val find_goal = find 0 end;
   316 
   317 fun get_goal state =
   318   let val (ctxt, (_, {using, goal, ...})) = find_goal state
   319   in (ctxt, (using, goal)) end;
   320 
   321 
   322 
   323 (** pretty_state **)
   324 
   325 fun pretty_facts _ _ NONE = []
   326   | pretty_facts s ctxt (SOME ths) =
   327       [Pretty.big_list (s ^ "this:") (map (Display.pretty_thm ctxt) ths), Pretty.str ""];
   328 
   329 fun pretty_state nr state =
   330   let
   331     val {context = ctxt, facts, mode, goal = _} = current state;
   332     val verbose = Config.get ctxt Proof_Context.verbose;
   333 
   334     fun levels_up 0 = ""
   335       | levels_up 1 = "1 level up"
   336       | levels_up i = string_of_int i ^ " levels up";
   337 
   338     fun subgoals 0 = ""
   339       | subgoals 1 = "1 subgoal"
   340       | subgoals n = string_of_int n ^ " subgoals";
   341 
   342     fun description strs =
   343       (case filter_out (fn s => s = "") strs of [] => ""
   344       | strs' => enclose " (" ")" (commas strs'));
   345 
   346     fun prt_goal (SOME (_, (i,
   347       {statement = ((_, pos), _, _), messages, using, goal, before_qed = _, after_qed = _}))) =
   348           pretty_facts "using " ctxt
   349             (if mode <> Backward orelse null using then NONE else SOME using) @
   350           [Pretty.str ("goal" ^
   351             description [levels_up (i div 2), subgoals (Thm.nprems_of goal)] ^ ":")] @
   352           Goal_Display.pretty_goals ctxt goal @
   353           (map (fn msg => Position.setmp_thread_data pos msg ()) (rev messages))
   354       | prt_goal NONE = [];
   355 
   356     val prt_ctxt =
   357       if verbose orelse mode = Forward then Proof_Context.pretty_context ctxt
   358       else if mode = Backward then Proof_Context.pretty_ctxt ctxt
   359       else [];
   360   in
   361     [Pretty.str ("proof (" ^ mode_name mode ^ "): step " ^ string_of_int nr ^
   362       (if verbose then ", depth " ^ string_of_int (level state div 2 - 1) else "")),
   363       Pretty.str ""] @
   364     (if null prt_ctxt then [] else prt_ctxt @ [Pretty.str ""]) @
   365     (if verbose orelse mode = Forward then
   366        pretty_facts "" ctxt facts @ prt_goal (try find_goal state)
   367      else if mode = Chain then pretty_facts "picking " ctxt facts
   368      else prt_goal (try find_goal state))
   369   end;
   370 
   371 fun pretty_goals main state =
   372   let
   373     val (_, (_, goal)) = get_goal state;
   374     val ctxt = context_of state
   375       |> Config.put Goal_Display.show_main_goal main
   376       |> Config.put Goal_Display.goals_total false;
   377   in Goal_Display.pretty_goals ctxt goal end;
   378 
   379 
   380 
   381 (** proof steps **)
   382 
   383 (* refine via method *)
   384 
   385 local
   386 
   387 fun goalN i = "goal" ^ string_of_int i;
   388 fun goals st = map goalN (1 upto Thm.nprems_of st);
   389 
   390 fun no_goal_cases st = map (rpair NONE) (goals st);
   391 
   392 fun goal_cases st =
   393   Rule_Cases.make_common (Thm.theory_of_thm st, Thm.prop_of st) (map (rpair [] o rpair []) (goals st));
   394 
   395 fun apply_method current_context meth state =
   396   let
   397     val (goal_ctxt, (_, {statement, messages = _, using, goal, before_qed, after_qed})) =
   398       find_goal state;
   399     val ctxt = if current_context then context_of state else goal_ctxt;
   400   in
   401     Method.apply meth ctxt using goal |> Seq.map (fn (meth_cases, goal') =>
   402       state
   403       |> map_goal
   404           (Proof_Context.add_cases false (no_goal_cases goal @ goal_cases goal') #>
   405            Proof_Context.add_cases true meth_cases)
   406           (K (statement, [], using, goal', before_qed, after_qed)))
   407   end;
   408 
   409 fun select_goals n meth state =
   410   state
   411   |> (#2 o #2 o get_goal)
   412   |> ALLGOALS Goal.conjunction_tac
   413   |> Seq.maps (fn goal =>
   414     state
   415     |> Seq.lift set_goal (Goal.extract 1 n goal |> Seq.maps (Goal.conjunction_tac 1))
   416     |> Seq.maps meth
   417     |> Seq.maps (fn state' => state'
   418       |> Seq.lift set_goal (Goal.retrofit 1 n (#2 (#2 (get_goal state'))) goal))
   419     |> Seq.maps (apply_method true (K Method.succeed)));
   420 
   421 fun apply_text cc text state =
   422   let
   423     val thy = theory_of state;
   424 
   425     fun eval (Method.Basic m) = apply_method cc m
   426       | eval (Method.Source src) = apply_method cc (Method.method thy src)
   427       | eval (Method.Source_i src) = apply_method cc (Method.method_i thy src)
   428       | eval (Method.Then txts) = Seq.EVERY (map eval txts)
   429       | eval (Method.Orelse txts) = Seq.FIRST (map eval txts)
   430       | eval (Method.Try txt) = Seq.TRY (eval txt)
   431       | eval (Method.Repeat1 txt) = Seq.REPEAT1 (eval txt)
   432       | eval (Method.SelectGoals (n, txt)) = select_goals n (eval txt);
   433   in eval text state end;
   434 
   435 in
   436 
   437 val refine = apply_text true;
   438 val refine_end = apply_text false;
   439 fun refine_insert ths = Seq.hd o refine (Method.Basic (K (Method.insert ths)));
   440 
   441 end;
   442 
   443 
   444 (* refine via sub-proof *)
   445 
   446 local
   447 
   448 fun finish_tac 0 = K all_tac
   449   | finish_tac n =
   450       Goal.norm_hhf_tac THEN'
   451       SUBGOAL (fn (goal, i) =>
   452         if can Logic.unprotect (Logic.strip_assums_concl goal) then
   453           Tactic.etac Drule.protectI i THEN finish_tac (n - 1) i
   454         else finish_tac (n - 1) (i + 1));
   455 
   456 fun goal_tac rule =
   457   Goal.norm_hhf_tac THEN'
   458   Tactic.rtac rule THEN'
   459   finish_tac (Thm.nprems_of rule);
   460 
   461 fun FINDGOAL tac st =
   462   let fun find i n = if i > n then Seq.fail else Seq.APPEND (tac i, find (i + 1) n)
   463   in find 1 (Thm.nprems_of st) st end;
   464 
   465 in
   466 
   467 fun refine_goals print_rule inner raw_rules state =
   468   let
   469     val (outer, (_, goal)) = get_goal state;
   470     fun refine rule st = (print_rule outer rule; FINDGOAL (goal_tac rule) st);
   471   in
   472     raw_rules
   473     |> Proof_Context.goal_export inner outer
   474     |> (fn rules => Seq.lift set_goal (EVERY (map refine rules) goal) state)
   475   end;
   476 
   477 end;
   478 
   479 
   480 (* conclude_goal *)
   481 
   482 fun conclude_goal ctxt goal propss =
   483   let
   484     val thy = Proof_Context.theory_of ctxt;
   485     val string_of_term = Syntax.string_of_term ctxt;
   486     val string_of_thm = Display.string_of_thm ctxt;
   487 
   488     val ngoals = Thm.nprems_of goal;
   489     val _ = ngoals = 0 orelse error (Pretty.string_of (Pretty.chunks
   490       (Goal_Display.pretty_goals ctxt goal @
   491         [Pretty.str (string_of_int ngoals ^ " unsolved goal(s)!")])));
   492 
   493     val extra_hyps = Assumption.extra_hyps ctxt goal;
   494     val _ = null extra_hyps orelse
   495       error ("Additional hypotheses:\n" ^ cat_lines (map string_of_term extra_hyps));
   496 
   497     fun lost_structure () = error ("Lost goal structure:\n" ^ string_of_thm goal);
   498 
   499     val th = Goal.conclude
   500       (if length (flat propss) > 1 then Thm.norm_proof goal else goal)
   501       handle THM _ => lost_structure ();
   502     val goal_propss = filter_out null propss;
   503     val results =
   504       Conjunction.elim_balanced (length goal_propss) th
   505       |> map2 Conjunction.elim_balanced (map length goal_propss)
   506       handle THM _ => lost_structure ();
   507     val _ = Unify.matches_list thy (flat goal_propss) (map Thm.prop_of (flat results)) orelse
   508       error ("Proved a different theorem:\n" ^ string_of_thm th);
   509     val _ = Thm.check_shyps (Variable.sorts_of ctxt) th;
   510 
   511     fun recover_result ([] :: pss) thss = [] :: recover_result pss thss
   512       | recover_result (_ :: pss) (ths :: thss) = ths :: recover_result pss thss
   513       | recover_result [] [] = []
   514       | recover_result _ _ = lost_structure ();
   515   in recover_result propss results end;
   516 
   517 
   518 (* goal views -- corresponding to methods *)
   519 
   520 fun raw_goal state =
   521   let val (ctxt, (facts, goal)) = get_goal state
   522   in {context = ctxt, facts = facts, goal = goal} end;
   523 
   524 val goal = raw_goal o refine_insert [];
   525 
   526 fun simple_goal state =
   527   let
   528     val (_, (facts, _)) = get_goal state;
   529     val (ctxt, (_, goal)) = get_goal (refine_insert facts state);
   530   in {context = ctxt, goal = goal} end;
   531 
   532 fun status_markup state =
   533   (case try goal state of
   534     SOME {goal, ...} => Isabelle_Markup.proof_state (Thm.nprems_of goal)
   535   | NONE => Markup.empty);
   536 
   537 
   538 
   539 (*** structured proof commands ***)
   540 
   541 (** context elements **)
   542 
   543 (* let bindings *)
   544 
   545 local
   546 
   547 fun gen_bind bind args state =
   548   state
   549   |> assert_forward
   550   |> map_context (bind true args #> snd)
   551   |> put_facts NONE;
   552 
   553 in
   554 
   555 val let_bind = gen_bind Proof_Context.match_bind_i;
   556 val let_bind_cmd = gen_bind Proof_Context.match_bind;
   557 
   558 end;
   559 
   560 
   561 (* concrete syntax *)
   562 
   563 local
   564 
   565 fun gen_write prep_arg mode args =
   566   assert_forward
   567   #> map_context (fn ctxt => ctxt |> Proof_Context.notation true mode (map (prep_arg ctxt) args))
   568   #> put_facts NONE;
   569 
   570 in
   571 
   572 val write = gen_write (K I);
   573 
   574 val write_cmd =
   575   gen_write (fn ctxt => fn (c, mx) =>
   576     (Proof_Context.read_const ctxt false (Mixfix.mixfixT mx) c, mx));
   577 
   578 end;
   579 
   580 
   581 (* fix *)
   582 
   583 local
   584 
   585 fun gen_fix prep_vars args =
   586   assert_forward
   587   #> map_context (fn ctxt => snd (Proof_Context.add_fixes (fst (prep_vars args ctxt)) ctxt))
   588   #> put_facts NONE;
   589 
   590 in
   591 
   592 val fix = gen_fix Proof_Context.cert_vars;
   593 val fix_cmd = gen_fix Proof_Context.read_vars;
   594 
   595 end;
   596 
   597 
   598 (* assume etc. *)
   599 
   600 local
   601 
   602 fun gen_assume asm prep_att exp args state =
   603   state
   604   |> assert_forward
   605   |> map_context_result (asm exp (Attrib.map_specs (map (prep_att (theory_of state))) args))
   606   |> these_factss [] |> #2;
   607 
   608 in
   609 
   610 val assm = gen_assume Proof_Context.add_assms_i (K I);
   611 val assm_cmd = gen_assume Proof_Context.add_assms Attrib.attribute;
   612 val assume = assm Assumption.assume_export;
   613 val assume_cmd = assm_cmd Assumption.assume_export;
   614 val presume = assm Assumption.presume_export;
   615 val presume_cmd = assm_cmd Assumption.presume_export;
   616 
   617 end;
   618 
   619 
   620 (* def *)
   621 
   622 local
   623 
   624 fun gen_def prep_att prep_vars prep_binds args state =
   625   let
   626     val _ = assert_forward state;
   627     val thy = theory_of state;
   628 
   629     val (raw_name_atts, (raw_vars, raw_rhss)) = args |> split_list ||> split_list;
   630     val name_atts = map (apsnd (map (prep_att thy))) raw_name_atts;
   631   in
   632     state
   633     |> map_context_result (prep_vars (map (fn (x, mx) => (x, NONE, mx)) raw_vars))
   634     |>> map (fn (x, _, mx) => (x, mx))
   635     |-> (fn vars =>
   636       map_context_result (prep_binds false (map swap raw_rhss))
   637       #-> (fn rhss => map_context_result (Local_Defs.add_defs (vars ~~ (name_atts ~~ rhss)))))
   638     |-> (put_facts o SOME o map (#2 o #2))
   639   end;
   640 
   641 in
   642 
   643 val def = gen_def (K I) Proof_Context.cert_vars Proof_Context.match_bind_i;
   644 val def_cmd = gen_def Attrib.attribute Proof_Context.read_vars Proof_Context.match_bind;
   645 
   646 end;
   647 
   648 
   649 
   650 (** facts **)
   651 
   652 (* chain *)
   653 
   654 fun clean_facts ctxt =
   655   put_facts (SOME (filter_out Thm.is_dummy (the_facts ctxt))) ctxt;
   656 
   657 val chain =
   658   assert_forward
   659   #> clean_facts
   660   #> enter_chain;
   661 
   662 fun chain_facts facts =
   663   put_facts (SOME facts)
   664   #> chain;
   665 
   666 
   667 (* note etc. *)
   668 
   669 fun no_binding args = map (pair (Binding.empty, [])) args;
   670 
   671 local
   672 
   673 fun gen_thmss more_facts opt_chain opt_result prep_atts prep_fact args state =
   674   state
   675   |> assert_forward
   676   |> map_context_result (Proof_Context.note_thmss ""
   677     (Attrib.map_facts_refs
   678       (map (prep_atts (theory_of state)))
   679       (prep_fact (context_of state)) args))
   680   |> these_factss (more_facts state)
   681   ||> opt_chain
   682   |> opt_result;
   683 
   684 in
   685 
   686 val note_thmss = gen_thmss (K []) I #2 (K I) (K I);
   687 val note_thmss_cmd = gen_thmss (K []) I #2 Attrib.attribute Proof_Context.get_fact;
   688 
   689 val from_thmss = gen_thmss (K []) chain #2 (K I) (K I) o no_binding;
   690 val from_thmss_cmd = gen_thmss (K []) chain #2 Attrib.attribute Proof_Context.get_fact o no_binding;
   691 
   692 val with_thmss = gen_thmss the_facts chain #2 (K I) (K I) o no_binding;
   693 val with_thmss_cmd = gen_thmss the_facts chain #2 Attrib.attribute Proof_Context.get_fact o no_binding;
   694 
   695 val local_results = gen_thmss (K []) I I (K I) (K I) o map (apsnd Thm.simple_fact);
   696 
   697 end;
   698 
   699 
   700 (* using/unfolding *)
   701 
   702 local
   703 
   704 fun gen_using f g prep_att prep_fact args state =
   705   state
   706   |> assert_backward
   707   |> map_context_result
   708     (Proof_Context.note_thmss ""
   709       (Attrib.map_facts_refs (map (prep_att (theory_of state))) (prep_fact (context_of state))
   710         (no_binding args)))
   711   |> (fn (named_facts, state') =>
   712     state' |> map_goal I (fn (statement, _, using, goal, before_qed, after_qed) =>
   713       let
   714         val ctxt = context_of state';
   715         val ths = maps snd named_facts;
   716       in (statement, [], f ctxt ths using, g ctxt ths goal, before_qed, after_qed) end));
   717 
   718 fun append_using _ ths using = using @ filter_out Thm.is_dummy ths;
   719 fun unfold_using ctxt ths = map (Local_Defs.unfold ctxt ths);
   720 val unfold_goals = Local_Defs.unfold_goals;
   721 
   722 in
   723 
   724 val using = gen_using append_using (K (K I)) (K I) (K I);
   725 val using_cmd = gen_using append_using (K (K I)) Attrib.attribute Proof_Context.get_fact;
   726 val unfolding = gen_using unfold_using unfold_goals (K I) (K I);
   727 val unfolding_cmd = gen_using unfold_using unfold_goals Attrib.attribute Proof_Context.get_fact;
   728 
   729 end;
   730 
   731 
   732 (* case *)
   733 
   734 local
   735 
   736 fun qualified_binding a =
   737   Binding.qualify true (Long_Name.qualifier a) (Binding.name (Long_Name.base_name a));
   738 
   739 fun gen_invoke_case prep_att (name, xs, raw_atts) state =
   740   let
   741     val atts = map (prep_att (theory_of state)) raw_atts;
   742     val (asms, state') = state |> map_context_result (fn ctxt =>
   743       ctxt |> Proof_Context.apply_case (Proof_Context.get_case ctxt name xs));
   744     val assumptions = asms |> map (fn (a, ts) => ((qualified_binding a, atts), map (rpair []) ts));
   745   in
   746     state'
   747     |> assume assumptions
   748     |> bind_terms Auto_Bind.no_facts
   749     |> `the_facts |-> (fn thms => note_thmss [((Binding.name name, []), [(thms, [])])])
   750   end;
   751 
   752 in
   753 
   754 val invoke_case = gen_invoke_case (K I);
   755 val invoke_case_cmd = gen_invoke_case Attrib.attribute;
   756 
   757 end;
   758 
   759 
   760 
   761 (** proof structure **)
   762 
   763 (* blocks *)
   764 
   765 val begin_block =
   766   assert_forward
   767   #> open_block
   768   #> put_goal NONE
   769   #> open_block;
   770 
   771 val next_block =
   772   assert_forward
   773   #> close_block
   774   #> open_block
   775   #> put_goal NONE
   776   #> put_facts NONE;
   777 
   778 fun end_block state =
   779   state
   780   |> assert_forward
   781   |> assert_bottom false
   782   |> close_block
   783   |> assert_current_goal false
   784   |> close_block
   785   |> export_facts state;
   786 
   787 
   788 (* global notepad *)
   789 
   790 val begin_notepad =
   791   init
   792   #> open_block
   793   #> map_context (Variable.set_body true)
   794   #> open_block;
   795 
   796 val end_notepad =
   797   assert_forward
   798   #> assert_bottom true
   799   #> close_block
   800   #> assert_current_goal false
   801   #> close_block
   802   #> context_of;
   803 
   804 
   805 (* sub-proofs *)
   806 
   807 fun proof opt_text =
   808   assert_backward
   809   #> refine (the_default Method.default_text opt_text)
   810   #> Seq.map (using_facts [] #> enter_forward);
   811 
   812 fun end_proof bot txt state =
   813   state
   814   |> assert_forward
   815   |> assert_bottom bot
   816   |> close_block
   817   |> assert_current_goal true
   818   |> using_facts []
   819   |> `before_qed |-> (refine o the_default Method.succeed_text)
   820   |> Seq.maps (refine (Method.finish_text txt));
   821 
   822 fun check_result msg sq =
   823   (case Seq.pull sq of
   824     NONE => error msg
   825   | SOME (s, _) => s);
   826 
   827 fun check_finish sq = check_result "Failed to finish proof" sq;
   828 
   829 
   830 (* unstructured refinement *)
   831 
   832 fun defer i = assert_no_chain #> refine (Method.Basic (K (Method.defer i)));
   833 fun prefer i = assert_no_chain #> refine (Method.Basic (K (Method.prefer i)));
   834 
   835 fun apply text = assert_backward #> refine text #> Seq.map (using_facts []);
   836 fun apply_end text = assert_forward #> refine_end text;
   837 
   838 
   839 
   840 (** goals **)
   841 
   842 (* generic goals *)
   843 
   844 local
   845 
   846 val is_var =
   847   can (dest_TVar o Logic.dest_type o Logic.dest_term) orf
   848   can (dest_Var o Logic.dest_term);
   849 
   850 fun implicit_vars props =
   851   let
   852     val (var_props, _) = take_prefix is_var props;
   853     val explicit_vars = fold Term.add_vars var_props [];
   854     val vars = filter_out (member (op =) explicit_vars) (fold Term.add_vars props []);
   855   in map (Logic.mk_term o Var) vars end;
   856 
   857 fun refine_terms n =
   858   refine (Method.Basic (K (RAW_METHOD
   859     (K (HEADGOAL (PRECISE_CONJUNCTS n
   860       (HEADGOAL (CONJUNCTS (ALLGOALS (rtac Drule.termI))))))))))
   861   #> Seq.hd;
   862 
   863 in
   864 
   865 fun generic_goal prepp kind before_qed after_qed raw_propp state =
   866   let
   867     val thy = theory_of state;
   868     val cert = Thm.cterm_of thy;
   869     val chaining = can assert_chain state;
   870     val pos = Position.thread_data ();
   871 
   872     val ((propss, after_ctxt), goal_state) =
   873       state
   874       |> assert_forward_or_chain
   875       |> enter_forward
   876       |> open_block
   877       |> map_context_result (prepp raw_propp);
   878     val props = flat propss;
   879 
   880     val vars = implicit_vars props;
   881     val propss' = vars :: propss;
   882     val goal_propss = filter_out null propss';
   883     val goal =
   884       cert (Logic.mk_conjunction_balanced (map Logic.mk_conjunction_balanced goal_propss))
   885       |> Thm.weaken_sorts (Variable.sorts_of (context_of goal_state));
   886     val statement = ((kind, pos), propss', Thm.term_of goal);
   887     val after_qed' = after_qed |>> (fn after_local =>
   888       fn results => map_context after_ctxt #> after_local results);
   889   in
   890     goal_state
   891     |> map_context (init_context #> Variable.set_body true)
   892     |> put_goal (SOME (make_goal (statement, [], [], Goal.init goal, before_qed, after_qed')))
   893     |> map_context (Proof_Context.auto_bind_goal props)
   894     |> chaining ? (`the_facts #-> using_facts)
   895     |> put_facts NONE
   896     |> open_block
   897     |> put_goal NONE
   898     |> enter_backward
   899     |> not (null vars) ? refine_terms (length goal_propss)
   900     |> null props ? (refine (Method.Basic Method.assumption) #> Seq.hd)
   901   end;
   902 
   903 fun generic_qed after_ctxt state =
   904   let
   905     val (goal_ctxt, {statement = (_, stmt, _), goal, after_qed, ...}) = current_goal state;
   906     val outer_state = state |> close_block;
   907     val outer_ctxt = context_of outer_state;
   908 
   909     val props =
   910       flat (tl stmt)
   911       |> Variable.exportT_terms goal_ctxt outer_ctxt;
   912     val results =
   913       tl (conclude_goal goal_ctxt goal stmt)
   914       |> burrow (Proof_Context.export goal_ctxt outer_ctxt);
   915   in
   916     outer_state
   917     |> map_context (after_ctxt props)
   918     |> pair (after_qed, results)
   919   end;
   920 
   921 end;
   922 
   923 
   924 (* local goals *)
   925 
   926 fun local_goal print_results prep_att prepp kind before_qed after_qed stmt state =
   927   let
   928     val thy = theory_of state;
   929     val ((names, attss), propp) =
   930       Attrib.map_specs (map (prep_att thy)) stmt |> split_list |>> split_list;
   931 
   932     fun after_qed' results =
   933       local_results ((names ~~ attss) ~~ results)
   934       #-> (fn res => tap (fn st => print_results (context_of st) ((kind, ""), res) : unit))
   935       #> after_qed results;
   936   in
   937     state
   938     |> generic_goal prepp kind before_qed (after_qed', K I) propp
   939     |> tap (Variable.warn_extra_tfrees (context_of state) o context_of)
   940   end;
   941 
   942 fun local_qeds txt =
   943   end_proof false txt
   944   #> Seq.map (generic_qed Proof_Context.auto_bind_facts #->
   945     (fn ((after_qed, _), results) => after_qed results));
   946 
   947 fun local_qed txt = local_qeds txt #> check_finish;
   948 
   949 
   950 (* global goals *)
   951 
   952 fun prepp_auto_fixes prepp args =
   953   prepp args #>
   954   (fn ((propss, a), ctxt) => ((propss, a), (fold o fold) Variable.auto_fixes propss ctxt));
   955 
   956 fun global_goal prepp before_qed after_qed propp =
   957   init #>
   958   generic_goal (prepp_auto_fixes prepp) "" before_qed (K I, after_qed) propp;
   959 
   960 val theorem = global_goal Proof_Context.bind_propp_schematic_i;
   961 val theorem_cmd = global_goal Proof_Context.bind_propp_schematic;
   962 
   963 fun global_qeds txt =
   964   end_proof true txt
   965   #> Seq.map (generic_qed (K I) #> (fn (((_, after_qed), results), state) =>
   966     after_qed results (context_of state)));
   967 
   968 fun global_qed txt = global_qeds txt #> check_finish;
   969 
   970 
   971 (* concluding steps *)
   972 
   973 fun terminal_proof qed initial terminal =
   974   proof (SOME initial) #> Seq.maps (qed terminal) #> check_finish;
   975 
   976 fun local_terminal_proof (text, opt_text) = terminal_proof local_qeds text (opt_text, true);
   977 val local_default_proof = local_terminal_proof (Method.default_text, NONE);
   978 val local_immediate_proof = local_terminal_proof (Method.this_text, NONE);
   979 fun local_skip_proof int = local_terminal_proof (Method.sorry_text int, NONE);
   980 val local_done_proof = terminal_proof local_qeds Method.done_text (NONE, false);
   981 
   982 fun global_terminal_proof (text, opt_text) = terminal_proof global_qeds text (opt_text, true);
   983 val global_default_proof = global_terminal_proof (Method.default_text, NONE);
   984 val global_immediate_proof = global_terminal_proof (Method.this_text, NONE);
   985 fun global_skip_proof int = global_terminal_proof (Method.sorry_text int, NONE);
   986 val global_done_proof = terminal_proof global_qeds Method.done_text (NONE, false);
   987 
   988 
   989 (* common goal statements *)
   990 
   991 local
   992 
   993 fun gen_have prep_att prepp before_qed after_qed stmt int =
   994   local_goal (Proof_Display.print_results Isabelle_Markup.state int)
   995     prep_att prepp "have" before_qed after_qed stmt;
   996 
   997 fun gen_show prep_att prepp before_qed after_qed stmt int state =
   998   let
   999     val testing = Unsynchronized.ref false;
  1000     val rule = Unsynchronized.ref (NONE: thm option);
  1001     fun fail_msg ctxt =
  1002       "Local statement will fail to refine any pending goal" ::
  1003       (case ! rule of NONE => [] | SOME th => [Proof_Display.string_of_rule ctxt "Failed" th])
  1004       |> cat_lines;
  1005 
  1006     fun print_results ctxt res =
  1007       if ! testing then ()
  1008       else Proof_Display.print_results Isabelle_Markup.state int ctxt res;
  1009     fun print_rule ctxt th =
  1010       if ! testing then rule := SOME th
  1011       else if int then
  1012         writeln
  1013           (Markup.markup Isabelle_Markup.state (Proof_Display.string_of_rule ctxt "Successful" th))
  1014       else ();
  1015     val test_proof =
  1016       try (local_skip_proof true)
  1017       |> Unsynchronized.setmp testing true
  1018       |> Exn.interruptible_capture;
  1019 
  1020     fun after_qed' results =
  1021       refine_goals print_rule (context_of state) (flat results)
  1022       #> check_result "Failed to refine any pending goal"
  1023       #> after_qed results;
  1024   in
  1025     state
  1026     |> local_goal print_results prep_att prepp "show" before_qed after_qed' stmt
  1027     |> int ? (fn goal_state =>
  1028       (case test_proof goal_state of
  1029         Exn.Res (SOME _) => goal_state
  1030       | Exn.Res NONE => error (fail_msg (context_of goal_state))
  1031       | Exn.Exn exn => raise Exn.EXCEPTIONS ([exn, ERROR (fail_msg (context_of goal_state))])))
  1032   end;
  1033 
  1034 in
  1035 
  1036 val have = gen_have (K I) Proof_Context.bind_propp_i;
  1037 val have_cmd = gen_have Attrib.attribute Proof_Context.bind_propp;
  1038 val show = gen_show (K I) Proof_Context.bind_propp_i;
  1039 val show_cmd = gen_show Attrib.attribute Proof_Context.bind_propp;
  1040 
  1041 end;
  1042 
  1043 
  1044 
  1045 (** future proofs **)
  1046 
  1047 (* relevant proof states *)
  1048 
  1049 fun is_schematic t =
  1050   Term.exists_subterm Term.is_Var t orelse
  1051   Term.exists_type (Term.exists_subtype Term.is_TVar) t;
  1052 
  1053 fun schematic_goal state =
  1054   let val (_, (_, {statement = (_, _, prop), ...})) = find_goal state
  1055   in is_schematic prop end;
  1056 
  1057 fun is_relevant state =
  1058   (case try find_goal state of
  1059     NONE => true
  1060   | SOME (_, (_, {statement = (_, _, prop), goal, ...})) =>
  1061       is_schematic prop orelse not (Logic.protect prop aconv Thm.concl_of goal));
  1062 
  1063 
  1064 (* full proofs *)
  1065 
  1066 local
  1067 
  1068 fun future_proof done get_context fork_proof state =
  1069   let
  1070     val _ = assert_backward state;
  1071     val (goal_ctxt, (_, goal)) = find_goal state;
  1072     val {statement as (kind, _, prop), messages, using, goal, before_qed, after_qed} = goal;
  1073     val goal_txt = prop :: map Thm.term_of (Assumption.all_assms_of goal_ctxt);
  1074 
  1075     val _ = is_relevant state andalso error "Cannot fork relevant proof";
  1076 
  1077     val prop' = Logic.protect prop;
  1078     val statement' = (kind, [[], [prop']], prop');
  1079     val goal' = Thm.adjust_maxidx_thm (Thm.maxidx_of goal)
  1080       (Drule.comp_no_flatten (goal, Thm.nprems_of goal) 1 Drule.protectI);
  1081 
  1082     fun after_local' [[th]] = put_thms false (Auto_Bind.thisN, SOME [th]);
  1083     fun after_global' [[th]] = Proof_Context.put_thms false (Auto_Bind.thisN, SOME [th]);
  1084     val after_qed' = (after_local', after_global');
  1085     val this_name = Proof_Context.full_name goal_ctxt (Binding.name Auto_Bind.thisN);
  1086 
  1087     val result_ctxt =
  1088       state
  1089       |> map_contexts (fold Variable.declare_term goal_txt)
  1090       |> map_goal I (K (statement', messages, using, goal', before_qed, after_qed'))
  1091       |> fork_proof;
  1092 
  1093     val future_thm = result_ctxt |> Future.map (fn (_, x) =>
  1094       Proof_Context.get_fact_single (get_context x) (Facts.named this_name));
  1095     val finished_goal = Goal.future_result goal_ctxt future_thm prop';
  1096     val state' =
  1097       state
  1098       |> map_goal I (K (statement, messages, using, finished_goal, NONE, after_qed))
  1099       |> done;
  1100   in (Future.map #1 result_ctxt, state') end;
  1101 
  1102 in
  1103 
  1104 fun local_future_proof x = future_proof local_done_proof context_of x;
  1105 fun global_future_proof x = future_proof global_done_proof I x;
  1106 
  1107 end;
  1108 
  1109 
  1110 (* terminal proofs *)
  1111 
  1112 local
  1113 
  1114 fun future_terminal_proof proof1 proof2 meths int state =
  1115   if int orelse is_relevant state orelse not (Goal.local_future_enabled ())
  1116   then proof1 meths state
  1117   else snd (proof2 (fn state' =>
  1118     Goal.fork_name "Proof.future_terminal_proof" (fn () => ((), proof1 meths state'))) state);
  1119 
  1120 in
  1121 
  1122 fun local_future_terminal_proof x =
  1123   future_terminal_proof local_terminal_proof local_future_proof x;
  1124 
  1125 fun global_future_terminal_proof x =
  1126   future_terminal_proof global_terminal_proof global_future_proof x;
  1127 
  1128 end;
  1129 
  1130 end;
  1131