blanchet@35826: (* Title: HOL/Tools/Sledgehammer/sledgehammer_proof_reconstruct.ML wenzelm@33310: Author: Lawrence C Paulson and Claire Quigley, Cambridge University Computer Laboratory blanchet@36392: Author: Jasmin Blanchette, TU Muenchen paulson@21978: wenzelm@33310: Transfer of proofs from external provers. wenzelm@33310: *) wenzelm@33310: blanchet@35826: signature SLEDGEHAMMER_PROOF_RECONSTRUCT = paulson@24425: sig blanchet@36281: type minimize_command = string list -> string blanchet@36393: type name_pool = Sledgehammer_FOL_Clause.name_pool blanchet@36281: paulson@25492: val chained_hint: string paulson@24425: val invert_const: string -> string paulson@24425: val invert_type_const: string -> string blanchet@36909: val num_type_args: theory -> string -> int paulson@24425: val strip_prefix: string -> string -> string option blanchet@36063: val metis_line: int -> int -> string list -> string blanchet@36223: val metis_proof_text: blanchet@36287: minimize_command * string * string vector * thm * int blanchet@36281: -> string * string list blanchet@36223: val isar_proof_text: blanchet@36909: name_pool option * bool * bool * int * Proof.context * int list list blanchet@36287: -> minimize_command * string * string vector * thm * int blanchet@36287: -> string * string list blanchet@36223: val proof_text: blanchet@36909: bool -> name_pool option * bool * bool * int * Proof.context * int list list blanchet@36287: -> minimize_command * string * string vector * thm * int blanchet@36287: -> string * string list paulson@24425: end; paulson@21978: blanchet@35826: structure Sledgehammer_Proof_Reconstruct : SLEDGEHAMMER_PROOF_RECONSTRUCT = paulson@21978: struct paulson@21978: blanchet@36478: open Sledgehammer_Util blanchet@35865: open Sledgehammer_FOL_Clause blanchet@36909: open Sledgehammer_HOL_Clause blanchet@35865: open Sledgehammer_Fact_Preprocessor paulson@21978: blanchet@36281: type minimize_command = string list -> string blanchet@36281: blanchet@36291: fun is_ident_char c = Char.isAlphaNum c orelse c = #"_" blanchet@36392: fun is_head_digit s = Char.isDigit (String.sub (s, 0)) blanchet@36291: blanchet@36551: (* Hack: Could return false positives (e.g., a user happens to declare a blanchet@36551: constant called "SomeTheory.sko_means_shoe_in_$wedish". *) blanchet@36551: val is_skolem_const_name = blanchet@36551: Long_Name.base_name blanchet@36551: #> String.isPrefix skolem_prefix andf String.isSubstring skolem_infix blanchet@36551: krauss@36607: val index_in_shape : int -> int list list -> int = krauss@36607: find_index o exists o curry (op =) blanchet@36402: fun is_axiom_clause_number thm_names num = num <= Vector.length thm_names blanchet@36570: fun is_conjecture_clause_number conjecture_shape num = blanchet@36570: index_in_shape num conjecture_shape >= 0 blanchet@36291: blanchet@36393: fun ugly_name NONE s = s blanchet@36393: | ugly_name (SOME the_pool) s = blanchet@36393: case Symtab.lookup (snd the_pool) s of blanchet@36393: SOME s' => s' blanchet@36393: | NONE => s blanchet@36393: blanchet@36491: fun smart_lambda v t = blanchet@36551: Abs (case v of blanchet@36551: Const (s, _) => blanchet@36551: List.last (space_explode skolem_infix (Long_Name.base_name s)) blanchet@36551: | Var ((s, _), _) => s blanchet@36551: | Free (s, _) => s blanchet@36551: | _ => "", fastype_of v, abstract_over (v, t)) blanchet@36491: fun forall_of v t = HOLogic.all_const (fastype_of v) $ smart_lambda v t blanchet@36491: blanchet@36491: datatype ('a, 'b, 'c, 'd, 'e) raw_step = blanchet@36491: Definition of 'a * 'b * 'c | blanchet@36491: Inference of 'a * 'd * 'e list blanchet@36491: paulson@21978: (**** PARSING OF TSTP FORMAT ****) paulson@21978: blanchet@36548: fun strip_spaces_in_list [] = "" blanchet@36548: | strip_spaces_in_list [c1] = if Char.isSpace c1 then "" else str c1 blanchet@36548: | strip_spaces_in_list [c1, c2] = blanchet@36548: strip_spaces_in_list [c1] ^ strip_spaces_in_list [c2] blanchet@36548: | strip_spaces_in_list (c1 :: c2 :: c3 :: cs) = blanchet@36548: if Char.isSpace c1 then blanchet@36548: strip_spaces_in_list (c2 :: c3 :: cs) blanchet@36548: else if Char.isSpace c2 then blanchet@36548: if Char.isSpace c3 then blanchet@36548: strip_spaces_in_list (c1 :: c3 :: cs) blanchet@36548: else blanchet@36548: str c1 ^ (if forall is_ident_char [c1, c3] then " " else "") ^ blanchet@36548: strip_spaces_in_list (c3 :: cs) blanchet@36548: else blanchet@36548: str c1 ^ strip_spaces_in_list (c2 :: c3 :: cs) blanchet@36548: val strip_spaces = strip_spaces_in_list o String.explode blanchet@36548: blanchet@36291: (* Syntax trees, either term list or formulae *) blanchet@36486: datatype node = IntLeaf of int | StrNode of string * node list paulson@21978: blanchet@36548: fun str_leaf s = StrNode (s, []) paulson@21978: blanchet@36486: fun scons (x, y) = StrNode ("cons", [x, y]) blanchet@36548: val slist_of = List.foldl scons (str_leaf "nil") paulson@21978: paulson@21978: (*Strings enclosed in single quotes, e.g. filenames*) blanchet@36392: val parse_quoted = $$ "'" |-- Scan.repeat (~$$ "'") --| $$ "'" >> implode; paulson@21978: paulson@21978: (*Integer constants, typically proof line numbers*) blanchet@36392: val parse_integer = Scan.many1 is_head_digit >> (the o Int.fromString o implode) paulson@21978: blanchet@36548: val parse_dollar_name = blanchet@36548: Scan.repeat ($$ "$") -- Symbol.scan_id >> (fn (ss, s) => implode ss ^ s) blanchet@36548: blanchet@36369: (* needed for SPASS's output format *) blanchet@36548: fun repair_name _ "$true" = "c_True" blanchet@36548: | repair_name _ "$false" = "c_False" blanchet@36559: | repair_name _ "$$e" = "c_equal" (* seen in Vampire 11 proofs *) blanchet@36548: | repair_name _ "equal" = "c_equal" (* probably not needed *) blanchet@36393: | repair_name pool s = ugly_name pool s blanchet@36392: (* Generalized first-order terms, which include file names, numbers, etc. *) blanchet@36393: (* The "x" argument is not strictly necessary, but without it Poly/ML loops blanchet@36393: forever at compile time. *) blanchet@36393: fun parse_term pool x = blanchet@36548: (parse_quoted >> str_leaf blanchet@36486: || parse_integer >> IntLeaf blanchet@36548: || (parse_dollar_name >> repair_name pool) blanchet@36486: -- Scan.optional ($$ "(" |-- parse_terms pool --| $$ ")") [] >> StrNode blanchet@36393: || $$ "(" |-- parse_term pool --| $$ ")" blanchet@36393: || $$ "[" |-- Scan.optional (parse_terms pool) [] --| $$ "]" >> slist_of) x blanchet@36393: and parse_terms pool x = blanchet@36393: (parse_term pool ::: Scan.repeat ($$ "," |-- parse_term pool)) x paulson@21978: blanchet@36486: fun negate_node u = StrNode ("c_Not", [u]) blanchet@36486: fun equate_nodes u1 u2 = StrNode ("c_equal", [u1, u2]) paulson@21978: blanchet@36392: (* Apply equal or not-equal to a term. *) blanchet@36486: fun repair_predicate_term (u, NONE) = u blanchet@36486: | repair_predicate_term (u1, SOME (NONE, u2)) = equate_nodes u1 u2 blanchet@36486: | repair_predicate_term (u1, SOME (SOME _, u2)) = blanchet@36486: negate_node (equate_nodes u1 u2) blanchet@36393: fun parse_predicate_term pool = blanchet@36393: parse_term pool -- Scan.option (Scan.option ($$ "!") --| $$ "=" blanchet@36393: -- parse_term pool) blanchet@36393: >> repair_predicate_term blanchet@36393: fun parse_literal pool x = blanchet@36486: ($$ "~" |-- parse_literal pool >> negate_node || parse_predicate_term pool) x blanchet@36393: fun parse_literals pool = blanchet@36393: parse_literal pool ::: Scan.repeat ($$ "|" |-- parse_literal pool) blanchet@36548: fun parse_parenthesized_literals pool = blanchet@36548: $$ "(" |-- parse_literals pool --| $$ ")" || parse_literals pool blanchet@36393: fun parse_clause pool = blanchet@36548: parse_parenthesized_literals pool blanchet@36548: ::: Scan.repeat ($$ "|" |-- parse_parenthesized_literals pool) blanchet@36548: >> List.concat blanchet@36291: blanchet@36486: fun ints_of_node (IntLeaf n) = cons n blanchet@36486: | ints_of_node (StrNode (_, us)) = fold ints_of_node us blanchet@36392: val parse_tstp_annotations = blanchet@36393: Scan.optional ($$ "," |-- parse_term NONE blanchet@36393: --| Scan.option ($$ "," |-- parse_terms NONE) blanchet@36486: >> (fn source => ints_of_node source [])) [] blanchet@36486: blanchet@36486: fun parse_definition pool = blanchet@36486: $$ "(" |-- parse_literal NONE --| Scan.this_string "<=>" blanchet@36486: -- parse_clause pool --| $$ ")" blanchet@36291: blanchet@36486: (* Syntax: cnf(, , ). blanchet@36486: The could be an identifier, but we assume integers. *) blanchet@36486: fun finish_tstp_definition_line (num, (u, us)) = Definition (num, u, us) blanchet@36486: fun finish_tstp_inference_line ((num, us), deps) = Inference (num, us, deps) blanchet@36393: fun parse_tstp_line pool = blanchet@36486: ((Scan.this_string "fof" -- $$ "(") |-- parse_integer --| $$ "," blanchet@36486: --| Scan.this_string "definition" --| $$ "," -- parse_definition pool blanchet@36486: --| parse_tstp_annotations --| $$ ")" --| $$ "." blanchet@36486: >> finish_tstp_definition_line) blanchet@36486: || ((Scan.this_string "cnf" -- $$ "(") |-- parse_integer --| $$ "," blanchet@36486: --| Symbol.scan_id --| $$ "," -- parse_clause pool blanchet@36486: -- parse_tstp_annotations --| $$ ")" --| $$ "." blanchet@36486: >> finish_tstp_inference_line) blanchet@36291: blanchet@36291: (**** PARSING OF SPASS OUTPUT ****) blanchet@36291: blanchet@36392: (* SPASS returns clause references of the form "x.y". We ignore "y", whose role blanchet@36392: is not clear anyway. *) blanchet@36392: val parse_dot_name = parse_integer --| $$ "." --| parse_integer paulson@21978: blanchet@36392: val parse_spass_annotations = blanchet@36392: Scan.optional ($$ ":" |-- Scan.repeat (parse_dot_name blanchet@36392: --| Scan.option ($$ ","))) [] blanchet@36291: blanchet@36574: (* It is not clear why some literals are followed by sequences of stars and/or blanchet@36574: pluses. We ignore them. *) blanchet@36574: fun parse_decorated_predicate_term pool = blanchet@36562: parse_predicate_term pool --| Scan.repeat ($$ "*" || $$ "+" || $$ " ") blanchet@36291: blanchet@36393: fun parse_horn_clause pool = blanchet@36574: Scan.repeat (parse_decorated_predicate_term pool) --| $$ "|" --| $$ "|" blanchet@36574: -- Scan.repeat (parse_decorated_predicate_term pool) --| $$ "-" --| $$ ">" blanchet@36574: -- Scan.repeat (parse_decorated_predicate_term pool) blanchet@36558: >> (fn (([], []), []) => [str_leaf "c_False"] blanchet@36558: | ((clauses1, clauses2), clauses3) => blanchet@36558: map negate_node (clauses1 @ clauses2) @ clauses3) paulson@21978: blanchet@36558: (* Syntax: [0:] blanchet@36558: || -> . *) blanchet@36486: fun finish_spass_line ((num, deps), us) = Inference (num, us, deps) blanchet@36402: fun parse_spass_line pool = blanchet@36392: parse_integer --| $$ "[" --| $$ "0" --| $$ ":" --| Symbol.scan_id blanchet@36558: -- parse_spass_annotations --| $$ "]" -- parse_horn_clause pool --| $$ "." blanchet@36486: >> finish_spass_line blanchet@36291: blanchet@36548: fun parse_line pool = parse_tstp_line pool || parse_spass_line pool blanchet@36548: fun parse_lines pool = Scan.repeat1 (parse_line pool) blanchet@36548: fun parse_proof pool = blanchet@36548: fst o Scan.finite Symbol.stopper blanchet@36548: (Scan.error (!! (fn _ => raise Fail "unrecognized ATP output") blanchet@36548: (parse_lines pool))) blanchet@36548: o explode o strip_spaces paulson@21978: paulson@21978: (**** INTERPRETATION OF TSTP SYNTAX TREES ****) paulson@21978: blanchet@36909: exception NODE of node list paulson@21978: paulson@21978: (*If string s has the prefix s1, return the result of deleting it.*) wenzelm@23139: fun strip_prefix s1 s = immler@31038: if String.isPrefix s1 s blanchet@35865: then SOME (undo_ascii_of (String.extract (s, size s1, NONE))) paulson@21978: else NONE; paulson@21978: paulson@21978: (*Invert the table of translations between Isabelle and ATPs*) paulson@21978: val type_const_trans_table_inv = blanchet@35865: Symtab.make (map swap (Symtab.dest type_const_trans_table)); paulson@21978: paulson@21978: fun invert_type_const c = paulson@21978: case Symtab.lookup type_const_trans_table_inv c of paulson@21978: SOME c' => c' paulson@21978: | NONE => c; paulson@21978: blanchet@36909: (* Type variables are given the basic sort "HOL.type". Some will later be blanchet@36909: constrained by information from type literals, or by type inference. *) blanchet@36967: fun type_from_node _ (u as IntLeaf _) = raise NODE [u] blanchet@36967: | type_from_node tfrees (u as StrNode (a, us)) = blanchet@36967: let val Ts = map (type_from_node tfrees) us in blanchet@36486: case strip_prefix tconst_prefix a of blanchet@36486: SOME b => Type (invert_type_const b, Ts) blanchet@36486: | NONE => blanchet@36486: if not (null us) then blanchet@36909: raise NODE [u] (* only "tconst"s have type arguments *) blanchet@36486: else case strip_prefix tfree_prefix a of blanchet@36967: SOME b => blanchet@36967: let val s = "'" ^ b in blanchet@36967: TFree (s, AList.lookup (op =) tfrees s |> the_default HOLogic.typeS) blanchet@36967: end blanchet@36486: | NONE => blanchet@36486: case strip_prefix tvar_prefix a of blanchet@36967: SOME b => TVar (("'" ^ b, 0), HOLogic.typeS) blanchet@36967: | NONE => blanchet@36967: (* Variable from the ATP, say "X1" *) wenzelm@37145: Type_Infer.param 0 (a, HOLogic.typeS) blanchet@36486: end paulson@21978: paulson@21978: (*Invert the table of translations between Isabelle and ATPs*) paulson@21978: val const_trans_table_inv = blanchet@36402: Symtab.update ("fequal", @{const_name "op ="}) blanchet@36402: (Symtab.make (map swap (Symtab.dest const_trans_table))) paulson@21978: blanchet@36402: fun invert_const c = c |> Symtab.lookup const_trans_table_inv |> the_default c paulson@21978: paulson@21978: (*The number of type arguments of a constant, zero if it's monomorphic*) blanchet@36909: fun num_type_args thy s = blanchet@36909: length (Sign.const_typargs thy (s, Sign.the_const_type thy s)) paulson@21978: blanchet@36486: fun fix_atp_variable_name s = blanchet@36486: let blanchet@36486: fun subscript_name s n = s ^ nat_subscript n blanchet@36486: val s = String.map Char.toLower s blanchet@36486: in blanchet@36486: case space_explode "_" s of blanchet@36486: [_] => (case take_suffix Char.isDigit (String.explode s) of blanchet@36486: (cs1 as _ :: _, cs2 as _ :: _) => blanchet@36486: subscript_name (String.implode cs1) blanchet@36486: (the (Int.fromString (String.implode cs2))) blanchet@36486: | (_, _) => s) blanchet@36486: | [s1, s2] => (case Int.fromString s2 of blanchet@36486: SOME n => subscript_name s1 n blanchet@36486: | NONE => s) blanchet@36486: | _ => s blanchet@36486: end blanchet@36486: blanchet@36909: (* First-order translation. No types are known for variables. "HOLogic.typeT" blanchet@36909: should allow them to be inferred.*) blanchet@36967: fun term_from_node thy full_types tfrees = blanchet@36909: let blanchet@36909: fun aux opt_T args u = blanchet@36909: case u of blanchet@36909: IntLeaf _ => raise NODE [u] blanchet@36909: | StrNode ("hBOOL", [u1]) => aux (SOME @{typ bool}) [] u1 blanchet@36909: | StrNode ("hAPP", [u1, u2]) => aux opt_T (u2 :: args) u1 blanchet@36909: | StrNode ("c_Not", [u1]) => @{const Not} $ aux (SOME @{typ bool}) [] u1 blanchet@36909: | StrNode (a, us) => blanchet@36909: if a = type_wrapper_name then blanchet@36909: case us of blanchet@36967: [term_u, typ_u] => blanchet@36967: aux (SOME (type_from_node tfrees typ_u)) args term_u blanchet@36909: | _ => raise NODE us blanchet@36909: else case strip_prefix const_prefix a of blanchet@36909: SOME "equal" => blanchet@36909: list_comb (Const (@{const_name "op ="}, HOLogic.typeT), blanchet@36909: map (aux NONE []) us) blanchet@36909: | SOME b => blanchet@36909: let blanchet@36909: val c = invert_const b blanchet@36909: val num_type_args = num_type_args thy c blanchet@36909: val actual_num_type_args = if full_types then 0 else num_type_args blanchet@36909: val num_term_args = length us - actual_num_type_args blanchet@36909: val ts = map (aux NONE []) (take num_term_args us @ args) blanchet@36909: val t = blanchet@36909: Const (c, if full_types then blanchet@36909: case opt_T of blanchet@36909: SOME T => map fastype_of ts ---> T blanchet@36909: | NONE => blanchet@36909: if num_type_args = 0 then blanchet@36909: Sign.const_instance thy (c, []) blanchet@36909: else blanchet@36909: raise Fail ("no type information for " ^ quote c) blanchet@36909: else blanchet@36909: (* Extra args from "hAPP" come after any arguments blanchet@36909: given directly to the constant. *) blanchet@36909: Sign.const_instance thy (c, blanchet@36967: map (type_from_node tfrees) blanchet@36967: (drop num_term_args us))) blanchet@36909: in list_comb (t, ts) end blanchet@36909: | NONE => (* a free or schematic variable *) blanchet@36909: let blanchet@36909: val ts = map (aux NONE []) (us @ args) blanchet@36909: val T = map fastype_of ts ---> HOLogic.typeT blanchet@36909: val t = blanchet@36909: case strip_prefix fixed_var_prefix a of blanchet@36909: SOME b => Free (b, T) blanchet@36909: | NONE => blanchet@36909: case strip_prefix schematic_var_prefix a of blanchet@36967: SOME b => Var ((b, 0), T) blanchet@36909: | NONE => blanchet@36909: (* Variable from the ATP, say "X1" *) blanchet@36967: Var ((fix_atp_variable_name a, 0), T) blanchet@36909: in list_comb (t, ts) end blanchet@36909: in aux end paulson@21978: blanchet@36392: (* Type class literal applied to a type. Returns triple of polarity, class, blanchet@36392: type. *) blanchet@36967: fun type_constraint_from_node pos tfrees (StrNode ("c_Not", [u])) = blanchet@36967: type_constraint_from_node (not pos) tfrees u blanchet@36967: | type_constraint_from_node pos tfrees u = case u of blanchet@36909: IntLeaf _ => raise NODE [u] blanchet@36486: | StrNode (a, us) => blanchet@36967: (case (strip_prefix class_prefix a, blanchet@36967: map (type_from_node tfrees) us) of blanchet@36486: (SOME b, [T]) => (pos, b, T) blanchet@36909: | _ => raise NODE [u]) paulson@21978: paulson@21978: (** Accumulate type constraints in a clause: negative type literals **) paulson@21978: blanchet@36485: fun add_var (key, z) = Vartab.map_default (key, []) (cons z) paulson@21978: blanchet@36909: fun add_type_constraint (false, cl, TFree (a ,_)) = add_var ((a, ~1), cl) blanchet@36909: | add_type_constraint (false, cl, TVar (ix, _)) = add_var (ix, cl) blanchet@36909: | add_type_constraint _ = I paulson@21978: blanchet@36491: fun is_positive_literal (@{const Not} $ _) = false blanchet@36402: | is_positive_literal t = true blanchet@36402: blanchet@36485: fun negate_term thy (Const (@{const_name All}, T) $ Abs (s, T', t')) = blanchet@36402: Const (@{const_name Ex}, T) $ Abs (s, T', negate_term thy t') blanchet@36402: | negate_term thy (Const (@{const_name Ex}, T) $ Abs (s, T', t')) = blanchet@36402: Const (@{const_name All}, T) $ Abs (s, T', negate_term thy t') blanchet@36402: | negate_term thy (@{const "op -->"} $ t1 $ t2) = blanchet@36402: @{const "op &"} $ t1 $ negate_term thy t2 blanchet@36402: | negate_term thy (@{const "op &"} $ t1 $ t2) = blanchet@36402: @{const "op |"} $ negate_term thy t1 $ negate_term thy t2 blanchet@36402: | negate_term thy (@{const "op |"} $ t1 $ t2) = blanchet@36402: @{const "op &"} $ negate_term thy t1 $ negate_term thy t2 blanchet@36486: | negate_term _ (@{const Not} $ t) = t blanchet@36486: | negate_term _ t = @{const Not} $ t blanchet@36402: blanchet@36402: fun clause_for_literals _ [] = HOLogic.false_const blanchet@36402: | clause_for_literals _ [lit] = lit blanchet@36402: | clause_for_literals thy lits = blanchet@36402: case List.partition is_positive_literal lits of blanchet@36402: (pos_lits as _ :: _, neg_lits as _ :: _) => blanchet@36402: @{const "op -->"} blanchet@36402: $ foldr1 HOLogic.mk_conj (map (negate_term thy) neg_lits) blanchet@36402: $ foldr1 HOLogic.mk_disj pos_lits blanchet@36402: | _ => foldr1 HOLogic.mk_disj lits blanchet@36402: blanchet@36402: (* Final treatment of the list of "real" literals from a clause. blanchet@36402: No "real" literals means only type information. *) blanchet@36402: fun finish_clause _ [] = HOLogic.true_const blanchet@36402: | finish_clause thy lits = blanchet@36402: lits |> filter_out (curry (op =) HOLogic.false_const) |> rev blanchet@36402: |> clause_for_literals thy paulson@22491: paulson@21978: (*Accumulate sort constraints in vt, with "real" literals in lits.*) blanchet@36967: fun lits_of_nodes thy full_types tfrees = blanchet@36967: let blanchet@36967: fun aux (vt, lits) [] = (vt, finish_clause thy lits) blanchet@36967: | aux (vt, lits) (u :: us) = blanchet@36967: aux (add_type_constraint blanchet@36967: (type_constraint_from_node true tfrees u) vt, lits) us blanchet@36967: handle NODE _ => blanchet@36967: aux (vt, term_from_node thy full_types tfrees (SOME @{typ bool}) blanchet@36967: [] u :: lits) us blanchet@36967: in aux end paulson@21978: blanchet@36967: (* Update TVars with detected sort constraints. It's not totally clear when blanchet@36967: this code is necessary. *) blanchet@36967: fun repair_tvar_sorts vt = blanchet@36556: let blanchet@36556: fun do_type (Type (a, Ts)) = Type (a, map do_type Ts) blanchet@36556: | do_type (TVar (xi, s)) = TVar (xi, the_default s (Vartab.lookup vt xi)) blanchet@36967: | do_type (TFree z) = TFree z blanchet@36556: fun do_term (Const (a, T)) = Const (a, do_type T) blanchet@36556: | do_term (Free (a, T)) = Free (a, do_type T) blanchet@36556: | do_term (Var (xi, T)) = Var (xi, do_type T) blanchet@36556: | do_term (t as Bound _) = t blanchet@36556: | do_term (Abs (a, T, t)) = Abs (a, do_type T, do_term t) blanchet@36556: | do_term (t1 $ t2) = do_term t1 $ do_term t2 blanchet@36556: in not (Vartab.is_empty vt) ? do_term end blanchet@36551: blanchet@36551: fun unskolemize_term t = blanchet@36909: Term.add_consts t [] blanchet@36909: |> filter (is_skolem_const_name o fst) |> map Const blanchet@36909: |> rpair t |-> fold forall_of paulson@21978: blanchet@36555: val combinator_table = blanchet@36555: [(@{const_name COMBI}, @{thm COMBI_def_raw}), blanchet@36555: (@{const_name COMBK}, @{thm COMBK_def_raw}), blanchet@36555: (@{const_name COMBB}, @{thm COMBB_def_raw}), blanchet@36555: (@{const_name COMBC}, @{thm COMBC_def_raw}), blanchet@36555: (@{const_name COMBS}, @{thm COMBS_def_raw})] blanchet@36555: blanchet@36555: fun uncombine_term (t1 $ t2) = betapply (pairself uncombine_term (t1, t2)) blanchet@36555: | uncombine_term (Abs (s, T, t')) = Abs (s, T, uncombine_term t') blanchet@36555: | uncombine_term (t as Const (x as (s, _))) = blanchet@36555: (case AList.lookup (op =) combinator_table s of blanchet@36555: SOME thm => thm |> prop_of |> specialize_type @{theory} x |> Logic.dest_equals |> snd blanchet@36555: | NONE => t) blanchet@36555: | uncombine_term t = t blanchet@36555: blanchet@36486: (* Interpret a list of syntax trees as a clause, given by "real" literals and blanchet@36486: sort constraints. "vt" holds the initial sort constraints, from the blanchet@36486: conjecture clauses. *) blanchet@36967: fun clause_of_nodes ctxt full_types tfrees us = blanchet@36909: let blanchet@36909: val thy = ProofContext.theory_of ctxt blanchet@36967: val (vt, t) = lits_of_nodes thy full_types tfrees (Vartab.empty, []) us blanchet@36967: in repair_tvar_sorts vt t end blanchet@36556: fun check_formula ctxt = wenzelm@37145: Type_Infer.constrain @{typ bool} blanchet@36486: #> Syntax.check_term (ProofContext.set_mode ProofContext.mode_schematic ctxt) paulson@21978: paulson@21978: paulson@21978: (**** Translation of TSTP files to Isar Proofs ****) paulson@21978: blanchet@36486: fun unvarify_term (Var ((s, 0), T)) = Free (s, T) blanchet@36486: | unvarify_term t = raise TERM ("unvarify_term: non-Var", [t]) paulson@21978: blanchet@36967: fun decode_line full_types tfrees (Definition (num, u, us)) ctxt = blanchet@36486: let blanchet@36967: val t1 = clause_of_nodes ctxt full_types tfrees [u] blanchet@36551: val vars = snd (strip_comb t1) blanchet@36486: val frees = map unvarify_term vars blanchet@36486: val unvarify_args = subst_atomic (vars ~~ frees) blanchet@36967: val t2 = clause_of_nodes ctxt full_types tfrees us blanchet@36551: val (t1, t2) = blanchet@36551: HOLogic.eq_const HOLogic.typeT $ t1 $ t2 blanchet@36556: |> unvarify_args |> uncombine_term |> check_formula ctxt blanchet@36555: |> HOLogic.dest_eq blanchet@36486: in blanchet@36551: (Definition (num, t1, t2), blanchet@36551: fold Variable.declare_term (maps OldTerm.term_frees [t1, t2]) ctxt) blanchet@36486: end blanchet@36967: | decode_line full_types tfrees (Inference (num, us, deps)) ctxt = blanchet@36551: let blanchet@36967: val t = us |> clause_of_nodes ctxt full_types tfrees blanchet@36556: |> unskolemize_term |> uncombine_term |> check_formula ctxt blanchet@36551: in blanchet@36551: (Inference (num, t, deps), blanchet@36551: fold Variable.declare_term (OldTerm.term_frees t) ctxt) blanchet@36486: end blanchet@36967: fun decode_lines ctxt full_types tfrees lines = blanchet@36967: fst (fold_map (decode_line full_types tfrees) lines ctxt) paulson@21978: blanchet@36486: fun aint_inference _ (Definition _) = true blanchet@36486: | aint_inference t (Inference (_, t', _)) = not (t aconv t') blanchet@36486: blanchet@36486: (* No "real" literals means only type information (tfree_tcs, clsrel, or blanchet@36486: clsarity). *) blanchet@36486: val is_only_type_information = curry (op aconv) HOLogic.true_const blanchet@36486: blanchet@36486: fun replace_one_dep (old, new) dep = if dep = old then new else [dep] blanchet@36486: fun replace_deps_in_line _ (line as Definition _) = line blanchet@36486: | replace_deps_in_line p (Inference (num, t, deps)) = blanchet@36486: Inference (num, t, fold (union (op =) o replace_one_dep p) deps []) paulson@21978: paulson@22491: (*Discard axioms; consolidate adjacent lines that prove the same clause, since they differ paulson@22491: only in type information.*) blanchet@36551: fun add_line _ _ (line as Definition _) lines = line :: lines blanchet@36551: | add_line conjecture_shape thm_names (Inference (num, t, [])) lines = blanchet@36570: (* No dependencies: axiom, conjecture clause, or internal axioms or blanchet@36570: definitions (Vampire). *) blanchet@36486: if is_axiom_clause_number thm_names num then blanchet@36486: (* Axioms are not proof lines. *) blanchet@36486: if is_only_type_information t then blanchet@36486: map (replace_deps_in_line (num, [])) lines blanchet@36486: (* Is there a repetition? If so, replace later line by earlier one. *) blanchet@36486: else case take_prefix (aint_inference t) lines of blanchet@36486: (_, []) => lines (*no repetition of proof line*) blanchet@36486: | (pre, Inference (num', _, _) :: post) => blanchet@36486: pre @ map (replace_deps_in_line (num', [num])) post blanchet@36570: else if is_conjecture_clause_number conjecture_shape num then blanchet@36486: Inference (num, t, []) :: lines blanchet@36551: else blanchet@36570: map (replace_deps_in_line (num, [])) lines blanchet@36551: | add_line _ _ (Inference (num, t, deps)) lines = blanchet@36486: (* Type information will be deleted later; skip repetition test. *) blanchet@36486: if is_only_type_information t then blanchet@36486: Inference (num, t, deps) :: lines blanchet@36486: (* Is there a repetition? If so, replace later line by earlier one. *) blanchet@36486: else case take_prefix (aint_inference t) lines of blanchet@36486: (* FIXME: Doesn't this code risk conflating proofs involving different blanchet@36486: types?? *) blanchet@36486: (_, []) => Inference (num, t, deps) :: lines blanchet@36486: | (pre, Inference (num', t', _) :: post) => blanchet@36486: Inference (num, t', deps) :: blanchet@36486: pre @ map (replace_deps_in_line (num', [num])) post paulson@22044: blanchet@36486: (* Recursively delete empty lines (type information) from the proof. *) blanchet@36486: fun add_nontrivial_line (Inference (num, t, [])) lines = blanchet@36486: if is_only_type_information t then delete_dep num lines blanchet@36486: else Inference (num, t, []) :: lines blanchet@36486: | add_nontrivial_line line lines = line :: lines blanchet@36395: and delete_dep num lines = blanchet@36486: fold_rev add_nontrivial_line (map (replace_deps_in_line (num, [])) lines) [] blanchet@36486: blanchet@36560: (* ATPs sometimes reuse free variable names in the strangest ways. Surprisingly, blanchet@36560: removing the offending lines often does the trick. *) blanchet@36560: fun is_bad_free frees (Free x) = not (member (op =) frees x) blanchet@36560: | is_bad_free _ _ = false paulson@22470: blanchet@36570: (* Vampire is keen on producing these. *) blanchet@36570: fun is_trivial_formula (@{const Not} $ (Const (@{const_name "op ="}, _) blanchet@36570: $ t1 $ t2)) = (t1 aconv t2) blanchet@36570: | is_trivial_formula t = false blanchet@36570: blanchet@36570: fun add_desired_line _ _ _ _ _ (line as Definition _) (j, lines) = blanchet@36560: (j, line :: lines) blanchet@36924: | add_desired_line ctxt isar_shrink_factor conjecture_shape thm_names frees blanchet@36570: (Inference (num, t, deps)) (j, lines) = blanchet@36402: (j + 1, blanchet@36570: if is_axiom_clause_number thm_names num orelse blanchet@36570: is_conjecture_clause_number conjecture_shape num orelse blanchet@36570: (not (is_only_type_information t) andalso blanchet@36570: null (Term.add_tvars t []) andalso blanchet@36570: not (exists_subterm (is_bad_free frees) t) andalso blanchet@36570: not (is_trivial_formula t) andalso blanchet@36570: (null lines orelse (* last line must be kept *) blanchet@36924: (length deps >= 2 andalso j mod isar_shrink_factor = 0))) then blanchet@36570: Inference (num, t, deps) :: lines (* keep line *) blanchet@36402: else blanchet@36570: map (replace_deps_in_line (num, deps)) lines) (* drop line *) paulson@21978: blanchet@36402: (** EXTRACTING LEMMAS **) paulson@21979: blanchet@36223: (* A list consisting of the first number in each line is returned. blanchet@36395: TSTP: Interesting lines have the form "cnf(108, axiom, ...)", where the blanchet@36223: number (108) is extracted. blanchet@36395: SPASS: Lines have the form "108[0:Inp] ...", where the first number (108) is blanchet@36223: extracted. *) blanchet@36402: fun extract_clause_numbers_in_atp_proof atp_proof = blanchet@35865: let blanchet@36395: val tokens_of = String.tokens (not o is_ident_char) blanchet@36402: fun extract_num ("cnf" :: num :: "axiom" :: _) = Int.fromString num blanchet@36395: | extract_num (num :: "0" :: "Inp" :: _) = Int.fromString num blanchet@36395: | extract_num _ = NONE blanchet@36402: in atp_proof |> split_lines |> map_filter (extract_num o tokens_of) end wenzelm@33310: blanchet@36968: (* Used to label theorems chained into the goal. *) blanchet@36395: val chained_hint = "sledgehammer_chained" blanchet@35865: blanchet@36063: fun apply_command _ 1 = "by " blanchet@36063: | apply_command 1 _ = "apply " blanchet@36063: | apply_command i _ = "prefer " ^ string_of_int i ^ " apply " blanchet@36570: fun metis_command i n [] = apply_command i n ^ "metis" blanchet@36570: | metis_command i n ss = blanchet@36570: apply_command i n ^ "(metis " ^ space_implode " " ss ^ ")" blanchet@36063: fun metis_line i n xs = blanchet@36063: "Try this command: " ^ blanchet@36063: Markup.markup Markup.sendback (metis_command i n xs) ^ ".\n" blanchet@36281: fun minimize_line _ [] = "" blanchet@36281: | minimize_line minimize_command facts = blanchet@36281: case minimize_command facts of blanchet@36281: "" => "" blanchet@36281: | command => blanchet@36065: "To minimize the number of lemmas, try this command: " ^ blanchet@36281: Markup.markup Markup.sendback command ^ ".\n" immler@31840: krauss@36606: fun metis_proof_text (minimize_command, atp_proof, thm_names, goal, i) = blanchet@36063: let blanchet@36231: val lemmas = blanchet@36402: atp_proof |> extract_clause_numbers_in_atp_proof blanchet@36402: |> filter (is_axiom_clause_number thm_names) blanchet@36402: |> map (fn i => Vector.sub (thm_names, i - 1)) blanchet@36402: |> filter_out (fn s => s = "??.unknown" orelse s = chained_hint) blanchet@36402: |> sort_distinct string_ord blanchet@36063: val n = Logic.count_prems (prop_of goal) blanchet@36395: in (metis_line i n lemmas ^ minimize_line minimize_command lemmas, lemmas) end immler@31037: blanchet@36486: (** Isar proof construction and manipulation **) blanchet@36486: blanchet@36486: fun merge_fact_sets (ls1, ss1) (ls2, ss2) = blanchet@36486: (union (op =) ls1 ls2, union (op =) ss1 ss2) blanchet@36402: blanchet@36402: type label = string * int blanchet@36402: type facts = label list * string list blanchet@36402: blanchet@36402: datatype qualifier = Show | Then | Moreover | Ultimately blanchet@36291: blanchet@36402: datatype step = blanchet@36478: Fix of (string * typ) list | blanchet@36486: Let of term * term | blanchet@36402: Assume of label * term | blanchet@36402: Have of qualifier list * label * term * byline blanchet@36402: and byline = blanchet@36564: ByMetis of facts | blanchet@36402: CaseSplit of step list list * facts blanchet@36402: blanchet@36574: fun smart_case_split [] facts = ByMetis facts blanchet@36574: | smart_case_split proofs facts = CaseSplit (proofs, facts) blanchet@36574: blanchet@36402: val raw_prefix = "X" blanchet@36402: val assum_prefix = "A" blanchet@36402: val fact_prefix = "F" blanchet@36402: blanchet@36570: fun string_for_label (s, num) = s ^ string_of_int num blanchet@36570: blanchet@36475: fun add_fact_from_dep thm_names num = blanchet@36475: if is_axiom_clause_number thm_names num then blanchet@36480: apsnd (insert (op =) (Vector.sub (thm_names, num - 1))) blanchet@36475: else blanchet@36480: apfst (insert (op =) (raw_prefix, num)) blanchet@36402: blanchet@36491: fun forall_vars t = fold_rev forall_of (map Var (Term.add_vars t [])) t blanchet@36491: blanchet@36486: fun step_for_line _ _ (Definition (num, t1, t2)) = Let (t1, t2) blanchet@36486: | step_for_line _ _ (Inference (num, t, [])) = Assume ((raw_prefix, num), t) blanchet@36486: | step_for_line thm_names j (Inference (num, t, deps)) = blanchet@36486: Have (if j = 1 then [Show] else [], (raw_prefix, num), blanchet@36491: forall_vars t, blanchet@36564: ByMetis (fold (add_fact_from_dep thm_names) deps ([], []))) blanchet@36291: blanchet@36967: fun proof_from_atp_proof pool ctxt full_types tfrees isar_shrink_factor blanchet@36967: atp_proof conjecture_shape thm_names params frees = blanchet@36402: let blanchet@36486: val lines = blanchet@36574: atp_proof ^ "$" (* the $ sign acts as a sentinel *) blanchet@36548: |> parse_proof pool blanchet@36967: |> decode_lines ctxt full_types tfrees blanchet@36551: |> rpair [] |-> fold_rev (add_line conjecture_shape thm_names) blanchet@36486: |> rpair [] |-> fold_rev add_nontrivial_line blanchet@36924: |> rpair (0, []) |-> fold_rev (add_desired_line ctxt isar_shrink_factor blanchet@36570: conjecture_shape thm_names frees) blanchet@36486: |> snd blanchet@36402: in blanchet@36909: (if null params then [] else [Fix params]) @ blanchet@36486: map2 (step_for_line thm_names) (length lines downto 1) lines blanchet@36402: end blanchet@36402: blanchet@36402: val indent_size = 2 blanchet@36402: val no_label = ("", ~1) blanchet@36402: blanchet@36402: fun no_show qs = not (member (op =) qs Show) blanchet@36402: blanchet@36402: (* When redirecting proofs, we keep information about the labels seen so far in blanchet@36402: the "backpatches" data structure. The first component indicates which facts blanchet@36402: should be associated with forthcoming proof steps. The second component is a blanchet@36402: pair ("keep_ls", "drop_ls"), where "keep_ls" are the labels to keep and blanchet@36402: "drop_ls" are those that should be dropped in a case split. *) blanchet@36402: type backpatches = (label * facts) list * (label list * label list) blanchet@36402: blanchet@36556: fun used_labels_of_step (Have (_, _, _, by)) = blanchet@36402: (case by of blanchet@36564: ByMetis (ls, _) => ls blanchet@36556: | CaseSplit (proofs, (ls, _)) => blanchet@36556: fold (union (op =) o used_labels_of) proofs ls) blanchet@36556: | used_labels_of_step _ = [] blanchet@36556: and used_labels_of proof = fold (union (op =) o used_labels_of_step) proof [] blanchet@36402: blanchet@36402: fun new_labels_of_step (Fix _) = [] blanchet@36486: | new_labels_of_step (Let _) = [] blanchet@36402: | new_labels_of_step (Assume (l, _)) = [l] blanchet@36402: | new_labels_of_step (Have (_, l, _, _)) = [l] blanchet@36402: val new_labels_of = maps new_labels_of_step blanchet@36402: blanchet@36402: val join_proofs = blanchet@36402: let blanchet@36402: fun aux _ [] = NONE blanchet@36402: | aux proof_tail (proofs as (proof1 :: _)) = blanchet@36402: if exists null proofs then blanchet@36402: NONE blanchet@36402: else if forall (curry (op =) (hd proof1) o hd) (tl proofs) then blanchet@36402: aux (hd proof1 :: proof_tail) (map tl proofs) blanchet@36402: else case hd proof1 of blanchet@36402: Have ([], l, t, by) => blanchet@36402: if forall (fn Have ([], l', t', _) :: _ => (l, t) = (l', t') blanchet@36402: | _ => false) (tl proofs) andalso blanchet@36402: not (exists (member (op =) (maps new_labels_of proofs)) blanchet@36556: (used_labels_of proof_tail)) then blanchet@36402: SOME (l, t, map rev proofs, proof_tail) blanchet@36402: else blanchet@36402: NONE blanchet@36402: | _ => NONE blanchet@36402: in aux [] o map rev end blanchet@36402: blanchet@36402: fun case_split_qualifiers proofs = blanchet@36402: case length proofs of blanchet@36402: 0 => [] blanchet@36402: | 1 => [Then] blanchet@36402: | _ => [Ultimately] blanchet@36402: blanchet@36491: fun redirect_proof thy conjecture_shape hyp_ts concl_t proof = wenzelm@33310: let blanchet@36402: val concl_ls = map (pair raw_prefix) (List.last conjecture_shape) blanchet@36551: fun find_hyp num = nth hyp_ts (index_in_shape num conjecture_shape) blanchet@36402: fun first_pass ([], contra) = ([], contra) blanchet@36491: | first_pass ((step as Fix _) :: proof, contra) = blanchet@36491: first_pass (proof, contra) |>> cons step blanchet@36491: | first_pass ((step as Let _) :: proof, contra) = blanchet@36491: first_pass (proof, contra) |>> cons step blanchet@36551: | first_pass ((step as Assume (l as (_, num), t)) :: proof, contra) = blanchet@36402: if member (op =) concl_ls l then blanchet@36491: first_pass (proof, contra ||> cons step) blanchet@36402: else blanchet@36551: first_pass (proof, contra) |>> cons (Assume (l, find_hyp num)) blanchet@36564: | first_pass ((step as Have (qs, l, t, ByMetis (ls, ss))) :: proof, blanchet@36491: contra) = blanchet@36402: if exists (member (op =) (fst contra)) ls then blanchet@36491: first_pass (proof, contra |>> cons l ||> cons step) blanchet@36402: else blanchet@36491: first_pass (proof, contra) |>> cons step blanchet@36402: | first_pass _ = raise Fail "malformed proof" blanchet@36402: val (proof_top, (contra_ls, contra_proof)) = blanchet@36402: first_pass (proof, (concl_ls, [])) blanchet@36402: val backpatch_label = the_default ([], []) oo AList.lookup (op =) o fst blanchet@36402: fun backpatch_labels patches ls = blanchet@36402: fold merge_fact_sets (map (backpatch_label patches) ls) ([], []) blanchet@36402: fun second_pass end_qs ([], assums, patches) = blanchet@36402: ([Have (end_qs, no_label, blanchet@36402: if length assums < length concl_ls then blanchet@36491: clause_for_literals thy (map (negate_term thy o fst) assums) blanchet@36402: else blanchet@36402: concl_t, blanchet@36564: ByMetis (backpatch_labels patches (map snd assums)))], patches) blanchet@36402: | second_pass end_qs (Assume (l, t) :: proof, assums, patches) = blanchet@36402: second_pass end_qs (proof, (t, l) :: assums, patches) blanchet@36564: | second_pass end_qs (Have (qs, l, t, ByMetis (ls, ss)) :: proof, assums, blanchet@36402: patches) = blanchet@36402: if member (op =) (snd (snd patches)) l andalso blanchet@36402: not (AList.defined (op =) (fst patches) l) then blanchet@36402: second_pass end_qs (proof, assums, patches ||> apsnd (append ls)) blanchet@36402: else blanchet@36402: (case List.partition (member (op =) contra_ls) ls of blanchet@36402: ([contra_l], co_ls) => blanchet@36402: if no_show qs then blanchet@36402: second_pass end_qs blanchet@36402: (proof, assums, blanchet@36402: patches |>> cons (contra_l, (l :: co_ls, ss))) blanchet@36402: |>> cons (if member (op =) (fst (snd patches)) l then blanchet@36491: Assume (l, negate_term thy t) blanchet@36402: else blanchet@36491: Have (qs, l, negate_term thy t, blanchet@36564: ByMetis (backpatch_label patches l))) blanchet@36402: else blanchet@36402: second_pass end_qs (proof, assums, blanchet@36402: patches |>> cons (contra_l, (co_ls, ss))) blanchet@36402: | (contra_ls as _ :: _, co_ls) => blanchet@36402: let blanchet@36402: val proofs = blanchet@36402: map_filter blanchet@36402: (fn l => blanchet@36402: if member (op =) concl_ls l then blanchet@36402: NONE blanchet@36402: else blanchet@36402: let blanchet@36402: val drop_ls = filter (curry (op <>) l) contra_ls blanchet@36402: in blanchet@36402: second_pass [] blanchet@36402: (proof, assums, blanchet@36402: patches ||> apfst (insert (op =) l) blanchet@36402: ||> apsnd (union (op =) drop_ls)) blanchet@36402: |> fst |> SOME blanchet@36402: end) contra_ls blanchet@36402: val facts = (co_ls, []) blanchet@36402: in blanchet@36402: (case join_proofs proofs of blanchet@36402: SOME (l, t, proofs, proof_tail) => blanchet@36402: Have (case_split_qualifiers proofs @ blanchet@36402: (if null proof_tail then end_qs else []), l, t, blanchet@36574: smart_case_split proofs facts) :: proof_tail blanchet@36402: | NONE => blanchet@36402: [Have (case_split_qualifiers proofs @ end_qs, no_label, blanchet@36574: concl_t, smart_case_split proofs facts)], blanchet@36402: patches) blanchet@36402: end blanchet@36402: | _ => raise Fail "malformed proof") blanchet@36402: | second_pass _ _ = raise Fail "malformed proof" blanchet@36486: val proof_bottom = blanchet@36486: second_pass [Show] (contra_proof, [], ([], ([], []))) |> fst blanchet@36402: in proof_top @ proof_bottom end blanchet@36402: blanchet@36402: val kill_duplicate_assumptions_in_proof = blanchet@36402: let blanchet@36402: fun relabel_facts subst = blanchet@36402: apfst (map (fn l => AList.lookup (op =) subst l |> the_default l)) blanchet@36491: fun do_step (step as Assume (l, t)) (proof, subst, assums) = blanchet@36402: (case AList.lookup (op aconv) assums t of blanchet@36967: SOME l' => (proof, (l, l') :: subst, assums) blanchet@36491: | NONE => (step :: proof, subst, (t, l) :: assums)) blanchet@36402: | do_step (Have (qs, l, t, by)) (proof, subst, assums) = blanchet@36402: (Have (qs, l, t, blanchet@36402: case by of blanchet@36564: ByMetis facts => ByMetis (relabel_facts subst facts) blanchet@36402: | CaseSplit (proofs, facts) => blanchet@36402: CaseSplit (map do_proof proofs, relabel_facts subst facts)) :: blanchet@36402: proof, subst, assums) blanchet@36491: | do_step step (proof, subst, assums) = (step :: proof, subst, assums) blanchet@36402: and do_proof proof = fold do_step proof ([], [], []) |> #1 |> rev blanchet@36402: in do_proof end blanchet@36402: blanchet@36402: val then_chain_proof = blanchet@36402: let blanchet@36402: fun aux _ [] = [] blanchet@36491: | aux _ ((step as Assume (l, _)) :: proof) = step :: aux l proof blanchet@36402: | aux l' (Have (qs, l, t, by) :: proof) = blanchet@36402: (case by of blanchet@36564: ByMetis (ls, ss) => blanchet@36402: Have (if member (op =) ls l' then blanchet@36402: (Then :: qs, l, t, blanchet@36564: ByMetis (filter_out (curry (op =) l') ls, ss)) blanchet@36402: else blanchet@36564: (qs, l, t, ByMetis (ls, ss))) blanchet@36402: | CaseSplit (proofs, facts) => blanchet@36402: Have (qs, l, t, CaseSplit (map (aux no_label) proofs, facts))) :: blanchet@36402: aux l proof blanchet@36491: | aux _ (step :: proof) = step :: aux no_label proof blanchet@36402: in aux no_label end blanchet@36402: blanchet@36402: fun kill_useless_labels_in_proof proof = blanchet@36402: let blanchet@36556: val used_ls = used_labels_of proof blanchet@36402: fun do_label l = if member (op =) used_ls l then l else no_label blanchet@36556: fun do_step (Assume (l, t)) = Assume (do_label l, t) blanchet@36556: | do_step (Have (qs, l, t, by)) = blanchet@36402: Have (qs, do_label l, t, blanchet@36402: case by of blanchet@36402: CaseSplit (proofs, facts) => blanchet@36556: CaseSplit (map (map do_step) proofs, facts) blanchet@36402: | _ => by) blanchet@36556: | do_step step = step blanchet@36556: in map do_step proof end blanchet@36402: blanchet@36402: fun prefix_for_depth n = replicate_string (n + 1) blanchet@36402: blanchet@36402: val relabel_proof = blanchet@36402: let blanchet@36402: fun aux _ _ _ [] = [] blanchet@36402: | aux subst depth (next_assum, next_fact) (Assume (l, t) :: proof) = blanchet@36402: if l = no_label then blanchet@36402: Assume (l, t) :: aux subst depth (next_assum, next_fact) proof blanchet@36402: else blanchet@36402: let val l' = (prefix_for_depth depth assum_prefix, next_assum) in blanchet@36402: Assume (l', t) :: blanchet@36402: aux ((l, l') :: subst) depth (next_assum + 1, next_fact) proof blanchet@36402: end blanchet@36402: | aux subst depth (next_assum, next_fact) (Have (qs, l, t, by) :: proof) = blanchet@36402: let blanchet@36402: val (l', subst, next_fact) = blanchet@36402: if l = no_label then blanchet@36402: (l, subst, next_fact) blanchet@36402: else blanchet@36402: let blanchet@36402: val l' = (prefix_for_depth depth fact_prefix, next_fact) blanchet@36402: in (l', (l, l') :: subst, next_fact + 1) end blanchet@36570: val relabel_facts = blanchet@36570: apfst (map (fn l => blanchet@36570: case AList.lookup (op =) subst l of blanchet@36570: SOME l' => l' blanchet@36570: | NONE => raise Fail ("unknown label " ^ blanchet@36570: quote (string_for_label l)))) blanchet@36402: val by = blanchet@36402: case by of blanchet@36564: ByMetis facts => ByMetis (relabel_facts facts) blanchet@36402: | CaseSplit (proofs, facts) => blanchet@36402: CaseSplit (map (aux subst (depth + 1) (1, 1)) proofs, blanchet@36402: relabel_facts facts) blanchet@36402: in blanchet@36402: Have (qs, l', t, by) :: blanchet@36402: aux subst depth (next_assum, next_fact) proof blanchet@36402: end blanchet@36491: | aux subst depth nextp (step :: proof) = blanchet@36491: step :: aux subst depth nextp proof blanchet@36402: in aux [] 0 (1, 1) end blanchet@36402: blanchet@36488: fun string_for_proof ctxt i n = blanchet@36402: let blanchet@36478: fun fix_print_mode f = wenzelm@37146: Print_Mode.setmp (filter (curry (op =) Symbol.xsymbolsN) blanchet@36478: (print_mode_value ())) f blanchet@36402: fun do_indent ind = replicate_string (ind * indent_size) " " blanchet@36478: fun do_free (s, T) = blanchet@36478: maybe_quote s ^ " :: " ^ blanchet@36478: maybe_quote (fix_print_mode (Syntax.string_of_typ ctxt) T) blanchet@36570: fun do_label l = if l = no_label then "" else string_for_label l ^ ": " blanchet@36402: fun do_have qs = blanchet@36402: (if member (op =) qs Moreover then "moreover " else "") ^ blanchet@36402: (if member (op =) qs Ultimately then "ultimately " else "") ^ blanchet@36402: (if member (op =) qs Then then blanchet@36402: if member (op =) qs Show then "thus" else "hence" blanchet@36402: else blanchet@36402: if member (op =) qs Show then "show" else "have") blanchet@36478: val do_term = maybe_quote o fix_print_mode (Syntax.string_of_term ctxt) blanchet@36570: fun do_facts (ls, ss) = blanchet@36570: let blanchet@36570: val ls = ls |> sort_distinct (prod_ord string_ord int_ord) blanchet@36570: val ss = ss |> sort_distinct string_ord blanchet@36570: in metis_command 1 1 (map string_for_label ls @ ss) end blanchet@36478: and do_step ind (Fix xs) = blanchet@36478: do_indent ind ^ "fix " ^ space_implode " and " (map do_free xs) ^ "\n" blanchet@36486: | do_step ind (Let (t1, t2)) = blanchet@36486: do_indent ind ^ "let " ^ do_term t1 ^ " = " ^ do_term t2 ^ "\n" blanchet@36402: | do_step ind (Assume (l, t)) = blanchet@36402: do_indent ind ^ "assume " ^ do_label l ^ do_term t ^ "\n" blanchet@36564: | do_step ind (Have (qs, l, t, ByMetis facts)) = blanchet@36402: do_indent ind ^ do_have qs ^ " " ^ blanchet@36479: do_label l ^ do_term t ^ " " ^ do_facts facts ^ "\n" blanchet@36402: | do_step ind (Have (qs, l, t, CaseSplit (proofs, facts))) = blanchet@36402: space_implode (do_indent ind ^ "moreover\n") blanchet@36402: (map (do_block ind) proofs) ^ blanchet@36479: do_indent ind ^ do_have qs ^ " " ^ do_label l ^ do_term t ^ " " ^ blanchet@36478: do_facts facts ^ "\n" blanchet@36402: and do_steps prefix suffix ind steps = blanchet@36402: let val s = implode (map (do_step ind) steps) in blanchet@36402: replicate_string (ind * indent_size - size prefix) " " ^ prefix ^ blanchet@36402: String.extract (s, ind * indent_size, blanchet@36402: SOME (size s - ind * indent_size - 1)) ^ blanchet@36402: suffix ^ "\n" blanchet@36402: end blanchet@36402: and do_block ind proof = do_steps "{ " " }" (ind + 1) proof blanchet@36564: (* One-step proofs are pointless; better use the Metis one-liner blanchet@36564: directly. *) blanchet@36564: and do_proof [Have (_, _, _, ByMetis _)] = "" blanchet@36564: | do_proof proof = blanchet@36480: (if i <> 1 then "prefer " ^ string_of_int i ^ "\n" else "") ^ blanchet@36480: do_indent 0 ^ "proof -\n" ^ blanchet@36480: do_steps "" "" 1 proof ^ blanchet@36480: do_indent 0 ^ (if n <> 1 then "next" else "qed") ^ "\n" blanchet@36488: in do_proof end blanchet@36402: blanchet@36924: fun isar_proof_text (pool, debug, full_types, isar_shrink_factor, ctxt, blanchet@36909: conjecture_shape) krauss@36606: (minimize_command, atp_proof, thm_names, goal, i) = blanchet@36402: let blanchet@36402: val thy = ProofContext.theory_of ctxt blanchet@36909: val (params, hyp_ts, concl_t) = strip_subgoal goal i blanchet@36909: val frees = fold Term.add_frees (concl_t :: hyp_ts) [] blanchet@36967: val tfrees = fold Term.add_tfrees (concl_t :: hyp_ts) [] blanchet@36402: val n = Logic.count_prems (prop_of goal) blanchet@36223: val (one_line_proof, lemma_names) = blanchet@36402: metis_proof_text (minimize_command, atp_proof, thm_names, goal, i) blanchet@36283: fun isar_proof_for () = blanchet@36967: case proof_from_atp_proof pool ctxt full_types tfrees isar_shrink_factor blanchet@36924: atp_proof conjecture_shape thm_names params blanchet@36924: frees blanchet@36491: |> redirect_proof thy conjecture_shape hyp_ts concl_t blanchet@36402: |> kill_duplicate_assumptions_in_proof blanchet@36402: |> then_chain_proof blanchet@36402: |> kill_useless_labels_in_proof blanchet@36402: |> relabel_proof blanchet@36488: |> string_for_proof ctxt i n of blanchet@36283: "" => "" blanchet@36402: | proof => "\nStructured proof:\n" ^ Markup.markup Markup.sendback proof blanchet@35868: val isar_proof = blanchet@36402: if debug then blanchet@36283: isar_proof_for () blanchet@36283: else blanchet@36283: try isar_proof_for () blanchet@36287: |> the_default "Warning: The Isar proof construction failed.\n" blanchet@36283: in (one_line_proof ^ isar_proof, lemma_names) end paulson@21978: blanchet@36557: fun proof_text isar_proof isar_params other_params = blanchet@36557: (if isar_proof then isar_proof_text isar_params else metis_proof_text) blanchet@36557: other_params blanchet@36223: immler@31038: end;