File ‹TPTP_Parser/tptp_reconstruct.ML›
signature TPTP_RECONSTRUCT =
sig
datatype formula_kind =
Conjunctive of bool option
| Disjunctive of bool option
| Biimplicational of bool option
| Negative of bool option
| Existential of bool option * typ
| Universal of bool option * typ
| Equational of bool option * typ
| Atomic of bool option
| Implicational of bool option
type formula_meaning =
(string *
{role : TPTP_Syntax.role,
fmla : term,
source_inf_opt : TPTP_Proof.source_info option})
type proof_annotation =
{problem_name : TPTP_Problem_Name.problem_name,
skolem_defs : (string * Binding.binding) list,
defs : (string * Binding.binding) list,
axs : (string * Binding.binding) list,
meta : formula_meaning list}
type rule_info =
{inference_name : string,
inference_fmla : term,
parents : string list}
exception UNPOLARISED of term
val remove_polarity : bool -> term -> term * bool
val interpret_bindings :
TPTP_Problem_Name.problem_name -> theory -> TPTP_Proof.parent_detail list -> (string * term) list -> (string * term) list
val diff_and_instantiate : Proof.context -> thm -> term -> term -> thm
val strip_top_all_vars : (string * typ) list -> term -> (string * typ) list * term
val strip_top_All_vars : term -> (string * typ) list * term
val strip_top_All_var : term -> (string * typ) * term
val new_consts_between : term -> term -> term list
val get_pannot_of_prob : theory -> TPTP_Problem_Name.problem_name -> proof_annotation
val inference_at_node : 'a -> TPTP_Problem_Name.problem_name -> formula_meaning list -> string -> rule_info option
val node_info : (string * 'a) list -> ('a -> 'b) -> string -> 'b
type step_id = string
datatype rolling_stock =
Step of step_id
| Assumed
| Unconjoin
| Split of step_id *
step_id *
step_id list
| Synth_step of step_id
| Annotated_step of step_id * string
| Definition of step_id
| Axiom of step_id
| Caboose
val import_thm : bool -> Path.T list -> Path.T -> (proof_annotation -> theory -> proof_annotation * theory) -> theory -> theory
val get_fmlas_of_prob : theory -> TPTP_Problem_Name.problem_name -> TPTP_Interpret.tptp_formula_meaning list
val structure_fmla_meaning : 'a * 'b * 'c * 'd -> 'a * {fmla: 'c, role: 'b, source_inf_opt: 'd}
val make_skeleton : Proof.context -> proof_annotation -> rolling_stock list
val naive_reconstruct_tacs :
(Proof.context -> TPTP_Problem_Name.problem_name -> step_id -> thm) ->
TPTP_Problem_Name.problem_name -> Proof.context -> (rolling_stock * term option * (thm * tactic) option) list
val naive_reconstruct_tac :
Proof.context -> (Proof.context -> TPTP_Problem_Name.problem_name -> step_id -> thm) -> TPTP_Problem_Name.problem_name -> tactic
val reconstruct : Proof.context -> (TPTP_Problem_Name.problem_name -> tactic) -> TPTP_Problem_Name.problem_name -> thm
end
structure TPTP_Reconstruct : TPTP_RECONSTRUCT =
struct
open TPTP_Reconstruct_Library
open TPTP_Syntax
fun get_fmlas_of_prob thy prob_name : TPTP_Interpret.tptp_formula_meaning list =
AList.lookup (op =) (TPTP_Interpret.get_manifests thy) prob_name
|> the |> #3 ;
type formula_meaning =
(string *
{role : TPTP_Syntax.role,
fmla : term,
source_inf_opt : TPTP_Proof.source_info option})
fun apply_to_parent_info f
(n, {role, fmla, source_inf_opt}) =
let
val source_inf_opt' =
case source_inf_opt of
NONE => NONE
| SOME (TPTP_Proof.Inference (inf_name, sinfos, pinfos)) =>
SOME (TPTP_Proof.Inference (inf_name, sinfos, f pinfos))
in
(n, {role = role, fmla = fmla, source_inf_opt = source_inf_opt'})
end
fun structure_fmla_meaning (s, r, t, info) =
(s, {role = r, fmla = t, source_inf_opt = info})
type proof_annotation =
{problem_name : TPTP_Problem_Name.problem_name,
skolem_defs : (string * Binding.binding) list,
defs : (string * Binding.binding) list,
axs : (string * Binding.binding) list,
meta : formula_meaning list}
fun empty_pannot prob_name =
{problem_name = prob_name,
skolem_defs = [],
defs = [],
axs = [],
meta = []}
exception MANIFEST of TPTP_Problem_Name.problem_name * string
type manifest = TPTP_Problem_Name.problem_name * proof_annotation
fun manifest_eq ((prob_name1, _), (prob_name2, _)) = prob_name1 = prob_name2
structure TPTP_Reconstruction_Data = Theory_Data
(
type T = manifest list
val empty = []
fun merge data : T = Library.merge manifest_eq data
)
val get_manifests : theory -> manifest list = TPTP_Reconstruction_Data.get
fun update_manifest prob_name pannot thy =
let
val idx =
find_index
(fn (n, _) => n = prob_name)
(get_manifests thy)
val transf = (fn _ =>
(prob_name, pannot))
in
TPTP_Reconstruction_Data.map
(nth_map idx transf)
thy
end
fun get_pannot_of_prob thy prob_name : proof_annotation =
case AList.lookup (op =) (get_manifests thy) prob_name of
SOME pa => pa
| NONE => raise (MANIFEST (prob_name, "Could not find proof annotation"))
val inode_prefixK = "inode"
val bindK = "bind"
val split_preprocessingK = "split_preprocessing"
type tptp_reconstruction_state = {next_int : int}
structure TPTP_Reconstruction_Internal_Data = Theory_Data
(
type T = tptp_reconstruction_state
val empty = {next_int = 0}
fun merge data : T = snd data
)
fun get_next_int thy : int * theory =
let
val state = TPTP_Reconstruction_Internal_Data.get thy
val state' = {next_int = 1 + #next_int state}
in
(#next_int state,
TPTP_Reconstruction_Internal_Data.put state' thy)
end
val get_next_name =
get_next_int
#> apfst (fn i => inode_prefixK ^ Int.toString i)
exception NON_BINDING
fun interpret_bindings (prob_name : TPTP_Problem_Name.problem_name) thy bindings acc =
if null bindings then acc
else
case hd bindings of
TPTP_Proof.Bind (v, fmla) =>
let
val (type_map, const_map) =
case AList.lookup (op =) (TPTP_Interpret.get_manifests thy) prob_name of
NONE => raise (MANIFEST (prob_name, "Problem details not found in interpretation manifest"))
| SOME (type_map, const_map, _) => (type_map, const_map)
val config =
{cautious = true,
problem_name = SOME prob_name}
val result =
(v,
TPTP_Interpret.interpret_formula
config TPTP_Syntax.THF
const_map [] type_map fmla thy
|> fst)
in
interpret_bindings prob_name thy (tl bindings) (result :: acc)
end
| _ => raise NON_BINDING
type rule_info =
{inference_name : string,
inference_fmla : term,
parents : string list}
fun apply_binding thy prob_name orig_parent_fmla target_fmla bindings =
let
val bindings' = interpret_bindings prob_name thy bindings []
fun bind_free_vars var_ctxt t =
case t of
Const _ => t
| Var _ => t
| Bound _ => t
| Abs (x, ty, t') => Abs (x, ty, bind_free_vars (x :: var_ctxt) t')
| Free (x, ty) =>
let
val idx = find_index (fn y => y = x) var_ctxt
in
if idx > ~1 andalso
ty = dummyT then
Bound idx
else t
end
| t1 $ t2 => bind_free_vars var_ctxt t1 $ bind_free_vars var_ctxt t2
fun instantiate_bound (binding as (bound_var, body)) (initial as (var_ctxt, t)) =
case t of
Const _ => initial
| Free _ => initial
| Var _ => initial
| Bound _ => initial
| Abs _ => initial
| t1 $ (t2 as Abs (x, ty, t')) =>
if is_Const t1 then
if x = bound_var andalso
fst (dest_Const t1) = \<^const_name>‹All› then
let val body' = bind_free_vars var_ctxt body
in
(var_ctxt,
betapply (t2, body'))
end
else
let
val (var_ctxt', rest) = instantiate_bound binding (x :: var_ctxt, t')
in
(var_ctxt',
t1 $ Abs (x, ty, rest))
end
else initial
| t1 $ t2 =>
let
val (var_ctxt', rest) = instantiate_bound binding (var_ctxt, t2)
in
(var_ctxt', t1 $ rest)
end
fun safe_instantiate_bound (binding as (bound_var, body)) (var_ctxt, t) =
instantiate_bound binding
(var_ctxt, push_allvar_in bound_var t)
fun is_polymorphic tys =
if null tys then false
else
case hd tys of
Type (_, tys') => is_polymorphic (tl tys @ tys')
| TFree _ => true
| TVar _ => true
local
fun type_of_quantified_var' s ts =
if null ts then NONE
else
case hd ts of
Const _ => type_of_quantified_var' s (tl ts)
| Free _ => type_of_quantified_var' s (tl ts)
| Var _ => type_of_quantified_var' s (tl ts)
| Bound _ => type_of_quantified_var' s (tl ts)
| Abs (s', ty, t') =>
if s = s' then SOME ty
else type_of_quantified_var' s (t' :: tl ts)
| t1 $ t2 => type_of_quantified_var' s (t1 :: t2 :: tl ts)
in
fun type_of_quantified_var s =
single #> type_of_quantified_var' s
end
fun close_formula t =
let
val frees_monomorphised =
fold_aterms
(fn t => fn rest =>
if is_Free t then
let
val (s, ty) = dest_Free t
val ty' =
if ty = dummyT orelse is_polymorphic [ty] then
the (type_of_quantified_var s target_fmla)
else ty
in insert (op =) (t, Free (s, ty')) rest
end
else rest)
t []
in
Term.subst_free frees_monomorphised t
|> fold (fn (s, ty) => fn t =>
HOLogic.mk_all (s, ty, t))
(map (snd #> dest_Free) frees_monomorphised)
end
val _ = \<^assert> (length bindings' = 1)
in
fold safe_instantiate_bound bindings' ([], HOLogic.dest_Trueprop orig_parent_fmla)
|> snd
|> close_formula
|> single
|> Type_Infer_Context.infer_types (Context.proof_of (Context.Theory thy))
|> the_single
|> HOLogic.mk_Trueprop
|> rpair bindings'
end
exception RECONSTRUCT of string
datatype formula_kind =
Conjunctive of bool option
| Disjunctive of bool option
| Biimplicational of bool option
| Negative of bool option
| Existential of bool option * typ
| Universal of bool option * typ
| Equational of bool option * typ
| Atomic of bool option
| Implicational of bool option
exception UNPOLARISED of term
fun remove_polarity strict formula =
case try HOLogic.dest_eq formula of
NONE => if strict then raise (UNPOLARISED formula)
else (formula, true)
| SOME (x, p as \<^term>‹True›) => (x, true)
| SOME (x, p as \<^term>‹False›) => (x, false)
| SOME (x, _) =>
if strict then raise (UNPOLARISED formula)
else (formula, true)
fun flatten formula_kind formula =
let
fun is_conj (Const (\<^const_name>‹HOL.conj›, _) $ _ $ _) = true
| is_conj _ = false
fun is_disj (Const (\<^const_name>‹HOL.disj›, _) $ _ $ _) = true
| is_disj _ = false
fun is_iff (Const (\<^const_name>‹HOL.eq›, ty) $ _ $ _) =
ty = ([HOLogic.boolT, HOLogic.boolT] ---> HOLogic.boolT)
| is_iff _ = false
fun flatten' formula acc =
case formula of
Const (\<^const_name>‹HOL.conj›, _) $ t1 $ t2 =>
(case formula_kind of
Conjunctive _ =>
let
val left =
if is_conj t1 then flatten' t1 acc else (t1 :: acc)
in
if is_conj t2 then flatten' t2 left else (t2 :: left)
end
| _ => formula :: acc)
| Const (\<^const_name>‹HOL.disj›, _) $ t1 $ t2 =>
(case formula_kind of
Disjunctive _ =>
let
val left =
if is_disj t1 then flatten' t1 acc else (t1 :: acc)
in
if is_disj t2 then flatten' t2 left else (t2 :: left)
end
| _ => formula :: acc)
| Const (\<^const_name>‹HOL.eq›, ty) $ t1 $ t2 =>
if ty = ([HOLogic.boolT, HOLogic.boolT] ---> HOLogic.boolT) then
case formula_kind of
Biimplicational _ =>
let
val left =
if is_iff t1 then flatten' t1 acc else (t1 :: acc)
in
if is_iff t2 then flatten' t2 left else (t2 :: left)
end
| _ => formula :: acc
else formula :: acc
| _ => [formula]
val formula' = try_dest_Trueprop formula
in
case formula_kind of
Conjunctive (SOME _) =>
remove_polarity false formula'
|> fst
|> (fn t => flatten' t [])
| Disjunctive (SOME _) =>
remove_polarity false formula'
|> fst
|> (fn t => flatten' t [])
| Biimplicational (SOME _) =>
remove_polarity false formula'
|> fst
|> (fn t => flatten' t [])
| _ => flatten' formula' []
end
fun node_info fms projector node_name =
case AList.lookup (op =) fms node_name of
NONE =>
raise (RECONSTRUCT ("node " ^ node_name ^
" doesn't exist"))
| SOME info => projector info
fun dest_parent_infos filtered fms parent_infos : {name : string, details : TPTP_Proof.parent_detail list} list =
let
val filter_deps =
filter (fn {name, ...} =>
let
val role = node_info fms #role name
in role <> TPTP_Syntax.Role_Definition andalso
role <> TPTP_Syntax.Role_Axiom
end)
val parent_nodelist =
parent_infos
|> map (fn n =>
case n of
TPTP_Proof.Parent parent => {name = parent, details = []}
| TPTP_Proof.ParentWithDetails (parent, details) =>
{name = parent, details = details})
in
parent_nodelist
|> filtered ? filter_deps
end
fun parents_of_node fms n =
case node_info fms #source_inf_opt n of
NONE => []
| SOME (TPTP_Proof.File _) => []
| SOME (TPTP_Proof.Inference (_, _ : TPTP_Proof.useful_info_as list, parent_infos)) =>
dest_parent_infos false fms parent_infos
|> map #name
exception FIND_ANCESTOR_USING_RULE of string
fun find_ancestor_using_rule pannot inference_rule (fringe : string list) : string =
if null fringe then
raise (FIND_ANCESTOR_USING_RULE inference_rule)
else
case node_info (#meta pannot) #source_inf_opt (hd fringe) of
NONE => find_ancestor_using_rule pannot inference_rule (tl fringe)
| SOME (TPTP_Proof.File _) => find_ancestor_using_rule pannot inference_rule (tl fringe)
| SOME (TPTP_Proof.Inference (rule_name, _ : TPTP_Proof.useful_info_as list, parent_infos)) =>
if rule_name = inference_rule then hd fringe
else
find_ancestor_using_rule pannot inference_rule
(tl fringe @
map #name (dest_parent_infos true (#meta pannot) parent_infos))
fun inference_at_node thy (prob_name : TPTP_Problem_Name.problem_name)
(fms : formula_meaning list) from : rule_info option =
let
exception INFERENCE_AT_NODE of string
val fmla_of_node =
node_info fms #fmla
#> try_dest_Trueprop
fun build_inference_info rule_name parent_infos =
let
val _ = \<^assert> (not (null parent_infos))
val parent_nodes =
dest_parent_infos false fms parent_infos
|> map #name
val parent_fmlas = map fmla_of_node (rev parent_nodes)
val inference_term =
if null parent_fmlas then
fmla_of_node from
|> HOLogic.mk_Trueprop
else
Logic.mk_implies
(fold
(curry HOLogic.mk_conj)
(tl parent_fmlas)
(hd parent_fmlas)
|> HOLogic.mk_Trueprop,
fmla_of_node from |> HOLogic.mk_Trueprop)
in
SOME {inference_name = rule_name,
inference_fmla = inference_term,
parents = parent_nodes}
end
in
case node_info fms #source_inf_opt from of
NONE => NONE
| SOME (TPTP_Proof.File _) => NONE
| SOME (TPTP_Proof.Inference (rule_name, _ : TPTP_Proof.useful_info_as list, parent_infos)) =>
if List.null parent_infos then
raise (INFERENCE_AT_NODE
("empty parent list for node " ^
from ^ ": check proof format"))
else
build_inference_info rule_name parent_infos
end
fun make_elimination_rule_t ctxt major_prem prems_and_concs conclusion =
let
val thy = Proof_Context.theory_of ctxt
val minor_prems =
map (fn (v, conc) =>
Logic.mk_implies (v, HOLogic.mk_Trueprop conc))
prems_and_concs
in
(Logic.list_implies
(major_prem :: minor_prems,
conclusion))
end
fun simulate_split ctxt split_fmla minor_prem_assumptions conclusion =
let
val prems_and_concs =
ListPair.zip (minor_prem_assumptions, flatten (Conjunctive NONE) conclusion)
val rule_t = make_elimination_rule_t ctxt split_fmla prems_and_concs conclusion
val abstraction_subterms =
(map (try_dest_Trueprop #> remove_polarity true #> fst)
minor_prem_assumptions)
val abs_rule_t =
abstract
abstraction_subterms
rule_t
|> snd
val abs_rule_thm =
Goal.prove ctxt [] [] abs_rule_t
(fn pdata => HEADGOAL (blast_tac (#context pdata)))
|> Drule.export_without_context
in
diff_and_instantiate ctxt abs_rule_thm (Thm.prop_of abs_rule_thm) rule_t
end
type step_id = string
datatype rolling_stock =
Step of step_id
| Assumed
| Unconjoin
| Split of step_id *
step_id *
step_id list
| Synth_step of step_id
| Annotated_step of step_id * string
| Definition of step_id
| Axiom of step_id
| Caboose
fun stock_to_string (Step n) = n
| stock_to_string (Annotated_step (n, anno)) = n ^ "(" ^ anno ^ ")"
| stock_to_string _ = error "Stock is not a step"
fun filter_by_role tptp_role =
filter
(fn (_, info) =>
#role info = tptp_role)
fun filter_by_name node_name =
filter
(fn (n, _) =>
n = node_name)
exception NO_MARKER_NODE
fun proof_beginning_node fms =
let
val result =
cascaded_filter_single true
[filter_by_role TPTP_Syntax.Role_Conjecture,
filter_by_name "1"]
fms
in
case result of
SOME x => fst x
| NONE => raise NO_MARKER_NODE
end
fun proof_end_node fms =
fms
|> hd
|> fst
fun make_skeleton ctxt (pannot : proof_annotation) : rolling_stock list =
let
val thy = Proof_Context.theory_of ctxt
fun stock_is_ax_or_def (Axiom _) = true
| stock_is_ax_or_def (Definition _) = true
| stock_is_ax_or_def _ = false
fun stock_of n =
case node_info (#meta pannot) #role n of
TPTP_Syntax.Role_Definition => (true, Definition n)
| TPTP_Syntax.Role_Axiom => (true, Axiom n)
| _ => (false, Step n)
fun n_is_split_conjecture (inference_info : rule_info option) =
case inference_info of
NONE => false
| SOME inference_info => #inference_name inference_info = "split_conjecture"
exception SKELETON
fun check_parents stop_just_befores n =
let
val parents = parents_of_node (#meta pannot) n
in
if length parents = 1 then
AList.lookup (op =) stop_just_befores (the_single parents)
else
NONE
end
fun naive_skeleton' stop_just_befores n =
case check_parents stop_just_befores n of
SOME skel => skel
| NONE =>
let
val inference_info = inference_at_node thy (#problem_name pannot) (#meta pannot) n
in
if is_none inference_info then
if node_info (#meta pannot) #role n = TPTP_Syntax.Role_Definition then
[(Definition n), Assumed]
else if node_info (#meta pannot) #role n = TPTP_Syntax.Role_Axiom then
[Axiom n]
else raise SKELETON
else
let
val inference_info = the inference_info
val parents = #parents inference_info
in
if #inference_name inference_info = "solved_all_splits" andalso length parents > 1 then
let
val split_node =
find_ancestor_using_rule pannot "split_conjecture" [hd parents]
|> parents_of_node (#meta pannot)
|> the_single
val skeletons_up =
map (naive_skeleton' ((split_node, [Assumed]) :: stop_just_befores)) parents
in
Step n :: (Split (split_node, n, parents)) ::
naive_skeleton' stop_just_befores split_node @
flat skeletons_up @ [Assumed]
end
else if length parents > 1 then
let
val skeletons_up =
map (naive_skeleton' stop_just_befores) parents
in
Step n :: concat_between skeletons_up (SOME Unconjoin, NONE) @ [Assumed]
end
else
Step n :: naive_skeleton' stop_just_befores (the_single parents)
end
end
in
if List.null (#meta pannot) then []
else
naive_skeleton'
[(proof_beginning_node (#meta pannot), [Assumed])]
(proof_end_node (#meta pannot))
|> rev |> tl |> cons Caboose |> rev
end
exception SKELETON
local
val neg_eq_false =
@{lemma "!! P. (~ P) ==> (P = False)" by auto}
val solved_all_splits =
@{lemma "False = True ==> False" by auto}
fun skel_to_naive_tactic ctxt prover_tac prob_name skel memo = fn st =>
let
val thy = Proof_Context.theory_of ctxt
val pannot = get_pannot_of_prob thy prob_name
fun tac_and_memo node memo =
case AList.lookup (op =) memo node of
NONE =>
let
val tac =
prover_tac ctxt prob_name node
in (tac, (node, tac) :: memo) end
| SOME tac => (tac, memo)
fun rest skel' memo =
skel_to_naive_tactic ctxt prover_tac prob_name skel' memo
val tactic =
if null skel then
raise SKELETON
else
case hd skel of
Assumed => TRY (HEADGOAL (assume_tac ctxt)) THEN rest (tl skel) memo
| Caboose => TRY (HEADGOAL (assume_tac ctxt))
| Unconjoin => resolve_tac ctxt @{thms conjI} 1 THEN rest (tl skel) memo
| Split (split_node, solved_node, antes) =>
let
val split_fmla = node_info (#meta pannot) #fmla split_node
val conclusion =
(inference_at_node thy prob_name (#meta pannot) solved_node
|> the
|> #inference_fmla)
|> Logic.dest_implies
|> #1
val minor_prems_assumps =
map (fn ante => find_ancestor_using_rule pannot "split_conjecture" [ante]) antes
|> map (node_info (#meta pannot) #fmla)
val split_thm =
simulate_split ctxt split_fmla minor_prems_assumps conclusion
in
resolve_tac ctxt [split_thm] 1 THEN rest (tl skel) memo
end
| Step s =>
let
val (th, memo') = tac_and_memo s memo
in
resolve_tac ctxt [th] 1 THEN rest (tl skel) memo'
end
| Definition n =>
let
val def_thm =
case AList.lookup (op =) (#defs pannot) n of
NONE => error ("Did not find definition: " ^ n)
| SOME binding => Global_Theory.get_thm thy (Binding.name_of binding)
in
resolve_tac ctxt [def_thm] 1 THEN rest (tl skel) memo
end
| Axiom n =>
let
val ax_thm =
case AList.lookup (op =) (#axs pannot) n of
NONE => error ("Did not find axiom: " ^ n)
| SOME binding => Global_Theory.get_thm thy (Binding.name_of binding)
in
resolve_tac ctxt [ax_thm] 1 THEN rest (tl skel) memo
end
| _ => raise SKELETON
in tactic st end
fun skel_to_naive_tactic_dbg prover_tac ctxt prob_name skel (memo : (string * (thm * tactic) option) list) =
let
val thy = Proof_Context.theory_of ctxt
val pannot = get_pannot_of_prob thy prob_name
fun inference_name_of_node node =
case AList.lookup (op =) (#meta pannot) node of
NONE => (warning "Inference step lacks an inference name"; "(Shouldn't be here)")
| SOME info =>
case #source_inf_opt info of
SOME (TPTP_Proof.Inference (infname, _, _)) =>
infname
| _ => (warning "Inference step lacks an inference name"; "(Shouldn't be here)")
fun inference_fmla node =
case inference_at_node thy prob_name (#meta pannot) node of
NONE => NONE
| SOME {inference_fmla, ...} => SOME inference_fmla
fun rest memo' ctxt' = skel_to_naive_tactic_dbg prover_tac ctxt' prob_name (tl skel) memo'
val try_make_step =
(fn ctxt' =>
let
fun thm ctxt'' = prover_tac ctxt'' prob_name (hd skel |> stock_to_string)
val reconstructed_inference = thm ctxt'
fun rec_inf_tac st = HEADGOAL (resolve_tac ctxt' [thm ctxt']) st
in (reconstructed_inference,
rec_inf_tac)
end)
fun ignore_interpretation_exn f x = SOME (f x)
handle INTERPRET_INFERENCE => NONE
in
if List.null skel then
raise SKELETON
else
case hd skel of
Assumed =>
(hd skel,
Thm.prop_of @{thm asm_rl}
|> SOME,
SOME (@{thm asm_rl}, TRY (HEADGOAL (assume_tac ctxt)))) :: rest memo ctxt
| Caboose =>
[(Caboose,
Thm.prop_of @{thm asm_rl}
|> SOME,
SOME (@{thm asm_rl}, TRY (HEADGOAL (assume_tac ctxt))))]
| Unconjoin =>
(hd skel,
Thm.prop_of @{thm conjI}
|> SOME,
SOME (@{thm conjI}, resolve_tac ctxt @{thms conjI} 1)) :: rest memo ctxt
| Split (split_node, solved_node, antes) =>
let
val split_fmla = node_info (#meta pannot) #fmla split_node
val conclusion =
(inference_at_node thy prob_name (#meta pannot) solved_node
|> the
|> #inference_fmla)
|> Logic.dest_implies
|> #1
val minor_prems_assumps =
map (fn ante => find_ancestor_using_rule pannot "split_conjecture" [ante]) antes
|> map (node_info (#meta pannot) #fmla)
val split_thm =
simulate_split ctxt split_fmla minor_prems_assumps conclusion
in
(hd skel,
Thm.prop_of split_thm
|> SOME,
SOME (split_thm, resolve_tac ctxt [split_thm] 1)) :: rest memo ctxt
end
| Step node =>
let
val inference_name = inference_name_of_node node
val inference_fmla = inference_fmla node
val (inference_instance_thm, memo', ctxt') =
case AList.lookup (op =) memo node of
NONE =>
let
val (thm, ctxt') =
if is_some inference_fmla andalso
exceeds_tptp_max_term_size ctxt (Term.size_of_term (the inference_fmla)) then
(
warning ("Gave up on node " ^ node ^ " because of fmla size " ^
Int.toString (Term.size_of_term (the inference_fmla)));
(NONE, ctxt)
)
else
let
val maybe_thm = ignore_interpretation_exn try_make_step ctxt
in
(maybe_thm, ctxt)
end
in (thm, (node, thm) :: memo, ctxt') end
| SOME maybe_thm => (maybe_thm, memo, ctxt)
in
(Annotated_step (node, inference_name),
inference_fmla,
inference_instance_thm) :: rest memo' ctxt'
end
| Definition n =>
let
fun def_thm thy =
case AList.lookup (op =) (#defs pannot) n of
NONE => error ("Did not find definition: " ^ n)
| SOME binding => Global_Theory.get_thm thy (Binding.name_of binding)
in
(hd skel,
Thm.prop_of (def_thm thy)
|> SOME,
SOME (def_thm thy, HEADGOAL (resolve_tac ctxt [def_thm thy]))) :: rest memo ctxt
end
| Axiom n =>
let
val ax_thm =
case AList.lookup (op =) (#axs pannot) n of
NONE => error ("Did not find axiom: " ^ n)
| SOME binding => Global_Theory.get_thm thy (Binding.name_of binding)
in
(hd skel,
Thm.prop_of ax_thm
|> SOME,
SOME (ax_thm, resolve_tac ctxt [ax_thm] 1)) :: rest memo ctxt
end
end
fun sas_if_needed_tac ctxt prob_name =
let
val thy = Proof_Context.theory_of ctxt
val pannot = get_pannot_of_prob thy prob_name
val last_inference_info_opt =
find_first
(fn (_, info) => #role info = TPTP_Syntax.Role_Plain)
(#meta pannot)
val last_inference_info =
case last_inference_info_opt of
NONE => NONE
| SOME (_, info) => #source_inf_opt info
in
if is_some last_inference_info andalso
TPTP_Proof.is_inference_called "solved_all_splits"
(the last_inference_info)
then (@{thm asm_rl}, all_tac)
else (solved_all_splits, TRY (resolve_tac ctxt [solved_all_splits] 1))
end
in
fun naive_reconstruct_tac ctxt prover_tac prob_name =
let
val thy = Proof_Context.theory_of ctxt
in
resolve_tac ctxt @{thms ccontr} 1
THEN dresolve_tac ctxt [neg_eq_false] 1
THEN (sas_if_needed_tac ctxt prob_name |> #2)
THEN skel_to_naive_tactic ctxt prover_tac prob_name
(make_skeleton ctxt
(get_pannot_of_prob thy prob_name)) []
end
fun naive_reconstruct_tacs prover_tac prob_name ctxt =
let
val thy = Proof_Context.theory_of ctxt
in
(Synth_step "ccontr", Thm.prop_of @{thm ccontr} |> SOME,
SOME (@{thm ccontr}, resolve_tac ctxt @{thms ccontr} 1)) ::
(Synth_step "neg_eq_false", Thm.prop_of neg_eq_false |> SOME,
SOME (neg_eq_false, dresolve_tac ctxt [neg_eq_false] 1)) ::
(Synth_step "sas_if_needed_tac", Thm.prop_of @{thm asm_rl} |> SOME,
SOME (sas_if_needed_tac ctxt prob_name)) ::
skel_to_naive_tactic_dbg prover_tac ctxt prob_name
(make_skeleton ctxt
(get_pannot_of_prob thy prob_name)) []
end
end
fun reconstruct ctxt tactic prob_name =
let
val thy = Proof_Context.theory_of ctxt
val pannot = get_pannot_of_prob thy prob_name
val goal =
#meta pannot
|> filter (fn (_, info) =>
#role info = TPTP_Syntax.Role_Conjecture)
in
if null (#meta pannot) then
@{thm TrueI}
else if null goal then
raise (RECONSTRUCT "Proof lacks conjecture")
else
the_single goal
|> snd |> #fmla
|> (fn fmla => Goal.prove ctxt [] [] fmla (fn _ => tactic prob_name))
end
val ignore_consts =
[HOLogic.conj, HOLogic.disj, HOLogic.imp, HOLogic.Not]
fun new_consts_between t1 t2 =
filter
(fn n => not (exists (fn n' => n' = n) ignore_consts))
(list_diff (consts_in t2) (consts_in t1))
fun mk_bind_eq prob_name params ((n, ty), t) =
let
val bnd =
Binding.name (Long_Name.base_name n ^ "_def")
|> Binding.qualify false (TPTP_Problem_Name.mangle_problem_name prob_name)
val t' =
Term.list_comb (Const (n, ty), params)
|> rpair t
|> HOLogic.mk_eq
|> HOLogic.mk_Trueprop
|> fold Logic.all params
in
(bnd, t')
end
fun mk_bind_ax prob_name node t =
let
val bnd =
Binding.name node
|> Binding.qualify false (TPTP_Problem_Name.mangle_problem_name prob_name)
in
(bnd, t)
end
fun get_defn_components
(Const (\<^const_name>‹HOL.Trueprop›, _) $
(Const (\<^const_name>‹HOL.eq›, _) $
Const (name, ty) $ t)) = ((name, ty), t)
fun transf_pannot f (pannot : proof_annotation) : (theory * proof_annotation) =
let
val (thy', fms') = f (#meta pannot)
in
(thy',
{problem_name = #problem_name pannot,
skolem_defs = #skolem_defs pannot,
defs = #defs pannot,
axs = #axs pannot,
meta = fms'})
end
fun interpolate_binds prob_name thy fms : theory * formula_meaning list =
let
fun factor_out_bind target_node pinfo intermediate_thy =
case pinfo of
TPTP_Proof.ParentWithDetails (n, pdetails) =>
let
val (new_node_name, thy') = get_next_name intermediate_thy
val orig_fmla = node_info fms #fmla n
val target_fmla = node_info fms #fmla target_node
val new_node =
(new_node_name,
{role = TPTP_Syntax.Role_Plain,
fmla = apply_binding thy' prob_name orig_fmla target_fmla pdetails |> fst,
source_inf_opt =
SOME (TPTP_Proof.Inference (bindK, [], [pinfo]))})
in
((TPTP_Proof.Parent new_node_name, SOME new_node), thy')
end
| _ => ((pinfo, NONE), intermediate_thy)
fun process_nodes (step as (n, data)) (intermediate_thy, rest) =
case #source_inf_opt data of
SOME (TPTP_Proof.Inference (inf_name, sinfos, pinfos)) =>
let
val ((pinfos', parent_nodes), thy') =
fold_map (factor_out_bind n) pinfos intermediate_thy
|> apfst ListPair.unzip
val step' =
(n, {role = #role data, fmla = #fmla data,
source_inf_opt = SOME (TPTP_Proof.Inference (inf_name, sinfos, pinfos'))})
in (thy', fold_options parent_nodes @ step' :: rest) end
| _ => (intermediate_thy, step :: rest)
in
fold process_nodes fms (thy, [])
|> apsnd rev
end
exception PREPROCESS_SPLITS
fun preprocess_splits prob_name thy fms : theory * formula_meaning list =
let
fun transform_fmla i fmla_t =
case fmla_t of
Const (\<^const_name>‹HOL.All›, ty) $ Abs (s, ty', t') =>
let
val (i', fmla_ts) = transform_fmla i t'
in
if i' > i then
(i' + 1,
map (fn t =>
Const (\<^const_name>‹HOL.All›, ty) $ Abs (s, ty', t))
fmla_ts)
else (i, [fmla_t])
end
| Const (\<^const_name>‹HOL.Ex›, ty) $ Abs (s, ty', t') =>
if loose_bvar (t', 0) then
(i, [fmla_t])
else transform_fmla (i + 1) t'
| \<^term>‹HOL.Not› $ (\<^term>‹HOL.Not› $ t') =>
transform_fmla (i + 1) t'
| \<^term>‹HOL.conj› $ t1 $ t2 =>
let
val (i1, fmla_t1s) = transform_fmla (i + 1) t1
val (i2, fmla_t2s) = transform_fmla (i + 1) t2
in
(i1 + i2 - i, fmla_t1s @ fmla_t2s)
end
| Const (\<^const_name>‹HOL.eq›, ty) $ t1 $ t2 =>
let
val (T1, (T2, res)) =
dest_funT ty
|> apsnd dest_funT
in
if T1 = HOLogic.boolT andalso T2 = HOLogic.boolT andalso
res = HOLogic.boolT then
(i + 1,
[HOLogic.mk_imp (t1, t2),
HOLogic.mk_imp (t2, t1)])
else (i, [fmla_t])
end
| _ => (i, [fmla_t])
fun preprocess_split thy split_node_name fmla_t =
let
val (node_name, thy') = get_next_name thy
val (changes, fmla_conjs) =
transform_fmla 0 fmla_t
|> apsnd rev
val target_fmla =
fold (curry HOLogic.mk_conj) (tl fmla_conjs) (hd fmla_conjs)
val new_node =
(node_name,
{role = TPTP_Syntax.Role_Plain,
fmla =
HOLogic.mk_eq (target_fmla, \<^term>‹False›)
|> HOLogic.mk_Trueprop,
source_inf_opt =
SOME (TPTP_Proof.Inference (split_preprocessingK, [], [TPTP_Proof.Parent split_node_name]))})
in
if changes = 0 then NONE
else SOME (TPTP_Proof.Parent node_name, new_node, thy')
end
in
fold
(fn step as (n, data) => fn (intermediate_thy, redirections, rest) =>
case #source_inf_opt data of
SOME (TPTP_Proof.Inference
(inf_name, sinfos, pinfos)) =>
if inf_name <> "split_conjecture" then
(intermediate_thy, redirections, step :: rest)
else
let
val split_node_name =
case pinfos of
[TPTP_Proof.Parent n] => n
| _ => raise PREPROCESS_SPLITS
in
case AList.lookup (op =) redirections split_node_name of
SOME preprocessed_split_node_name =>
let
val step' =
apply_to_parent_info (fn _ => [TPTP_Proof.Parent preprocessed_split_node_name]) step
in (intermediate_thy, redirections, step' :: rest) end
| NONE =>
let
val split_fmla =
try_dest_Trueprop (node_info fms #fmla split_node_name)
|> remove_polarity true
|> fst
val preprocess_result =
preprocess_split intermediate_thy
split_node_name
split_fmla
in
if is_none preprocess_result then
(intermediate_thy,
(split_node_name, split_node_name) :: redirections,
step :: rest)
else
let
val (new_parent_info, new_parent_node, thy') = the preprocess_result
val step' =
(n, {role = #role data, fmla = #fmla data,
source_inf_opt = SOME (TPTP_Proof.Inference (inf_name, sinfos, [new_parent_info]))})
in
(thy',
(split_node_name, fst new_parent_node) :: redirections,
step' :: new_parent_node :: rest)
end
end
end
| _ => (intermediate_thy, redirections, step :: rest))
(rev fms)
(thy, [], [])
|> (fn (x, _, z) => (x, z))
end
exception DROP_REPEATED_QUANTIFICATION
fun drop_repeated_quantification thy (fms : formula_meaning list) : theory * formula_meaning list =
let
fun remove_repeated_quantification seen t =
case t of
Const (\<^const_name>‹HOL.All›, ty) $ Abs (s, ty', t') =>
let
val (seen_so_far, seen') =
case AList.lookup (op =) seen s of
NONE => (0, (s, 0) :: seen)
| SOME n => (n + 1, AList.update (op =) (s, n + 1) seen)
val (pre_final_t, final_seen) = remove_repeated_quantification seen' t'
val final_t =
case AList.lookup (op =) final_seen s of
NONE => raise DROP_REPEATED_QUANTIFICATION
| SOME n =>
if n > seen_so_far then pre_final_t
else Const (\<^const_name>‹HOL.All›, ty) $ Abs (s, ty', pre_final_t)
in (final_t, final_seen) end
| _ => (t, seen)
fun remove_repeated_quantification' (n, {role, fmla, source_inf_opt}) =
(n,
{role = role,
fmla =
try_dest_Trueprop fmla
|> remove_repeated_quantification []
|> fst
|> HOLogic.mk_Trueprop,
source_inf_opt = source_inf_opt})
in
(thy, map remove_repeated_quantification' fms)
end
fun node_is_inference fms rule_name node_name =
case node_info fms #source_inf_opt node_name of
NONE => false
| SOME (TPTP_Proof.File _) => false
| SOME (TPTP_Proof.Inference (rule_name', _, _)) => rule_name' = rule_name
datatype branch_info =
Split_free
| Split_present
| Coinconsistent of int
| No_info
type path = (branch_info * string list)
exception PRUNE_REDUNDANT_SPLITS
fun prune_redundant_splits prob_name thy fms : theory * formula_meaning list =
let
val initial_path = (No_info, [proof_end_node fms])
val end_node = proof_beginning_node fms
fun compute_path (path as ((info,
(n :: ns)) : path))
intermediate_thy =
case info of
Split_free => (([path], []), intermediate_thy)
| Coinconsistent branch_id =>
let
val parent_nodes = parents_of_node fms n
in
if exists (node_is_inference fms "split_conjecture") parent_nodes then
(([], [branch_id]), intermediate_thy)
else
list_prod [] parent_nodes (n :: ns)
|> map (fn ns' => (Coinconsistent branch_id, ns'))
|> (fn x => ((x, []), intermediate_thy))
end
| No_info =>
let
val parent_nodes = parents_of_node fms n
val (thy', new_branch_info) =
if node_is_inference fms "fo_atp_e" n orelse
node_is_inference fms "res" n then
let
val (i', intermediate_thy') = get_next_int intermediate_thy
in
(intermediate_thy', SOME (Coinconsistent i'))
end
else (intermediate_thy, NONE)
in
if exists (node_is_inference fms "split_conjecture") parent_nodes then
(([], []), thy')
else
list_prod [] parent_nodes (n :: ns)
|> map (fn ns' =>
let
val info =
if is_some new_branch_info then the new_branch_info
else
if hd ns' = end_node then Split_free else No_info
in (info, ns') end)
|> (fn x => ((x, []), thy'))
end
| _ => raise PRUNE_REDUNDANT_SPLITS
fun compute_paths intermediate_thy (paths : path list) =
if filter (fn (_, ns) => ns <> [] andalso hd ns = end_node) paths = paths then
(intermediate_thy, paths)
else
let
val filtered_paths = filter (fn (info, _) : path => info <> Split_present) paths
val (paths', thy') =
fold_map compute_path filtered_paths intermediate_thy
in
paths'
|> ListPair.unzip
|> (fn (paths, branch_ids) =>
(flat paths,
fold (Library.insert (op =)) (flat branch_ids) []))
|> (fn (paths, branch_ids) =>
filter (fn (info, _) =>
case info of
Coinconsistent branch_id => exists (fn x => x = branch_id) branch_ids
| _ => true) paths)
|> compute_paths thy'
end
val (thy', paths) =
compute_paths thy [initial_path]
|> apsnd
(filter (fn (branch_info, _) =>
case branch_info of
Split_free => true
| Coinconsistent _ => true
| _ => false))
fun path_to_fms ((_, nodes) : path) =
fold
(fn n => fn fms' =>
case AList.lookup (op =) fms' n of
SOME _ => fms'
| NONE =>
let
val node_info = the (AList.lookup (op =) fms n)
val source_info' =
case #source_inf_opt node_info of
NONE => error "Only the conjecture is an orphan"
| SOME (source_info as TPTP_Proof.File _) => source_info
| SOME (source_info as
TPTP_Proof.Inference (inference_name,
useful_infos : TPTP_Proof.useful_info_as list,
parent_infos)) =>
let
fun is_node_in_fms' parent_info =
let
val parent_nodename =
case parent_info of
TPTP_Proof.Parent n => n
| TPTP_Proof.ParentWithDetails (n, _) => n
in
case AList.lookup (op =) fms' parent_nodename of
NONE => false
| SOME _ => true
end
in
TPTP_Proof.Inference (inference_name,
useful_infos,
filter is_node_in_fms' parent_infos)
end
in
(n,
{role = #role node_info,
fmla = #fmla node_info,
source_inf_opt = SOME source_info'}) :: fms'
end)
nodes
[]
in
if null paths then (thy', fms) else
(thy',
hd paths
|> path_to_fms)
end
fun import_thm cautious path_prefixes file_name
(on_load : proof_annotation -> theory -> (proof_annotation * theory)) thy =
let
val prob_name =
Path.file_name file_name
|> TPTP_Problem_Name.parse_problem_name
val thy1 = TPTP_Interpret.import_file cautious path_prefixes file_name [] [] thy
val fms = get_fmlas_of_prob thy1 prob_name
in
if List.null fms then
(warning ("File " ^ Path.print file_name ^ " appears empty!");
TPTP_Reconstruction_Data.map (cons ((prob_name, empty_pannot prob_name))) thy1)
else
let
val defn_equations =
filter (fn (_, role, _, _) => role = TPTP_Syntax.Role_Definition) fms
|> map (fn (node, _, t, _) =>
(node,
get_defn_components t
|> mk_bind_eq prob_name []))
val axioms =
filter (fn (_, role, _, _) => role = TPTP_Syntax.Role_Axiom) fms
|> map (fn (node, _, t, _) =>
(node,
mk_bind_ax prob_name node t))
val thy2 =
fold
(fn bnd => fn thy =>
let
val ((name, thm), thy') = Thm.add_axiom_global bnd thy
in Global_Theory.add_thm ((#1 bnd, thm), []) thy' |> #2 end)
(map snd defn_equations @ map snd axioms)
thy1
val (thy3, pre_pannot) : theory * proof_annotation =
transf_pannot
(prune_redundant_splits prob_name thy2
#-> interpolate_binds prob_name
#-> preprocess_splits prob_name
#-> drop_repeated_quantification)
{problem_name = prob_name,
skolem_defs = [],
defs = map (apsnd fst) defn_equations,
axs = map (apsnd fst) axioms,
meta = map (fn (n, r, t, info) => (n, {role=r, fmla=t, source_inf_opt=info})) fms}
val thy4 = TPTP_Reconstruction_Data.map (cons ((prob_name, pre_pannot))) thy3
val (pannot, thy5) = on_load pre_pannot thy4
in TPTP_Reconstruction_Data.map (cons ((prob_name, pannot))) thy5 end
end
fun match_rules_of_current (pannot : proof_annotation) rules n =
case node_info (#meta pannot) #source_inf_opt n of
NONE => NONE
| SOME (TPTP_Proof.File _) => NONE
| SOME (TPTP_Proof.Inference (rule_name, _ : TPTP_Proof.useful_info_as list, _)) =>
if member (op =) rules rule_name then SOME rule_name else NONE
fun match_rules_of_immediate_previous (pannot : proof_annotation) strict rules n =
case node_info (#meta pannot) #source_inf_opt n of
NONE => null rules
| SOME (TPTP_Proof.File _) => null rules
| SOME (TPTP_Proof.Inference (rule_name, _ : TPTP_Proof.useful_info_as list, parent_infos)) =>
let
val matched_rules : string option list =
map (match_rules_of_current pannot rules)
(dest_parent_infos true (#meta pannot) parent_infos |> map #name)
in
if strict andalso member (op =) matched_rules NONE then false
else
fold
(fn (rule : string) => fn (st, matches : string option list) =>
if not st then (st, matches)
else
let
val idx = find_index (fn match => SOME rule = match) matches
in
if idx < 0 then (false, matches)
else
(st, nth_drop idx matches)
end)
rules
(true, matched_rules)
|> #1
end
end