src/HOL/TPTP/TPTP_Parser_Example.thy
author sultana
Wed, 04 Apr 2012 21:57:39 +0100
changeset 47366 f5a304ca037e
child 47518 b2f209258621
permissions -rw-r--r--
improved import_tptp to use standard TPTP directory structure; extended the TPTP testing theory to include an example of using the import_tptp command;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
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
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
     8
imports TPTP_Parser
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
uses "~~/src/HOL/ex/sledgehammer_tactics.ML"
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
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
    11
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
import_tptp "$TPTP_PROBLEMS_PATH/ALG/ALG001^5.p" (*FIXME multiple imports*)
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
    13
(*import_tptp "$TPTP_PROBLEMS_PATH/NUM/NUM021^1.p"*)
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
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
ML {*
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
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
    17
  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
    18
  |> 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
    19
  |> #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
    20
  |> #3 (*get formulas*)
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
*}
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
    22
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
    23
(*Display nicely.*)
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
ML {*
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
List.app (fn (n, role, _, 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
    26
  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
    27
    (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
    28
      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
    29
  ) (rev 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
    30
*}
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
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
ML {*
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
(*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
    34
 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
    35
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
    36
 (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
    37
   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
    38
     fun role_predicate (_, role, _, _, _) =
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
    39
       fold (fn r1 => fn b => role = r1 orelse b) roles false
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
   in filter role_predicate #> map (fn (n, _, _, t, _) => (n, t)) 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
    41
*}
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
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
ML {*
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
(*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
    45
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
    46
  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
    47
    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
    48
      extract_terms
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
       [TPTP_Syntax.Role_Definition (*FIXME include axioms, etc here*)]
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
       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
      |> 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
    52
    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
    53
    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
    54
      (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
    55
  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
    56
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
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
    58
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
    59
  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
    60
  (*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
    61
  {add = [], del = [], only = false} 1)
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
    62
*}
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
    63
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
    64
ML "auto_prove @{context} 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
    65
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
    66
sledgehammer_params [provers = z3_tptp leo2, debug]
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
    67
ML "sh_prove @{context} 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
    68
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
    69
end