src/HOL/ex/Sketch_and_Explore.thy
author Fabian Huch <huch@in.tum.de>
Sat, 16 Mar 2024 09:58:23 +0100
changeset 79907 01652fac1039
parent 79903 d3811cf07da6
child 80100 7506cb70ecfb
permissions -rw-r--r--
unused;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
70018
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
     1
(*   Title:   HOL/ex/Sketch_and_Explore.thy
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
     2
     Author:  Florian Haftmann, TU Muenchen
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
     3
*)
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
     4
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
     5
chapter \<open>Experimental commands \<^text>\<open>sketch\<close> and \<^text>\<open>explore\<close>\<close>
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
     6
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
     7
theory Sketch_and_Explore
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
     8
  imports Main \<comment> \<open>TODO: generalize existing sledgehammer functions to Pure\<close>
79901
54e9875e491f sketch & explore: replace functionality of `sketch` by more useful `nxsketch` (and remove `nxsketch`)
Simon Wimmer <wimmers@in.tum.de>
parents: 79900
diff changeset
     9
  keywords "sketch" "explore" :: diag
70018
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    10
begin
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    11
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    12
ML \<open>
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    13
fun split_clause t =
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    14
  let
74525
c960bfcb91db discontinued Term.dest_abs / Logic.dest_all, which are officially superseded by Variable.dest_abs etc., but there are also Term.dest_abs_global to recover existing tools easily;
wenzelm
parents: 70225
diff changeset
    15
    val (fixes, horn) = funpow_yield (length (Term.strip_all_vars t)) Logic.dest_all_global t;
70018
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    16
    val assms = Logic.strip_imp_prems horn;
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    17
    val concl = Logic.strip_imp_concl horn;
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    18
  in (fixes, assms, concl) end;
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    19
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    20
fun print_typ ctxt T =
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    21
  T
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    22
  |> Syntax.string_of_typ ctxt
79824
ce3a0d2c9aa7 eliminate odd aliases (see also 2746dfc9ceae);
wenzelm
parents: 79799
diff changeset
    23
  |> ATP_Util.maybe_quote ctxt;
