| author | wenzelm | 
| Tue, 18 Apr 2023 12:23:37 +0200 | |
| changeset 77869 | 1156aa9db7f5 | 
| parent 73684 | a63d76ba0a03 | 
| permissions | -rw-r--r-- | 
| 47366 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 1 | (* Title: HOL/TPTP/TPTP_Parser_Example.thy | 
| 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 2 | Author: Nik Sultana, Cambridge University Computer Laboratory | 
| 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 3 | |
| 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 4 | Example of importing a TPTP problem and trying to prove it in Isabelle/HOL. | 
| 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 5 | *) | 
| 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 6 | |
| 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 7 | theory TPTP_Parser_Example | 
| 47518 | 8 | imports TPTP_Parser TPTP_Interpret | 
| 47366 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 9 | begin | 
| 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 10 | |
| 53393 
5278312037f8
included axiom formulas (removing a FIXME) in the imported problem;
 sultana parents: 
48891diff
changeset | 11 | import_tptp "$TPTP/Problems/LCL/LCL414+1.p" | 
| 47366 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 12 | |
| 63167 | 13 | ML \<open> | 
| 47366 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 14 | val an_fmlas = | 
| 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 15 |   TPTP_Interpret.get_manifests @{theory}
 | 
| 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 16 | |> hd (*FIXME use named lookup*) | 
| 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 17 | |> #2 (*get problem contents*) | 
| 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 18 | |> #3 (*get formulas*) | 
| 63167 | 19 | \<close> | 
| 47366 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 20 | |
| 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 21 | (*Display nicely.*) | 
| 63167 | 22 | ML \<open> | 
| 47548 | 23 | List.app (fn (n, role, fmla, _) => | 
| 47366 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 24 | Pretty.writeln | 
| 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 25 |     (Pretty.block [Pretty.str ("\"" ^ n ^ "\"" ^ "(" ^
 | 
| 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 26 |       TPTP_Syntax.role_to_string role  ^ "): "), Syntax.pretty_term @{context} fmla])
 | 
| 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 27 | ) (rev an_fmlas) | 
| 63167 | 28 | \<close> | 
| 47366 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 29 | |
| 63167 | 30 | ML \<open> | 
| 47366 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 31 | (*Extract the (name, term) pairs of formulas having roles belonging to a | 
| 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 32 | user-supplied set*) | 
| 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 33 | fun extract_terms roles : TPTP_Interpret.tptp_formula_meaning list -> | 
| 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 34 | (string * term) list = | 
| 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 35 | let | 
| 47548 | 36 | fun role_predicate (_, role, _, _) = | 
| 47366 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 37 | fold (fn r1 => fn b => role = r1 orelse b) roles false | 
| 47548 | 38 | in filter role_predicate #> map (fn (n, _, t, _) => (n, t)) end | 
| 63167 | 39 | \<close> | 
| 47366 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 40 | |
| 63167 | 41 | ML \<open> | 
| 47366 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 42 | (*Use a given tactic on a goal*) | 
| 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 43 | fun prove_conjectures tactic ctxt an_fmlas = | 
| 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 44 | let | 
| 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 45 | val assumptions = | 
| 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 46 | extract_terms | 
| 53393 
5278312037f8
included axiom formulas (removing a FIXME) in the imported problem;
 sultana parents: 
48891diff
changeset | 47 | [TPTP_Syntax.Role_Definition, TPTP_Syntax.Role_Axiom] | 
| 47366 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 48 | an_fmlas | 
| 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 49 | |> map snd | 
| 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 50 | val goals = extract_terms [TPTP_Syntax.Role_Conjecture] an_fmlas | 
| 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 51 | fun driver (n, goal) = | 
| 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 52 | (n, Goal.prove ctxt [] assumptions goal (fn _ => tactic ctxt)) | 
| 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 53 | in map driver goals end | 
| 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 54 | |
| 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 55 | val auto_prove = prove_conjectures auto_tac | 
| 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 56 | val sh_prove = prove_conjectures (fn ctxt => | 
| 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 57 | Sledgehammer_Tactics.sledgehammer_with_metis_tac ctxt [] | 
| 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 58 | (*FIXME use relevance_override*) | 
| 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 59 |   {add = [], del = [], only = false} 1)
 | 
| 63167 | 60 | \<close> | 
| 47366 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 61 | |
| 63167 | 62 | ML \<open> | 
| 53393 
5278312037f8
included axiom formulas (removing a FIXME) in the imported problem;
 sultana parents: 
48891diff
changeset | 63 | @{assert} (is_some (try (auto_prove @{context}) an_fmlas) = false)
 | 
| 63167 | 64 | \<close> | 
| 47366 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 65 | |
| 55049 | 66 | sledgehammer_params [provers = z3, debug] | 
| 63167 | 67 | ML \<open> | 
| 53393 
5278312037f8
included axiom formulas (removing a FIXME) in the imported problem;
 sultana parents: 
48891diff
changeset | 68 | @{assert} (is_some (try (sh_prove @{context}) an_fmlas) = true)
 | 
| 63167 | 69 | \<close> | 
| 47366 
f5a304ca037e
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
 sultana parents: diff
changeset | 70 | |
| 55049 | 71 | end |