src/HOL/Tools/Sledgehammer/sledgehammer_proof_reconstruct.ML
changeset 39014 e820beaf7d8c
parent 38981 7cf8beb31e0f
parent 39013 c79e6d536267
child 39016 caad9d509bc4
child 39035 094848cf7ef3
equal deleted inserted replaced
38981:7cf8beb31e0f 39014:e820beaf7d8c
     1 (*  Title:      HOL/Tools/Sledgehammer/sledgehammer_proof_reconstruct.ML
       
     2     Author:     Lawrence C. Paulson, Cambridge University Computer Laboratory
       
     3     Author:     Claire Quigley, Cambridge University Computer Laboratory
       
     4     Author:     Jasmin Blanchette, TU Muenchen
       
     5 
       
     6 Transfer of proofs from external provers.
       
     7 *)
       
     8 
       
     9 signature SLEDGEHAMMER_PROOF_RECONSTRUCT =
       
    10 sig
       
    11   type locality = Sledgehammer_Fact_Filter.locality
       
    12   type minimize_command = string list -> string
       
    13   type metis_params =
       
    14     bool * minimize_command * string * (string * locality) list vector * thm
       
    15     * int
       
    16   type isar_params =
       
    17     string Symtab.table * bool * int * Proof.context * int list list
       
    18   type text_result = string * (string * locality) list
       
    19 
       
    20   val metis_proof_text : metis_params -> text_result
       
    21   val isar_proof_text : isar_params -> metis_params -> text_result
       
    22   val proof_text : bool -> isar_params -> metis_params -> text_result
       
    23 end;
       
    24 
       
    25 structure Sledgehammer_Proof_Reconstruct : SLEDGEHAMMER_PROOF_RECONSTRUCT =
       
    26 struct
       
    27 
       
    28 open ATP_Problem
       
    29 open Metis_Clauses
       
    30 open Sledgehammer_Util
       
    31 open Sledgehammer_Fact_Filter
       
    32 open Sledgehammer_Translate
       
    33 
       
    34 type minimize_command = string list -> string
       
    35 type metis_params =
       
    36   bool * minimize_command * string * (string * locality) list vector * thm * int
       
    37 type isar_params =
       
    38   string Symtab.table * bool * int * Proof.context * int list list
       
    39 type text_result = string * (string * locality) list
       
    40 
       
    41 (* Simple simplifications to ensure that sort annotations don't leave a trail of
       
    42    spurious "True"s. *)
       
    43 fun s_not @{const False} = @{const True}
       
    44   | s_not @{const True} = @{const False}
       
    45   | s_not (@{const Not} $ t) = t
       
    46   | s_not t = @{const Not} $ t
       
    47 fun s_conj (@{const True}, t2) = t2
       
    48   | s_conj (t1, @{const True}) = t1
       
    49   | s_conj p = HOLogic.mk_conj p
       
    50 fun s_disj (@{const False}, t2) = t2
       
    51   | s_disj (t1, @{const False}) = t1
       
    52   | s_disj p = HOLogic.mk_disj p
       
    53 fun s_imp (@{const True}, t2) = t2
       
    54   | s_imp (t1, @{const False}) = s_not t1
       
    55   | s_imp p = HOLogic.mk_imp p
       
    56 fun s_iff (@{const True}, t2) = t2
       
    57   | s_iff (t1, @{const True}) = t1
       
    58   | s_iff (t1, t2) = HOLogic.eq_const HOLogic.boolT $ t1 $ t2
       
    59 
       
    60 fun mk_anot (AConn (ANot, [phi])) = phi
       
    61   | mk_anot phi = AConn (ANot, [phi])
       
    62 fun mk_aconn c (phi1, phi2) = AConn (c, [phi1, phi2])
       
    63 
       
    64 fun index_in_shape x = find_index (exists (curry (op =) x))
       
    65 fun is_axiom_number axiom_names num =
       
    66   num > 0 andalso num <= Vector.length axiom_names andalso
       
    67   not (null (Vector.sub (axiom_names, num - 1)))
       
    68 fun is_conjecture_number conjecture_shape num =
       
    69   index_in_shape num conjecture_shape >= 0
       
    70 
       
    71 fun negate_term (Const (@{const_name All}, T) $ Abs (s, T', t')) =
       
    72     Const (@{const_name Ex}, T) $ Abs (s, T', negate_term t')
       
    73   | negate_term (Const (@{const_name Ex}, T) $ Abs (s, T', t')) =
       
    74     Const (@{const_name All}, T) $ Abs (s, T', negate_term t')
       
    75   | negate_term (@{const HOL.implies} $ t1 $ t2) =
       
    76     @{const HOL.conj} $ t1 $ negate_term t2
       
    77   | negate_term (@{const HOL.conj} $ t1 $ t2) =
       
    78     @{const HOL.disj} $ negate_term t1 $ negate_term t2
       
    79   | negate_term (@{const HOL.disj} $ t1 $ t2) =
       
    80     @{const HOL.conj} $ negate_term t1 $ negate_term t2
       
    81   | negate_term (@{const Not} $ t) = t
       
    82   | negate_term t = @{const Not} $ t
       
    83 
       
    84 datatype ('a, 'b, 'c, 'd, 'e) raw_step =
       
    85   Definition of 'a * 'b * 'c |
       
    86   Inference of 'a * 'd * 'e list
       
    87 
       
    88 fun raw_step_number (Definition (num, _, _)) = num
       
    89   | raw_step_number (Inference (num, _, _)) = num
       
    90 
       
    91 (**** PARSING OF TSTP FORMAT ****)
       
    92 
       
    93 (*Strings enclosed in single quotes, e.g. filenames*)
       
    94 val scan_quoted = $$ "'" |-- Scan.repeat (~$$ "'") --| $$ "'" >> implode;
       
    95 
       
    96 val scan_dollar_name =
       
    97   Scan.repeat ($$ "$") -- Symbol.scan_id >> (fn (ss, s) => implode ss ^ s)
       
    98 
       
    99 fun repair_name _ "$true" = "c_True"
       
   100   | repair_name _ "$false" = "c_False"
       
   101   | repair_name _ "$$e" = "c_equal" (* seen in Vampire proofs *)
       
   102   | repair_name _ "equal" = "c_equal" (* needed by SPASS? *)
       
   103   | repair_name pool s =
       
   104     case Symtab.lookup pool s of
       
   105       SOME s' => s'
       
   106     | NONE =>
       
   107       if String.isPrefix "sQ" s andalso String.isSuffix "_eqProxy" s then
       
   108         "c_equal" (* seen in Vampire proofs *)
       
   109       else
       
   110         s
       
   111 (* Generalized first-order terms, which include file names, numbers, etc. *)
       
   112 val parse_potential_integer =
       
   113   (scan_dollar_name || scan_quoted) >> K NONE
       
   114   || scan_integer >> SOME
       
   115 fun parse_annotation x =
       
   116   ((parse_potential_integer ::: Scan.repeat ($$ " " |-- parse_potential_integer)
       
   117     >> map_filter I) -- Scan.optional parse_annotation []
       
   118      >> uncurry (union (op =))
       
   119    || $$ "(" |-- parse_annotations --| $$ ")"
       
   120    || $$ "[" |-- parse_annotations --| $$ "]") x
       
   121 and parse_annotations x =
       
   122   (Scan.optional (parse_annotation
       
   123                   ::: Scan.repeat ($$ "," |-- parse_annotation)) []
       
   124    >> (fn numss => fold (union (op =)) numss [])) x
       
   125 
       
   126 (* Vampire proof lines sometimes contain needless information such as "(0:3)",
       
   127    which can be hard to disambiguate from function application in an LL(1)
       
   128    parser. As a workaround, we extend the TPTP term syntax with such detritus
       
   129    and ignore it. *)
       
   130 fun parse_vampire_detritus x =
       
   131   (scan_integer |-- $$ ":" --| scan_integer >> K []) x
       
   132 
       
   133 fun parse_term pool x =
       
   134   ((scan_dollar_name >> repair_name pool)
       
   135     -- Scan.optional ($$ "(" |-- (parse_vampire_detritus || parse_terms pool)
       
   136                       --| $$ ")") []
       
   137     --| Scan.optional ($$ "(" |-- parse_vampire_detritus --| $$ ")") []
       
   138    >> ATerm) x
       
   139 and parse_terms pool x =
       
   140   (parse_term pool ::: Scan.repeat ($$ "," |-- parse_term pool)) x
       
   141 
       
   142 fun parse_atom pool =
       
   143   parse_term pool -- Scan.option (Scan.option ($$ "!") --| $$ "="
       
   144                                   -- parse_term pool)
       
   145   >> (fn (u1, NONE) => AAtom u1
       
   146        | (u1, SOME (NONE, u2)) => AAtom (ATerm ("c_equal", [u1, u2]))
       
   147        | (u1, SOME (SOME _, u2)) =>
       
   148          mk_anot (AAtom (ATerm ("c_equal", [u1, u2]))))
       
   149 
       
   150 fun fo_term_head (ATerm (s, _)) = s
       
   151 
       
   152 (* TPTP formulas are fully parenthesized, so we don't need to worry about
       
   153    operator precedence. *)
       
   154 fun parse_formula pool x =
       
   155   (($$ "(" |-- parse_formula pool --| $$ ")"
       
   156     || ($$ "!" >> K AForall || $$ "?" >> K AExists)
       
   157        --| $$ "[" -- parse_terms pool --| $$ "]" --| $$ ":"
       
   158        -- parse_formula pool
       
   159        >> (fn ((q, ts), phi) => AQuant (q, map fo_term_head ts, phi))
       
   160     || $$ "~" |-- parse_formula pool >> mk_anot
       
   161     || parse_atom pool)
       
   162    -- Scan.option ((Scan.this_string "=>" >> K AImplies
       
   163                     || Scan.this_string "<=>" >> K AIff
       
   164                     || Scan.this_string "<~>" >> K ANotIff
       
   165                     || Scan.this_string "<=" >> K AIf
       
   166                     || $$ "|" >> K AOr || $$ "&" >> K AAnd)
       
   167                    -- parse_formula pool)
       
   168    >> (fn (phi1, NONE) => phi1
       
   169         | (phi1, SOME (c, phi2)) => mk_aconn c (phi1, phi2))) x
       
   170 
       
   171 val parse_tstp_extra_arguments =
       
   172   Scan.optional ($$ "," |-- parse_annotation
       
   173                  --| Scan.option ($$ "," |-- parse_annotations)) []
       
   174 
       
   175 (* Syntax: (fof|cnf)\(<num>, <formula_role>, <formula> <extra_arguments>\).
       
   176    The <num> could be an identifier, but we assume integers. *)
       
   177  fun parse_tstp_line pool =
       
   178    ((Scan.this_string "fof" || Scan.this_string "cnf") -- $$ "(")
       
   179      |-- scan_integer --| $$ "," -- Symbol.scan_id --| $$ ","
       
   180      -- parse_formula pool -- parse_tstp_extra_arguments --| $$ ")" --| $$ "."
       
   181     >> (fn (((num, role), phi), deps) =>
       
   182            case role of
       
   183              "definition" =>
       
   184              (case phi of
       
   185                 AConn (AIff, [phi1 as AAtom _, phi2]) =>
       
   186                 Definition (num, phi1, phi2)
       
   187               | AAtom (ATerm ("c_equal", _)) =>
       
   188                 Inference (num, phi, deps) (* Vampire's equality proxy axiom *)
       
   189               | _ => raise Fail "malformed definition")
       
   190            | _ => Inference (num, phi, deps))
       
   191 
       
   192 (**** PARSING OF VAMPIRE OUTPUT ****)
       
   193 
       
   194 (* Syntax: <num>. <formula> <annotation> *)
       
   195 fun parse_vampire_line pool =
       
   196   scan_integer --| $$ "." -- parse_formula pool -- parse_annotation
       
   197   >> (fn ((num, phi), deps) => Inference (num, phi, deps))
       
   198 
       
   199 (**** PARSING OF SPASS OUTPUT ****)
       
   200 
       
   201 (* SPASS returns clause references of the form "x.y". We ignore "y", whose role
       
   202    is not clear anyway. *)
       
   203 val parse_dot_name = scan_integer --| $$ "." --| scan_integer
       
   204 
       
   205 val parse_spass_annotations =
       
   206   Scan.optional ($$ ":" |-- Scan.repeat (parse_dot_name
       
   207                                          --| Scan.option ($$ ","))) []
       
   208 
       
   209 (* It is not clear why some literals are followed by sequences of stars and/or
       
   210    pluses. We ignore them. *)
       
   211 fun parse_decorated_atom pool =
       
   212   parse_atom pool --| Scan.repeat ($$ "*" || $$ "+" || $$ " ")
       
   213 
       
   214 fun mk_horn ([], []) = AAtom (ATerm ("c_False", []))
       
   215   | mk_horn ([], pos_lits) = foldr1 (mk_aconn AOr) pos_lits
       
   216   | mk_horn (neg_lits, []) = mk_anot (foldr1 (mk_aconn AAnd) neg_lits)
       
   217   | mk_horn (neg_lits, pos_lits) =
       
   218     mk_aconn AImplies (foldr1 (mk_aconn AAnd) neg_lits,
       
   219                        foldr1 (mk_aconn AOr) pos_lits)
       
   220 
       
   221 fun parse_horn_clause pool =
       
   222   Scan.repeat (parse_decorated_atom pool) --| $$ "|" --| $$ "|"
       
   223     -- Scan.repeat (parse_decorated_atom pool) --| $$ "-" --| $$ ">"
       
   224     -- Scan.repeat (parse_decorated_atom pool)
       
   225   >> (mk_horn o apfst (op @))
       
   226 
       
   227 (* Syntax: <num>[0:<inference><annotations>]
       
   228    <atoms> || <atoms> -> <atoms>. *)
       
   229 fun parse_spass_line pool =
       
   230   scan_integer --| $$ "[" --| $$ "0" --| $$ ":" --| Symbol.scan_id
       
   231     -- parse_spass_annotations --| $$ "]" -- parse_horn_clause pool --| $$ "."
       
   232   >> (fn ((num, deps), u) => Inference (num, u, deps))
       
   233 
       
   234 fun parse_line pool =
       
   235   parse_tstp_line pool || parse_vampire_line pool || parse_spass_line pool
       
   236 fun parse_lines pool = Scan.repeat1 (parse_line pool)
       
   237 fun parse_proof pool =
       
   238   fst o Scan.finite Symbol.stopper
       
   239             (Scan.error (!! (fn _ => raise Fail "unrecognized ATP output")
       
   240                             (parse_lines pool)))
       
   241   o explode o strip_spaces_except_between_ident_chars
       
   242 
       
   243 (**** INTERPRETATION OF TSTP SYNTAX TREES ****)
       
   244 
       
   245 exception FO_TERM of string fo_term list
       
   246 exception FORMULA of (string, string fo_term) formula list
       
   247 exception SAME of unit
       
   248 
       
   249 (* Type variables are given the basic sort "HOL.type". Some will later be
       
   250    constrained by information from type literals, or by type inference. *)
       
   251 fun type_from_fo_term tfrees (u as ATerm (a, us)) =
       
   252   let val Ts = map (type_from_fo_term tfrees) us in
       
   253     case strip_prefix_and_unascii type_const_prefix a of
       
   254       SOME b => Type (invert_const b, Ts)
       
   255     | NONE =>
       
   256       if not (null us) then
       
   257         raise FO_TERM [u]  (* only "tconst"s have type arguments *)
       
   258       else case strip_prefix_and_unascii tfree_prefix a of
       
   259         SOME b =>
       
   260         let val s = "'" ^ b in
       
   261           TFree (s, AList.lookup (op =) tfrees s |> the_default HOLogic.typeS)
       
   262         end
       
   263       | NONE =>
       
   264         case strip_prefix_and_unascii tvar_prefix a of
       
   265           SOME b => TVar (("'" ^ b, 0), HOLogic.typeS)
       
   266         | NONE =>
       
   267           (* Variable from the ATP, say "X1" *)
       
   268           Type_Infer.param 0 (a, HOLogic.typeS)
       
   269   end
       
   270 
       
   271 (* Type class literal applied to a type. Returns triple of polarity, class,
       
   272    type. *)
       
   273 fun type_constraint_from_term pos tfrees (u as ATerm (a, us)) =
       
   274   case (strip_prefix_and_unascii class_prefix a,
       
   275         map (type_from_fo_term tfrees) us) of
       
   276     (SOME b, [T]) => (pos, b, T)
       
   277   | _ => raise FO_TERM [u]
       
   278 
       
   279 (** Accumulate type constraints in a formula: negative type literals **)
       
   280 fun add_var (key, z)  = Vartab.map_default (key, []) (cons z)
       
   281 fun add_type_constraint (false, cl, TFree (a ,_)) = add_var ((a, ~1), cl)
       
   282   | add_type_constraint (false, cl, TVar (ix, _)) = add_var (ix, cl)
       
   283   | add_type_constraint _ = I
       
   284 
       
   285 fun repair_atp_variable_name f s =
       
   286   let
       
   287     fun subscript_name s n = s ^ nat_subscript n
       
   288     val s = String.map f s
       
   289   in
       
   290     case space_explode "_" s of
       
   291       [_] => (case take_suffix Char.isDigit (String.explode s) of
       
   292                 (cs1 as _ :: _, cs2 as _ :: _) =>
       
   293                 subscript_name (String.implode cs1)
       
   294                                (the (Int.fromString (String.implode cs2)))
       
   295               | (_, _) => s)
       
   296     | [s1, s2] => (case Int.fromString s2 of
       
   297                      SOME n => subscript_name s1 n
       
   298                    | NONE => s)
       
   299     | _ => s
       
   300   end
       
   301 
       
   302 (* First-order translation. No types are known for variables. "HOLogic.typeT"
       
   303    should allow them to be inferred. *)
       
   304 fun raw_term_from_pred thy full_types tfrees =
       
   305   let
       
   306     fun aux opt_T extra_us u =
       
   307       case u of
       
   308         ATerm ("hBOOL", [u1]) => aux (SOME @{typ bool}) [] u1
       
   309       | ATerm ("hAPP", [u1, u2]) => aux opt_T (u2 :: extra_us) u1
       
   310       | ATerm (a, us) =>
       
   311         if a = type_wrapper_name then
       
   312           case us of
       
   313             [typ_u, term_u] =>
       
   314             aux (SOME (type_from_fo_term tfrees typ_u)) extra_us term_u
       
   315           | _ => raise FO_TERM us
       
   316         else case strip_prefix_and_unascii const_prefix a of
       
   317           SOME "equal" =>
       
   318           list_comb (Const (@{const_name HOL.eq}, HOLogic.typeT),
       
   319                      map (aux NONE []) us)
       
   320         | SOME b =>
       
   321           let
       
   322             val c = invert_const b
       
   323             val num_type_args = num_type_args thy c
       
   324             val (type_us, term_us) =
       
   325               chop (if full_types then 0 else num_type_args) us
       
   326             (* Extra args from "hAPP" come after any arguments given directly to
       
   327                the constant. *)
       
   328             val term_ts = map (aux NONE []) term_us
       
   329             val extra_ts = map (aux NONE []) extra_us
       
   330             val t =
       
   331               Const (c, if full_types then
       
   332                           case opt_T of
       
   333                             SOME T => map fastype_of term_ts ---> T
       
   334                           | NONE =>
       
   335                             if num_type_args = 0 then
       
   336                               Sign.const_instance thy (c, [])
       
   337                             else
       
   338                               raise Fail ("no type information for " ^ quote c)
       
   339                         else
       
   340                           Sign.const_instance thy (c,
       
   341                               map (type_from_fo_term tfrees) type_us))
       
   342           in list_comb (t, term_ts @ extra_ts) end
       
   343         | NONE => (* a free or schematic variable *)
       
   344           let
       
   345             val ts = map (aux NONE []) (us @ extra_us)
       
   346             val T = map fastype_of ts ---> HOLogic.typeT
       
   347             val t =
       
   348               case strip_prefix_and_unascii fixed_var_prefix a of
       
   349                 SOME b => Free (b, T)
       
   350               | NONE =>
       
   351                 case strip_prefix_and_unascii schematic_var_prefix a of
       
   352                   SOME b => Var ((b, 0), T)
       
   353                 | NONE =>
       
   354                   if is_tptp_variable a then
       
   355                     Var ((repair_atp_variable_name Char.toLower a, 0), T)
       
   356                   else
       
   357                     (* Skolem constants? *)
       
   358                     Var ((repair_atp_variable_name Char.toUpper a, 0), T)
       
   359           in list_comb (t, ts) end
       
   360   in aux (SOME HOLogic.boolT) [] end
       
   361 
       
   362 fun term_from_pred thy full_types tfrees pos (u as ATerm (s, _)) =
       
   363   if String.isPrefix class_prefix s then
       
   364     add_type_constraint (type_constraint_from_term pos tfrees u)
       
   365     #> pair @{const True}
       
   366   else
       
   367     pair (raw_term_from_pred thy full_types tfrees u)
       
   368 
       
   369 val combinator_table =
       
   370   [(@{const_name COMBI}, @{thm COMBI_def_raw}),
       
   371    (@{const_name COMBK}, @{thm COMBK_def_raw}),
       
   372    (@{const_name COMBB}, @{thm COMBB_def_raw}),
       
   373    (@{const_name COMBC}, @{thm COMBC_def_raw}),
       
   374    (@{const_name COMBS}, @{thm COMBS_def_raw})]
       
   375 
       
   376 fun uncombine_term (t1 $ t2) = betapply (pairself uncombine_term (t1, t2))
       
   377   | uncombine_term (Abs (s, T, t')) = Abs (s, T, uncombine_term t')
       
   378   | uncombine_term (t as Const (x as (s, _))) =
       
   379     (case AList.lookup (op =) combinator_table s of
       
   380        SOME thm => thm |> prop_of |> specialize_type @{theory} x |> Logic.dest_equals |> snd
       
   381      | NONE => t)
       
   382   | uncombine_term t = t
       
   383 
       
   384 (* Update schematic type variables with detected sort constraints. It's not
       
   385    totally clear when this code is necessary. *)
       
   386 fun repair_tvar_sorts (t, tvar_tab) =
       
   387   let
       
   388     fun do_type (Type (a, Ts)) = Type (a, map do_type Ts)
       
   389       | do_type (TVar (xi, s)) =
       
   390         TVar (xi, the_default s (Vartab.lookup tvar_tab xi))
       
   391       | do_type (TFree z) = TFree z
       
   392     fun do_term (Const (a, T)) = Const (a, do_type T)
       
   393       | do_term (Free (a, T)) = Free (a, do_type T)
       
   394       | do_term (Var (xi, T)) = Var (xi, do_type T)
       
   395       | do_term (t as Bound _) = t
       
   396       | do_term (Abs (a, T, t)) = Abs (a, do_type T, do_term t)
       
   397       | do_term (t1 $ t2) = do_term t1 $ do_term t2
       
   398   in t |> not (Vartab.is_empty tvar_tab) ? do_term end
       
   399 
       
   400 fun quantify_over_free quant_s free_s body_t =
       
   401   case Term.add_frees body_t [] |> filter (curry (op =) free_s o fst) of
       
   402     [] => body_t
       
   403   | frees as (_, free_T) :: _ =>
       
   404     Abs (free_s, free_T, fold (curry abstract_over) (map Free frees) body_t)
       
   405 
       
   406 (* Interpret an ATP formula as a HOL term, extracting sort constraints as they
       
   407    appear in the formula. *)
       
   408 fun prop_from_formula thy full_types tfrees phi =
       
   409   let
       
   410     fun do_formula pos phi =
       
   411       case phi of
       
   412         AQuant (_, [], phi) => do_formula pos phi
       
   413       | AQuant (q, x :: xs, phi') =>
       
   414         do_formula pos (AQuant (q, xs, phi'))
       
   415         #>> quantify_over_free (case q of
       
   416                                   AForall => @{const_name All}
       
   417                                 | AExists => @{const_name Ex})
       
   418                                (repair_atp_variable_name Char.toLower x)
       
   419       | AConn (ANot, [phi']) => do_formula (not pos) phi' #>> s_not
       
   420       | AConn (c, [phi1, phi2]) =>
       
   421         do_formula (pos |> c = AImplies ? not) phi1
       
   422         ##>> do_formula pos phi2
       
   423         #>> (case c of
       
   424                AAnd => s_conj
       
   425              | AOr => s_disj
       
   426              | AImplies => s_imp
       
   427              | AIf => s_imp o swap
       
   428              | AIff => s_iff
       
   429              | ANotIff => s_not o s_iff)
       
   430       | AAtom tm => term_from_pred thy full_types tfrees pos tm
       
   431       | _ => raise FORMULA [phi]
       
   432   in repair_tvar_sorts (do_formula true phi Vartab.empty) end
       
   433 
       
   434 fun check_formula ctxt =
       
   435   Type_Infer.constrain HOLogic.boolT
       
   436   #> Syntax.check_term (ProofContext.set_mode ProofContext.mode_schematic ctxt)
       
   437 
       
   438 
       
   439 (**** Translation of TSTP files to Isar Proofs ****)
       
   440 
       
   441 fun unvarify_term (Var ((s, 0), T)) = Free (s, T)
       
   442   | unvarify_term t = raise TERM ("unvarify_term: non-Var", [t])
       
   443 
       
   444 fun decode_line full_types tfrees (Definition (num, phi1, phi2)) ctxt =
       
   445     let
       
   446       val thy = ProofContext.theory_of ctxt
       
   447       val t1 = prop_from_formula thy full_types tfrees phi1
       
   448       val vars = snd (strip_comb t1)
       
   449       val frees = map unvarify_term vars
       
   450       val unvarify_args = subst_atomic (vars ~~ frees)
       
   451       val t2 = prop_from_formula thy full_types tfrees phi2
       
   452       val (t1, t2) =
       
   453         HOLogic.eq_const HOLogic.typeT $ t1 $ t2
       
   454         |> unvarify_args |> uncombine_term |> check_formula ctxt
       
   455         |> HOLogic.dest_eq
       
   456     in
       
   457       (Definition (num, t1, t2),
       
   458        fold Variable.declare_term (maps OldTerm.term_frees [t1, t2]) ctxt)
       
   459     end
       
   460   | decode_line full_types tfrees (Inference (num, u, deps)) ctxt =
       
   461     let
       
   462       val thy = ProofContext.theory_of ctxt
       
   463       val t = u |> prop_from_formula thy full_types tfrees
       
   464                 |> uncombine_term |> check_formula ctxt
       
   465     in
       
   466       (Inference (num, t, deps),
       
   467        fold Variable.declare_term (OldTerm.term_frees t) ctxt)
       
   468     end
       
   469 fun decode_lines ctxt full_types tfrees lines =
       
   470   fst (fold_map (decode_line full_types tfrees) lines ctxt)
       
   471 
       
   472 fun is_same_inference _ (Definition _) = false
       
   473   | is_same_inference t (Inference (_, t', _)) = t aconv t'
       
   474 
       
   475 (* No "real" literals means only type information (tfree_tcs, clsrel, or
       
   476    clsarity). *)
       
   477 val is_only_type_information = curry (op aconv) HOLogic.true_const
       
   478 
       
   479 fun replace_one_dep (old, new) dep = if dep = old then new else [dep]
       
   480 fun replace_deps_in_line _ (line as Definition _) = line
       
   481   | replace_deps_in_line p (Inference (num, t, deps)) =
       
   482     Inference (num, t, fold (union (op =) o replace_one_dep p) deps [])
       
   483 
       
   484 (* Discard axioms; consolidate adjacent lines that prove the same formula, since
       
   485    they differ only in type information.*)
       
   486 fun add_line _ _ (line as Definition _) lines = line :: lines
       
   487   | add_line conjecture_shape axiom_names (Inference (num, t, [])) lines =
       
   488     (* No dependencies: axiom, conjecture, or (for Vampire) internal axioms or
       
   489        definitions. *)
       
   490     if is_axiom_number axiom_names num then
       
   491       (* Axioms are not proof lines. *)
       
   492       if is_only_type_information t then
       
   493         map (replace_deps_in_line (num, [])) lines
       
   494       (* Is there a repetition? If so, replace later line by earlier one. *)
       
   495       else case take_prefix (not o is_same_inference t) lines of
       
   496         (_, []) => lines (*no repetition of proof line*)
       
   497       | (pre, Inference (num', _, _) :: post) =>
       
   498         pre @ map (replace_deps_in_line (num', [num])) post
       
   499     else if is_conjecture_number conjecture_shape num then
       
   500       Inference (num, negate_term t, []) :: lines
       
   501     else
       
   502       map (replace_deps_in_line (num, [])) lines
       
   503   | add_line _ _ (Inference (num, t, deps)) lines =
       
   504     (* Type information will be deleted later; skip repetition test. *)
       
   505     if is_only_type_information t then
       
   506       Inference (num, t, deps) :: lines
       
   507     (* Is there a repetition? If so, replace later line by earlier one. *)
       
   508     else case take_prefix (not o is_same_inference t) lines of
       
   509       (* FIXME: Doesn't this code risk conflating proofs involving different
       
   510          types? *)
       
   511        (_, []) => Inference (num, t, deps) :: lines
       
   512      | (pre, Inference (num', t', _) :: post) =>
       
   513        Inference (num, t', deps) ::
       
   514        pre @ map (replace_deps_in_line (num', [num])) post
       
   515 
       
   516 (* Recursively delete empty lines (type information) from the proof. *)
       
   517 fun add_nontrivial_line (Inference (num, t, [])) lines =
       
   518     if is_only_type_information t then delete_dep num lines
       
   519     else Inference (num, t, []) :: lines
       
   520   | add_nontrivial_line line lines = line :: lines
       
   521 and delete_dep num lines =
       
   522   fold_rev add_nontrivial_line (map (replace_deps_in_line (num, [])) lines) []
       
   523 
       
   524 (* ATPs sometimes reuse free variable names in the strangest ways. Removing
       
   525    offending lines often does the trick. *)
       
   526 fun is_bad_free frees (Free x) = not (member (op =) frees x)
       
   527   | is_bad_free _ _ = false
       
   528 
       
   529 (* Vampire is keen on producing these. *)
       
   530 fun is_trivial_formula (@{const Not} $ (Const (@{const_name HOL.eq}, _)
       
   531                                         $ t1 $ t2)) = (t1 aconv t2)
       
   532   | is_trivial_formula _ = false
       
   533 
       
   534 fun add_desired_line _ _ _ _ (line as Definition (num, _, _)) (j, lines) =
       
   535     (j, line :: map (replace_deps_in_line (num, [])) lines)
       
   536   | add_desired_line isar_shrink_factor conjecture_shape axiom_names frees
       
   537                      (Inference (num, t, deps)) (j, lines) =
       
   538     (j + 1,
       
   539      if is_axiom_number axiom_names num orelse
       
   540         is_conjecture_number conjecture_shape num orelse
       
   541         (not (is_only_type_information t) andalso
       
   542          null (Term.add_tvars t []) andalso
       
   543          not (exists_subterm (is_bad_free frees) t) andalso
       
   544          not (is_trivial_formula t) andalso
       
   545          (null lines orelse (* last line must be kept *)
       
   546           (length deps >= 2 andalso j mod isar_shrink_factor = 0))) then
       
   547        Inference (num, t, deps) :: lines  (* keep line *)
       
   548      else
       
   549        map (replace_deps_in_line (num, deps)) lines)  (* drop line *)
       
   550 
       
   551 (** EXTRACTING LEMMAS **)
       
   552 
       
   553 (* Like "split_line", but ignores "\n" that follow a comma (as in SNARK's
       
   554    output). *)
       
   555 val split_proof_lines =
       
   556   let
       
   557     fun aux [] [] = []
       
   558       | aux line [] = [implode (rev line)]
       
   559       | aux line ("," :: "\n" :: rest) = aux ("," :: line) rest
       
   560       | aux line ("\n" :: rest) = aux line [] @ aux [] rest
       
   561       | aux line (s :: rest) = aux (s :: line) rest
       
   562   in aux [] o explode end
       
   563 
       
   564 (* A list consisting of the first number in each line is returned. For TSTP,
       
   565    interesting lines have the form "fof(108, axiom, ...)", where the number
       
   566    (108) is extracted. For SPASS, lines have the form "108[0:Inp] ...", where
       
   567    the first number (108) is extracted. For Vampire, we look for
       
   568    "108. ... [input]". *)
       
   569 fun used_facts_in_atp_proof axiom_names atp_proof =
       
   570   let
       
   571     fun axiom_names_at_index num =
       
   572       let val j = Int.fromString num |> the_default ~1 in
       
   573         if is_axiom_number axiom_names j then Vector.sub (axiom_names, j - 1)
       
   574         else []
       
   575       end
       
   576     val tokens_of =
       
   577       String.tokens (fn c => not (Char.isAlphaNum c) andalso c <> #"_")
       
   578     fun do_line (tag :: num :: "axiom" :: (rest as _ :: _)) =
       
   579         if tag = "cnf" orelse tag = "fof" then
       
   580           (case strip_prefix_and_unascii axiom_prefix (List.last rest) of
       
   581              SOME name =>
       
   582              if member (op =) rest "file" then
       
   583                ([(name, name |> find_first_in_list_vector axiom_names |> the)]
       
   584                 handle Option.Option =>
       
   585                        error ("No such fact: " ^ quote name ^ "."))
       
   586              else
       
   587                axiom_names_at_index num
       
   588            | NONE => axiom_names_at_index num)
       
   589         else
       
   590           []
       
   591       | do_line (num :: "0" :: "Inp" :: _) = axiom_names_at_index num
       
   592       | do_line (num :: rest) =
       
   593         (case List.last rest of "input" => axiom_names_at_index num | _ => [])
       
   594       | do_line _ = []
       
   595   in atp_proof |> split_proof_lines |> maps (do_line o tokens_of) end
       
   596 
       
   597 val indent_size = 2
       
   598 val no_label = ("", ~1)
       
   599 
       
   600 val raw_prefix = "X"
       
   601 val assum_prefix = "A"
       
   602 val fact_prefix = "F"
       
   603 
       
   604 fun string_for_label (s, num) = s ^ string_of_int num
       
   605 
       
   606 fun metis_using [] = ""
       
   607   | metis_using ls =
       
   608     "using " ^ space_implode " " (map string_for_label ls) ^ " "
       
   609 fun metis_apply _ 1 = "by "
       
   610   | metis_apply 1 _ = "apply "
       
   611   | metis_apply i _ = "prefer " ^ string_of_int i ^ " apply "
       
   612 fun metis_name full_types = if full_types then "metisFT" else "metis"
       
   613 fun metis_call full_types [] = metis_name full_types
       
   614   | metis_call full_types ss =
       
   615     "(" ^ metis_name full_types ^ " " ^ space_implode " " ss ^ ")"
       
   616 fun metis_command full_types i n (ls, ss) =
       
   617   metis_using ls ^ metis_apply i n ^ metis_call full_types ss
       
   618 fun metis_line full_types i n ss =
       
   619   "Try this command: " ^
       
   620   Markup.markup Markup.sendback (metis_command full_types i n ([], ss)) ^ "."
       
   621 fun minimize_line _ [] = ""
       
   622   | minimize_line minimize_command ss =
       
   623     case minimize_command ss of
       
   624       "" => ""
       
   625     | command =>
       
   626       "\nTo minimize the number of lemmas, try this: " ^
       
   627       Markup.markup Markup.sendback command ^ "."
       
   628 
       
   629 fun used_facts axiom_names =
       
   630   used_facts_in_atp_proof axiom_names
       
   631   #> List.partition (curry (op =) Chained o snd)
       
   632   #> pairself (sort_distinct (string_ord o pairself fst))
       
   633 
       
   634 fun metis_proof_text (full_types, minimize_command, atp_proof, axiom_names,
       
   635                       goal, i) =
       
   636   let
       
   637     val (chained_lemmas, other_lemmas) = used_facts axiom_names atp_proof
       
   638     val n = Logic.count_prems (prop_of goal)
       
   639   in
       
   640     (metis_line full_types i n (map fst other_lemmas) ^
       
   641      minimize_line minimize_command (map fst (other_lemmas @ chained_lemmas)),
       
   642      other_lemmas @ chained_lemmas)
       
   643   end
       
   644 
       
   645 (** Isar proof construction and manipulation **)
       
   646 
       
   647 fun merge_fact_sets (ls1, ss1) (ls2, ss2) =
       
   648   (union (op =) ls1 ls2, union (op =) ss1 ss2)
       
   649 
       
   650 type label = string * int
       
   651 type facts = label list * string list
       
   652 
       
   653 datatype qualifier = Show | Then | Moreover | Ultimately
       
   654 
       
   655 datatype step =
       
   656   Fix of (string * typ) list |
       
   657   Let of term * term |
       
   658   Assume of label * term |
       
   659   Have of qualifier list * label * term * byline
       
   660 and byline =
       
   661   ByMetis of facts |
       
   662   CaseSplit of step list list * facts
       
   663 
       
   664 fun smart_case_split [] facts = ByMetis facts
       
   665   | smart_case_split proofs facts = CaseSplit (proofs, facts)
       
   666 
       
   667 fun add_fact_from_dep axiom_names num =
       
   668   if is_axiom_number axiom_names num then
       
   669     apsnd (union (op =) (map fst (Vector.sub (axiom_names, num - 1))))
       
   670   else
       
   671     apfst (insert (op =) (raw_prefix, num))
       
   672 
       
   673 fun forall_of v t = HOLogic.all_const (fastype_of v) $ lambda v t
       
   674 fun forall_vars t = fold_rev forall_of (map Var (Term.add_vars t [])) t
       
   675 
       
   676 fun step_for_line _ _ (Definition (_, t1, t2)) = Let (t1, t2)
       
   677   | step_for_line _ _ (Inference (num, t, [])) = Assume ((raw_prefix, num), t)
       
   678   | step_for_line axiom_names j (Inference (num, t, deps)) =
       
   679     Have (if j = 1 then [Show] else [], (raw_prefix, num),
       
   680           forall_vars t,
       
   681           ByMetis (fold (add_fact_from_dep axiom_names) deps ([], [])))
       
   682 
       
   683 fun proof_from_atp_proof pool ctxt full_types tfrees isar_shrink_factor
       
   684                          atp_proof conjecture_shape axiom_names params frees =
       
   685   let
       
   686     val lines =
       
   687       atp_proof ^ "$" (* the $ sign acts as a sentinel (FIXME: needed?) *)
       
   688       |> parse_proof pool
       
   689       |> sort (int_ord o pairself raw_step_number)
       
   690       |> decode_lines ctxt full_types tfrees
       
   691       |> rpair [] |-> fold_rev (add_line conjecture_shape axiom_names)
       
   692       |> rpair [] |-> fold_rev add_nontrivial_line
       
   693       |> rpair (0, []) |-> fold_rev (add_desired_line isar_shrink_factor
       
   694                                              conjecture_shape axiom_names frees)
       
   695       |> snd
       
   696   in
       
   697     (if null params then [] else [Fix params]) @
       
   698     map2 (step_for_line axiom_names) (length lines downto 1) lines
       
   699   end
       
   700 
       
   701 (* When redirecting proofs, we keep information about the labels seen so far in
       
   702    the "backpatches" data structure. The first component indicates which facts
       
   703    should be associated with forthcoming proof steps. The second component is a
       
   704    pair ("assum_ls", "drop_ls"), where "assum_ls" are the labels that should
       
   705    become assumptions and "drop_ls" are the labels that should be dropped in a
       
   706    case split. *)
       
   707 type backpatches = (label * facts) list * (label list * label list)
       
   708 
       
   709 fun used_labels_of_step (Have (_, _, _, by)) =
       
   710     (case by of
       
   711        ByMetis (ls, _) => ls
       
   712      | CaseSplit (proofs, (ls, _)) =>
       
   713        fold (union (op =) o used_labels_of) proofs ls)
       
   714   | used_labels_of_step _ = []
       
   715 and used_labels_of proof = fold (union (op =) o used_labels_of_step) proof []
       
   716 
       
   717 fun new_labels_of_step (Fix _) = []
       
   718   | new_labels_of_step (Let _) = []
       
   719   | new_labels_of_step (Assume (l, _)) = [l]
       
   720   | new_labels_of_step (Have (_, l, _, _)) = [l]
       
   721 val new_labels_of = maps new_labels_of_step
       
   722 
       
   723 val join_proofs =
       
   724   let
       
   725     fun aux _ [] = NONE
       
   726       | aux proof_tail (proofs as (proof1 :: _)) =
       
   727         if exists null proofs then
       
   728           NONE
       
   729         else if forall (curry (op =) (hd proof1) o hd) (tl proofs) then
       
   730           aux (hd proof1 :: proof_tail) (map tl proofs)
       
   731         else case hd proof1 of
       
   732           Have ([], l, t, _) => (* FIXME: should we really ignore the "by"? *)
       
   733           if forall (fn Have ([], l', t', _) :: _ => (l, t) = (l', t')
       
   734                       | _ => false) (tl proofs) andalso
       
   735              not (exists (member (op =) (maps new_labels_of proofs))
       
   736                          (used_labels_of proof_tail)) then
       
   737             SOME (l, t, map rev proofs, proof_tail)
       
   738           else
       
   739             NONE
       
   740         | _ => NONE
       
   741   in aux [] o map rev end
       
   742 
       
   743 fun case_split_qualifiers proofs =
       
   744   case length proofs of
       
   745     0 => []
       
   746   | 1 => [Then]
       
   747   | _ => [Ultimately]
       
   748 
       
   749 fun redirect_proof conjecture_shape hyp_ts concl_t proof =
       
   750   let
       
   751     (* The first pass outputs those steps that are independent of the negated
       
   752        conjecture. The second pass flips the proof by contradiction to obtain a
       
   753        direct proof, introducing case splits when an inference depends on
       
   754        several facts that depend on the negated conjecture. *)
       
   755     fun find_hyp num =
       
   756       nth hyp_ts (index_in_shape num conjecture_shape)
       
   757       handle Subscript =>
       
   758              raise Fail ("Cannot find hypothesis " ^ Int.toString num)
       
   759      val concl_ls = map (pair raw_prefix) (List.last conjecture_shape)
       
   760      val canonicalize_labels =
       
   761        map (fn l => if member (op =) concl_ls l then hd concl_ls else l)
       
   762        #> distinct (op =)
       
   763      fun first_pass ([], contra) = ([], contra)
       
   764        | first_pass ((step as Fix _) :: proof, contra) =
       
   765          first_pass (proof, contra) |>> cons step
       
   766        | first_pass ((step as Let _) :: proof, contra) =
       
   767          first_pass (proof, contra) |>> cons step
       
   768        | first_pass ((step as Assume (l as (_, num), _)) :: proof, contra) =
       
   769          if member (op =) concl_ls l then
       
   770            first_pass (proof, contra ||> l = hd concl_ls ? cons step)
       
   771          else
       
   772            first_pass (proof, contra) |>> cons (Assume (l, find_hyp num))
       
   773        | first_pass (Have (qs, l, t, ByMetis (ls, ss)) :: proof, contra) =
       
   774          let
       
   775            val ls = canonicalize_labels ls
       
   776            val step = Have (qs, l, t, ByMetis (ls, ss))
       
   777          in
       
   778            if exists (member (op =) (fst contra)) ls then
       
   779              first_pass (proof, contra |>> cons l ||> cons step)
       
   780            else
       
   781              first_pass (proof, contra) |>> cons step
       
   782          end
       
   783        | first_pass _ = raise Fail "malformed proof"
       
   784     val (proof_top, (contra_ls, contra_proof)) =
       
   785       first_pass (proof, (concl_ls, []))
       
   786     val backpatch_label = the_default ([], []) oo AList.lookup (op =) o fst
       
   787     fun backpatch_labels patches ls =
       
   788       fold merge_fact_sets (map (backpatch_label patches) ls) ([], [])
       
   789     fun second_pass end_qs ([], assums, patches) =
       
   790         ([Have (end_qs, no_label, concl_t,
       
   791                 ByMetis (backpatch_labels patches (map snd assums)))], patches)
       
   792       | second_pass end_qs (Assume (l, t) :: proof, assums, patches) =
       
   793         second_pass end_qs (proof, (t, l) :: assums, patches)
       
   794       | second_pass end_qs (Have (qs, l, t, ByMetis (ls, ss)) :: proof, assums,
       
   795                             patches) =
       
   796         if member (op =) (snd (snd patches)) l andalso
       
   797            not (member (op =) (fst (snd patches)) l) andalso
       
   798            not (AList.defined (op =) (fst patches) l) then
       
   799           second_pass end_qs (proof, assums, patches ||> apsnd (append ls))
       
   800         else
       
   801           (case List.partition (member (op =) contra_ls) ls of
       
   802              ([contra_l], co_ls) =>
       
   803              if member (op =) qs Show then
       
   804                second_pass end_qs (proof, assums,
       
   805                                    patches |>> cons (contra_l, (co_ls, ss)))
       
   806              else
       
   807                second_pass end_qs
       
   808                            (proof, assums,
       
   809                             patches |>> cons (contra_l, (l :: co_ls, ss)))
       
   810                |>> cons (if member (op =) (fst (snd patches)) l then
       
   811                            Assume (l, negate_term t)
       
   812                          else
       
   813                            Have (qs, l, negate_term t,
       
   814                                  ByMetis (backpatch_label patches l)))
       
   815            | (contra_ls as _ :: _, co_ls) =>
       
   816              let
       
   817                val proofs =
       
   818                  map_filter
       
   819                      (fn l =>
       
   820                          if member (op =) concl_ls l then
       
   821                            NONE
       
   822                          else
       
   823                            let
       
   824                              val drop_ls = filter (curry (op <>) l) contra_ls
       
   825                            in
       
   826                              second_pass []
       
   827                                  (proof, assums,
       
   828                                   patches ||> apfst (insert (op =) l)
       
   829                                           ||> apsnd (union (op =) drop_ls))
       
   830                              |> fst |> SOME
       
   831                            end) contra_ls
       
   832                val (assumes, facts) =
       
   833                  if member (op =) (fst (snd patches)) l then
       
   834                    ([Assume (l, negate_term t)], (l :: co_ls, ss))
       
   835                  else
       
   836                    ([], (co_ls, ss))
       
   837              in
       
   838                (case join_proofs proofs of
       
   839                   SOME (l, t, proofs, proof_tail) =>
       
   840                   Have (case_split_qualifiers proofs @
       
   841                         (if null proof_tail then end_qs else []), l, t,
       
   842                         smart_case_split proofs facts) :: proof_tail
       
   843                 | NONE =>
       
   844                   [Have (case_split_qualifiers proofs @ end_qs, no_label,
       
   845                          concl_t, smart_case_split proofs facts)],
       
   846                 patches)
       
   847                |>> append assumes
       
   848              end
       
   849            | _ => raise Fail "malformed proof")
       
   850        | second_pass _ _ = raise Fail "malformed proof"
       
   851     val proof_bottom =
       
   852       second_pass [Show] (contra_proof, [], ([], ([], []))) |> fst
       
   853   in proof_top @ proof_bottom end
       
   854 
       
   855 (* FIXME: Still needed? Probably not. *)
       
   856 val kill_duplicate_assumptions_in_proof =
       
   857   let
       
   858     fun relabel_facts subst =
       
   859       apfst (map (fn l => AList.lookup (op =) subst l |> the_default l))
       
   860     fun do_step (step as Assume (l, t)) (proof, subst, assums) =
       
   861         (case AList.lookup (op aconv) assums t of
       
   862            SOME l' => (proof, (l, l') :: subst, assums)
       
   863          | NONE => (step :: proof, subst, (t, l) :: assums))
       
   864       | do_step (Have (qs, l, t, by)) (proof, subst, assums) =
       
   865         (Have (qs, l, t,
       
   866                case by of
       
   867                  ByMetis facts => ByMetis (relabel_facts subst facts)
       
   868                | CaseSplit (proofs, facts) =>
       
   869                  CaseSplit (map do_proof proofs, relabel_facts subst facts)) ::
       
   870          proof, subst, assums)
       
   871       | do_step step (proof, subst, assums) = (step :: proof, subst, assums)
       
   872     and do_proof proof = fold do_step proof ([], [], []) |> #1 |> rev
       
   873   in do_proof end
       
   874 
       
   875 val then_chain_proof =
       
   876   let
       
   877     fun aux _ [] = []
       
   878       | aux _ ((step as Assume (l, _)) :: proof) = step :: aux l proof
       
   879       | aux l' (Have (qs, l, t, by) :: proof) =
       
   880         (case by of
       
   881            ByMetis (ls, ss) =>
       
   882            Have (if member (op =) ls l' then
       
   883                    (Then :: qs, l, t,
       
   884                     ByMetis (filter_out (curry (op =) l') ls, ss))
       
   885                  else
       
   886                    (qs, l, t, ByMetis (ls, ss)))
       
   887          | CaseSplit (proofs, facts) =>
       
   888            Have (qs, l, t, CaseSplit (map (aux no_label) proofs, facts))) ::
       
   889         aux l proof
       
   890       | aux _ (step :: proof) = step :: aux no_label proof
       
   891   in aux no_label end
       
   892 
       
   893 fun kill_useless_labels_in_proof proof =
       
   894   let
       
   895     val used_ls = used_labels_of proof
       
   896     fun do_label l = if member (op =) used_ls l then l else no_label
       
   897     fun do_step (Assume (l, t)) = Assume (do_label l, t)
       
   898       | do_step (Have (qs, l, t, by)) =
       
   899         Have (qs, do_label l, t,
       
   900               case by of
       
   901                 CaseSplit (proofs, facts) =>
       
   902                 CaseSplit (map (map do_step) proofs, facts)
       
   903               | _ => by)
       
   904       | do_step step = step
       
   905   in map do_step proof end
       
   906 
       
   907 fun prefix_for_depth n = replicate_string (n + 1)
       
   908 
       
   909 val relabel_proof =
       
   910   let
       
   911     fun aux _ _ _ [] = []
       
   912       | aux subst depth (next_assum, next_fact) (Assume (l, t) :: proof) =
       
   913         if l = no_label then
       
   914           Assume (l, t) :: aux subst depth (next_assum, next_fact) proof
       
   915         else
       
   916           let val l' = (prefix_for_depth depth assum_prefix, next_assum) in
       
   917             Assume (l', t) ::
       
   918             aux ((l, l') :: subst) depth (next_assum + 1, next_fact) proof
       
   919           end
       
   920       | aux subst depth (next_assum, next_fact) (Have (qs, l, t, by) :: proof) =
       
   921         let
       
   922           val (l', subst, next_fact) =
       
   923             if l = no_label then
       
   924               (l, subst, next_fact)
       
   925             else
       
   926               let
       
   927                 val l' = (prefix_for_depth depth fact_prefix, next_fact)
       
   928               in (l', (l, l') :: subst, next_fact + 1) end
       
   929           val relabel_facts =
       
   930             apfst (map (fn l =>
       
   931                            case AList.lookup (op =) subst l of
       
   932                              SOME l' => l'
       
   933                            | NONE => raise Fail ("unknown label " ^
       
   934                                                  quote (string_for_label l))))
       
   935           val by =
       
   936             case by of
       
   937               ByMetis facts => ByMetis (relabel_facts facts)
       
   938             | CaseSplit (proofs, facts) =>
       
   939               CaseSplit (map (aux subst (depth + 1) (1, 1)) proofs,
       
   940                          relabel_facts facts)
       
   941         in
       
   942           Have (qs, l', t, by) ::
       
   943           aux subst depth (next_assum, next_fact) proof
       
   944         end
       
   945       | aux subst depth nextp (step :: proof) =
       
   946         step :: aux subst depth nextp proof
       
   947   in aux [] 0 (1, 1) end
       
   948 
       
   949 fun string_for_proof ctxt full_types i n =
       
   950   let
       
   951     fun fix_print_mode f x =
       
   952       setmp_CRITICAL show_no_free_types true
       
   953           (setmp_CRITICAL show_types true
       
   954                (Print_Mode.setmp (filter (curry (op =) Symbol.xsymbolsN)
       
   955                                          (print_mode_value ())) f)) x
       
   956     fun do_indent ind = replicate_string (ind * indent_size) " "
       
   957     fun do_free (s, T) =
       
   958       maybe_quote s ^ " :: " ^
       
   959       maybe_quote (fix_print_mode (Syntax.string_of_typ ctxt) T)
       
   960     fun do_label l = if l = no_label then "" else string_for_label l ^ ": "
       
   961     fun do_have qs =
       
   962       (if member (op =) qs Moreover then "moreover " else "") ^
       
   963       (if member (op =) qs Ultimately then "ultimately " else "") ^
       
   964       (if member (op =) qs Then then
       
   965          if member (op =) qs Show then "thus" else "hence"
       
   966        else
       
   967          if member (op =) qs Show then "show" else "have")
       
   968     val do_term = maybe_quote o fix_print_mode (Syntax.string_of_term ctxt)
       
   969     fun do_facts (ls, ss) =
       
   970       metis_command full_types 1 1
       
   971                     (ls |> sort_distinct (prod_ord string_ord int_ord),
       
   972                      ss |> sort_distinct string_ord)
       
   973     and do_step ind (Fix xs) =
       
   974         do_indent ind ^ "fix " ^ space_implode " and " (map do_free xs) ^ "\n"
       
   975       | do_step ind (Let (t1, t2)) =
       
   976         do_indent ind ^ "let " ^ do_term t1 ^ " = " ^ do_term t2 ^ "\n"
       
   977       | do_step ind (Assume (l, t)) =
       
   978         do_indent ind ^ "assume " ^ do_label l ^ do_term t ^ "\n"
       
   979       | do_step ind (Have (qs, l, t, ByMetis facts)) =
       
   980         do_indent ind ^ do_have qs ^ " " ^
       
   981         do_label l ^ do_term t ^ " " ^ do_facts facts ^ "\n"
       
   982       | do_step ind (Have (qs, l, t, CaseSplit (proofs, facts))) =
       
   983         space_implode (do_indent ind ^ "moreover\n")
       
   984                       (map (do_block ind) proofs) ^
       
   985         do_indent ind ^ do_have qs ^ " " ^ do_label l ^ do_term t ^ " " ^
       
   986         do_facts facts ^ "\n"
       
   987     and do_steps prefix suffix ind steps =
       
   988       let val s = implode (map (do_step ind) steps) in
       
   989         replicate_string (ind * indent_size - size prefix) " " ^ prefix ^
       
   990         String.extract (s, ind * indent_size,
       
   991                         SOME (size s - ind * indent_size - 1)) ^
       
   992         suffix ^ "\n"
       
   993       end
       
   994     and do_block ind proof = do_steps "{ " " }" (ind + 1) proof
       
   995     (* One-step proofs are pointless; better use the Metis one-liner
       
   996        directly. *)
       
   997     and do_proof [Have (_, _, _, ByMetis _)] = ""
       
   998       | do_proof proof =
       
   999         (if i <> 1 then "prefer " ^ string_of_int i ^ "\n" else "") ^
       
  1000         do_indent 0 ^ "proof -\n" ^
       
  1001         do_steps "" "" 1 proof ^
       
  1002         do_indent 0 ^ (if n <> 1 then "next" else "qed")
       
  1003   in do_proof end
       
  1004 
       
  1005 fun isar_proof_text (pool, debug, isar_shrink_factor, ctxt, conjecture_shape)
       
  1006                     (other_params as (full_types, _, atp_proof, axiom_names,
       
  1007                                       goal, i)) =
       
  1008   let
       
  1009     val (params, hyp_ts, concl_t) = strip_subgoal goal i
       
  1010     val frees = fold Term.add_frees (concl_t :: hyp_ts) []
       
  1011     val tfrees = fold Term.add_tfrees (concl_t :: hyp_ts) []
       
  1012     val n = Logic.count_prems (prop_of goal)
       
  1013     val (one_line_proof, lemma_names) = metis_proof_text other_params
       
  1014     fun isar_proof_for () =
       
  1015       case proof_from_atp_proof pool ctxt full_types tfrees isar_shrink_factor
       
  1016                                 atp_proof conjecture_shape axiom_names params
       
  1017                                 frees
       
  1018            |> redirect_proof conjecture_shape hyp_ts concl_t
       
  1019            |> kill_duplicate_assumptions_in_proof
       
  1020            |> then_chain_proof
       
  1021            |> kill_useless_labels_in_proof
       
  1022            |> relabel_proof
       
  1023            |> string_for_proof ctxt full_types i n of
       
  1024         "" => "\nNo structured proof available."
       
  1025       | proof => "\n\nStructured proof:\n" ^ Markup.markup Markup.sendback proof
       
  1026     val isar_proof =
       
  1027       if debug then
       
  1028         isar_proof_for ()
       
  1029       else
       
  1030         try isar_proof_for ()
       
  1031         |> the_default "\nWarning: The Isar proof construction failed."
       
  1032   in (one_line_proof ^ isar_proof, lemma_names) end
       
  1033 
       
  1034 fun proof_text isar_proof isar_params other_params =
       
  1035   (if isar_proof then isar_proof_text isar_params else metis_proof_text)
       
  1036       other_params
       
  1037 
       
  1038 end;