70018
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    24
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    25
fun print_term ctxt t =
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    26
  t
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    27
  |> singleton (Syntax.uncheck_terms ctxt)
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    28
  |> Sledgehammer_Isar_Annotate.annotate_types_in_term ctxt
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    29
  |> Print_Mode.setmp [] (Syntax.unparse_term ctxt #> Pretty.string_of)
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    30
  |> Sledgehammer_Util.simplify_spaces
79824
ce3a0d2c9aa7 eliminate odd aliases (see also 2746dfc9ceae);
wenzelm
parents: 79799
diff changeset
    31
  |> ATP_Util.maybe_quote ctxt;
70018
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    32
79902
0d5c41ea208a sketch & explore: reduce unnecessary type constraints
Simon Wimmer <wimmers@in.tum.de>
parents: 79901
diff changeset
    33
fun eigen_context_for_statement (params, assms, concl) ctxt =
70018
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    34
  let
79902
0d5c41ea208a sketch & explore: reduce unnecessary type constraints
Simon Wimmer <wimmers@in.tum.de>
parents: 79901
diff changeset
    35
    val fixes = map (fn (s, T) => (Binding.name s, SOME T, NoSyn)) params
0d5c41ea208a sketch & explore: reduce unnecessary type constraints
Simon Wimmer <wimmers@in.tum.de>
parents: 79901
diff changeset
    36
    val ctxt' = ctxt |> Variable.set_body false |> Proof_Context.add_fixes fixes |> snd
0d5c41ea208a sketch & explore: reduce unnecessary type constraints
Simon Wimmer <wimmers@in.tum.de>
parents: 79901
diff changeset
    37
  in ((params, assms, concl), ctxt') end;
70225
129757af1096 more correct simulation of eigen context for generated Isar statements
haftmann
parents: 70018
diff changeset
    38
129757af1096 more correct simulation of eigen context for generated Isar statements
haftmann
parents: 70018
diff changeset
    39
fun print_isar_skeleton ctxt indent keyword stmt =
129757af1096 more correct simulation of eigen context for generated Isar statements
haftmann
parents: 70018
diff changeset
    40
  let
129757af1096 more correct simulation of eigen context for generated Isar statements
haftmann
parents: 70018
diff changeset
    41
    val ((fixes, assms, concl), ctxt') = eigen_context_for_statement stmt ctxt;
70018
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    42
    val prefix = replicate_string indent " ";
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    43
    val prefix_sep = "\n" ^ prefix ^ "    and ";
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    44
    val show_s = prefix ^ keyword ^ " " ^ print_term ctxt' concl;
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    45
    val if_s = if null assms then NONE
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    46
      else SOME (prefix ^ "  if " ^ space_implode prefix_sep
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    47
        (map (print_term ctxt') assms));
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    48
    val for_s = if null fixes then NONE
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    49
      else SOME (prefix ^ "  for " ^ space_implode prefix_sep
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    50
        (map (fn (v, T) => v ^ " :: " ^ print_typ ctxt T) fixes));
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    51
    val s = cat_lines ([show_s] @ map_filter I [if_s, for_s] @
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    52
      [prefix ^ "  " ^ (if is_none if_s then "" else "using that ") ^ "sorry"]);
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    53
  in
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    54
    s
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    55
  end;
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    56
79901
54e9875e491f sketch & explore: replace functionality of `sketch` by more useful `nxsketch` (and remove `nxsketch`)
Simon Wimmer <wimmers@in.tum.de>
parents: 79900
diff changeset
    57
fun print_skeleton ctxt indent keyword stmt =
78820
b356019e8d49 Added the command nxsketch, which is like sketch except it separates these subgoals by "next" fills in the context using "fix"/"assume" rather than "for"/"if". Also the sketch command does the same if there is only one subgoal.
paulson <lp15@cam.ac.uk>
parents: 74525
diff changeset
    58
  let
b356019e8d49 Added the command nxsketch, which is like sketch except it separates these subgoals by "next" fills in the context using "fix"/"assume" rather than "for"/"if". Also the sketch command does the same if there is only one subgoal.
paulson <lp15@cam.ac.uk>
parents: 74525
diff changeset
    59
    val ((fixes, assms, concl), ctxt') = eigen_context_for_statement stmt ctxt;
b356019e8d49 Added the command nxsketch, which is like sketch except it separates these subgoals by "next" fills in the context using "fix"/"assume" rather than "for"/"if". Also the sketch command does the same if there is only one subgoal.
paulson <lp15@cam.ac.uk>
parents: 74525
diff changeset
    60
    val prefix = replicate_string indent " ";
79900
8f0ff2847ba8 sketch & explore: use Active.sendback_markup_command to preserve indentation of generated proof text
Simon Wimmer <wimmers@in.tum.de>
parents: 79824
diff changeset
    61
    val prefix_sep = "\n" ^ prefix ^ "  and ";
78820
b356019e8d49 Added the command nxsketch, which is like sketch except it separates these subgoals by "next" fills in the context using "fix"/"assume" rather than "for"/"if". Also the sketch command does the same if there is only one subgoal.
paulson <lp15@cam.ac.uk>
parents: 74525
diff changeset
    62
    val show_s = prefix ^ keyword ^ " " ^ print_term ctxt' concl;
b356019e8d49 Added the command nxsketch, which is like sketch except it separates these subgoals by "next" fills in the context using "fix"/"assume" rather than "for"/"if". Also the sketch command does the same if there is only one subgoal.
paulson <lp15@cam.ac.uk>
parents: 74525
diff changeset
    63
    val assumes_s = if null assms then NONE
79900
8f0ff2847ba8 sketch & explore: use Active.sendback_markup_command to preserve indentation of generated proof text
Simon Wimmer <wimmers@in.tum.de>
parents: 79824
diff changeset
    64
      else SOME (prefix ^ "assume " ^ space_implode prefix_sep
78820
b356019e8d49 Added the command nxsketch, which is like sketch except it separates these subgoals by "next" fills in the context using "fix"/"assume" rather than "for"/"if". Also the sketch command does the same if there is only one subgoal.
paulson <lp15@cam.ac.uk>
parents: 74525
diff changeset
    65
        (map (print_term ctxt') assms));
b356019e8d49 Added the command nxsketch, which is like sketch except it separates these subgoals by "next" fills in the context using "fix"/"assume" rather than "for"/"if". Also the sketch command does the same if there is only one subgoal.
paulson <lp15@cam.ac.uk>
parents: 74525
diff changeset
    66
    val fixes_s = if null fixes then NONE
b356019e8d49 Added the command nxsketch, which is like sketch except it separates these subgoals by "next" fills in the context using "fix"/"assume" rather than "for"/"if". Also the sketch command does the same if there is only one subgoal.
paulson <lp15@cam.ac.uk>
parents: 74525
diff changeset
    67
      else SOME (prefix ^ "fix " ^ space_implode prefix_sep
b356019e8d49 Added the command nxsketch, which is like sketch except it separates these subgoals by "next" fills in the context using "fix"/"assume" rather than "for"/"if". Also the sketch command does the same if there is only one subgoal.
paulson <lp15@cam.ac.uk>
parents: 74525
diff changeset
    68
        (map (fn (v, T) => v ^ " :: " ^ print_typ ctxt T) fixes));
b356019e8d49 Added the command nxsketch, which is like sketch except it separates these subgoals by "next" fills in the context using "fix"/"assume" rather than "for"/"if". Also the sketch command does the same if there is only one subgoal.
paulson <lp15@cam.ac.uk>
parents: 74525
diff changeset
    69
    val s = cat_lines (map_filter I [fixes_s, assumes_s] @ [show_s] @ [prefix ^ " sorry"]);
b356019e8d49 Added the command nxsketch, which is like sketch except it separates these subgoals by "next" fills in the context using "fix"/"assume" rather than "for"/"if". Also the sketch command does the same if there is only one subgoal.
paulson <lp15@cam.ac.uk>
parents: 74525
diff changeset
    70
  in
b356019e8d49 Added the command nxsketch, which is like sketch except it separates these subgoals by "next" fills in the context using "fix"/"assume" rather than "for"/"if". Also the sketch command does the same if there is only one subgoal.
paulson <lp15@cam.ac.uk>
parents: 74525
diff changeset
    71
    s
b356019e8d49 Added the command nxsketch, which is like sketch except it separates these subgoals by "next" fills in the context using "fix"/"assume" rather than "for"/"if". Also the sketch command does the same if there is only one subgoal.
paulson <lp15@cam.ac.uk>
parents: 74525
diff changeset
    72
  end;
b356019e8d49 Added the command nxsketch, which is like sketch except it separates these subgoals by "next" fills in the context using "fix"/"assume" rather than "for"/"if". Also the sketch command does the same if there is only one subgoal.
paulson <lp15@cam.ac.uk>
parents: 74525
diff changeset
    73
79901
54e9875e491f sketch & explore: replace functionality of `sketch` by more useful `nxsketch` (and remove `nxsketch`)
Simon Wimmer <wimmers@in.tum.de>
parents: 79900
diff changeset
    74
fun print_sketch ctxt method_text clauses =
54e9875e491f sketch & explore: replace functionality of `sketch` by more useful `nxsketch` (and remove `nxsketch`)
Simon Wimmer <wimmers@in.tum.de>
parents: 79900
diff changeset
    75
  "proof" ^ method_text :: separate "next" (map (print_skeleton ctxt 2 "show") clauses) @ ["qed"];
78820
b356019e8d49 Added the command nxsketch, which is like sketch except it separates these subgoals by "next" fills in the context using "fix"/"assume" rather than "for"/"if". Also the sketch command does the same if there is only one subgoal.
paulson <lp15@cam.ac.uk>
parents: 74525
diff changeset
    76
70018
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    77
fun print_exploration ctxt method_text [clause] =
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    78
    ["proof -", print_isar_skeleton ctxt 2 "have" clause,
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    79
      "  then show ?thesis", "    by" ^ method_text, "qed"]
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    80
  | print_exploration ctxt method_text (clause :: clauses) =
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    81
    "proof -" :: print_isar_skeleton ctxt 2 "have" clause
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    82
      :: map (print_isar_skeleton ctxt 2 "moreover have") clauses
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    83
      @ ["  ultimately show ?thesis", "    by" ^ method_text, "qed"];
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    84
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    85
fun coalesce_method_txt [] = ""
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    86
  | coalesce_method_txt [s] = s
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    87
  | coalesce_method_txt (s1 :: s2 :: ss) =
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    88
      if s1 = "(" orelse s1 = "["
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    89
        orelse s2 = ")" orelse s2 = "]" orelse s2= ":"
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    90
      then s1 ^ coalesce_method_txt (s2 :: ss)
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    91
      else s1 ^ " " ^ coalesce_method_txt (s2 :: ss);
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    92
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    93
fun print_proof_text_from_state print (some_method_ref : ((Method.text * Position.range) * Token.T list) option) state =
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    94
  let
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    95
    val state' = state
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    96
      |> Proof.proof (Option.map fst some_method_ref)
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    97
      |> Seq.the_result ""
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    98
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
    99
    val { context = ctxt, facts = _, goal } = Proof.goal state';
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
   100
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
   101
    val ctxt_print = fold (fn opt => Config.put opt false)
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
   102
      [show_markup, Printer.show_type_emphasis, show_types, show_sorts, show_consts] ctxt
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
   103
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
   104
    val method_text = case some_method_ref of
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
   105
        NONE => ""
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
   106
      | SOME (_, toks) => " " ^ coalesce_method_txt (map Token.unparse toks);
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
   107
        \<comment> \<open>TODO proper printing required\<close>
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
   108
    val goal_props = Logic.strip_imp_prems (Thm.prop_of goal);
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
   109
    val clauses = map split_clause goal_props;
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
   110
    val lines = if null clauses then
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
   111
      if is_none some_method_ref then ["  .."]
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
   112
      else ["  by" ^ method_text]
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
   113
      else print ctxt_print method_text clauses;
79900
8f0ff2847ba8 sketch & explore: use Active.sendback_markup_command to preserve indentation of generated proof text
Simon Wimmer <wimmers@in.tum.de>
parents: 79824
diff changeset
   114
    val message = Active.sendback_markup_command (cat_lines lines);
70018
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
   115
  in
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
   116
    (state |> tap (fn _ => Output.information message))
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
   117
  end
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
   118
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
   119
val sketch = print_proof_text_from_state print_sketch;
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
   120
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
   121
fun explore method_ref = print_proof_text_from_state print_exploration (SOME method_ref);
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
   122
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
   123
fun sketch_cmd some_method_text =
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
   124
  Toplevel.keep_proof (K () o sketch some_method_text o Toplevel.proof_of)
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
   125
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
   126
fun explore_cmd method_text =
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
   127
  Toplevel.keep_proof (K () o explore method_text o Toplevel.proof_of)
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
   128
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
   129
val _ =
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
   130
  Outer_Syntax.command \<^command_keyword>\<open>sketch\<close>
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
   131
    "print sketch of Isar proof text after method application"
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
   132
    (Scan.option (Scan.trace Method.parse) >> sketch_cmd);
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
   133
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
   134
val _ =
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
   135
  Outer_Syntax.command \<^command_keyword>\<open>explore\<close>
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
   136
    "explore proof obligations after method application as Isar proof text"
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
   137
    (Scan.trace Method.parse >> explore_cmd);
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
   138
\<close>
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
   139
571909ef3103 experimental commands for proof sketching and exploration
haftmann
parents:
diff changeset
   140
end