CookBook/Recipes/ExternalSolver.thy
author Christian Urban <urbanc@in.tum.de>
Tue, 27 Jan 2009 06:15:13 +0000
changeset 83 0fb5f91d5109
parent 81 8fda6b452f28
child 94 531e453c9d67
permissions -rw-r--r--
fixed typos
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
61
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
     1
theory ExternalSolver
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
     2
imports "../Base"
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
     3
begin
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
     4
81
8fda6b452f28 polished
Christian Urban <urbanc@in.tum.de>
parents: 79
diff changeset
     5
section {* Using an External Solver\label{rec:external} *} 
61
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
     6
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
     7
text {*
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
     8
  {\bf Problem:}
79
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
     9
  You want to use an external solver, because the solver might be more efficient 
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    10
  for deciding a certain class of formulae than Isabelle tactics.
61
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    11
  \smallskip
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    12
79
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    13
  {\bf Solution:} The easiest way to do this is by implementing an oracle.
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    14
  We will also construct proofs inside Isabelle by using the results produced 
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    15
  by the oracle.
61
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    16
  \smallskip
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    17
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    18
  \begin{readmore}
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    19
  A short introduction to oracles can be found in [isar-ref: no suitable label
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    20
  for section 3.11]. A simple example is given in 
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    21
  @{ML_file "FOL/ex/IffOracle"}.
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    22
  (TODO: add more references to the code)
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    23
  \end{readmore}
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    24
79
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    25
  For our explanation here, we will use the @{text metis} prover for proving
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    26
  propositional formulae. The general method will be roughly as follows: 
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    27
  Given a goal @{text G}, we transform it into the syntactical respresentation 
83
0fb5f91d5109 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 81
diff changeset
    28
  of @{text "metis"}, build the CNF of the negated formula and then let metis search 
0fb5f91d5109 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 81
diff changeset
    29
  for a refutation. @{text Metis} will either return the proved goal or raise 
0fb5f91d5109 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 81
diff changeset
    30
  an exception meaning 
79
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    31
  that it was unable to prove the goal (FIXME: is this so?).
61
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    32
79
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    33
  The translation function from Isabelle propositions into formulae of 
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    34
  @{text metis} is as follows:
61
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    35
  *}
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    36
79
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    37
ML{*fun trans t =
61
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    38
  (case t of
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    39
    @{term Trueprop} $ t => trans t
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    40
  | @{term True} => Metis.Formula.True
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    41
  | @{term False} => Metis.Formula.False
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    42
  | @{term Not} $ t => Metis.Formula.Not (trans t)
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    43
  | @{term "op &"} $ t1 $ t2 => Metis.Formula.And (trans t1, trans t2)
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    44
  | @{term "op |"} $ t1 $ t2 => Metis.Formula.Or (trans t1, trans t2)
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    45
  | @{term "op -->"} $ t1 $ t2 => Metis.Formula.Imp (trans t1, trans t2)
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    46
  | @{term "op = :: bool => bool => bool"} $ t1 $ t2 => 
79
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    47
                                     Metis.Formula.Iff (trans t1, trans t2)
61
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    48
  | Free (n, @{typ bool}) => Metis.Formula.Atom (n, [])
79
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    49
  | _ => error "inacceptable term")*}
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    50
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    51
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    52
text {* 
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    53
  An example is as follows:
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    54
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    55
@{ML_response [display,gray] "trans @{prop \"A \<and> B\"}" 
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    56
"Metis.Formula.And 
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    57
        (Metis.Formula.Atom (\"A\", []), Metis.Formula.Atom (\"B\", []))"}
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    58
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    59
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    60
  The next function computes the conjunctive-normal-form.
61
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    61
*}
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    62
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    63
79
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    64
ML %linenumbers{*fun make_cnfs fm =
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    65
       fm |> Metis.Formula.Not
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    66
          |> Metis.Normalize.cnf
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    67
          |> map Metis.Formula.stripConj
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    68
          |> map (map Metis.Formula.stripDisj)
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    69
          |> map (map (map Metis.Literal.fromFormula))
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    70
          |> map (map Metis.LiteralSet.fromList)
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    71
          |> map (map Metis.Thm.axiom)*}
61
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    72
79
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    73
text {* 
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    74
  (FIXME: Is there a deep reason why Metis.Normalize.cnf returns a list?)
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    75
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    76
  (FIXME: What does Line 8 do?)
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    77
83
0fb5f91d5109 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 81
diff changeset
    78
  (FIXME: Can this code be improved?)
79
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    79
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    80
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    81
  Setting up the resolution.
61
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    82
*}
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    83
79
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    84
ML{*fun refute cls =
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    85
 let val result = 
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    86
        Metis.Resolution.loop             
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    87
          (Metis.Resolution.new Metis.Resolution.default cls)
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    88
 in
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    89
   (case result of
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    90
      Metis.Resolution.Contradiction _ => true
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    91
    | Metis.Resolution.Satisfiable _ => false)
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    92
 end*}
61
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    93
79
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    94
text {* 
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    95
  Stringing the functions together.
61
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    96
*}
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
    97
79
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    98
ML{*fun solve f = List.all refute (make_cnfs f)*}
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
    99
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
   100
text {* Setting up the oracle*}
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
   101
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
   102
ML{*fun prop_dp (thy, t) = 
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
   103
        if solve (trans t) then (Thm.cterm_of thy t) 
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
   104
        else error "Proof failed."*}
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
   105
61
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
   106
oracle prop_oracle = prop_dp
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
   107
83
0fb5f91d5109 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 81
diff changeset
   108
text {* (FIXME: What does oracle do?) *}
79
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
   109
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
   110
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
   111
ML{*fun prop_oracle_tac ctxt = 
61
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
   112
  SUBGOAL (fn (goal, i) => 
79
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
   113
    (case (try prop_oracle (ProofContext.theory_of ctxt, goal)) of
61
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
   114
      SOME thm => rtac thm i
79
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
   115
    | NONE => no_tac))*}
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
   116
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
   117
text {*
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
   118
  (FIXME: The oracle returns a @{text cterm}. How is it possible
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
   119
  that I can apply this cterm with @{ML rtac}?)
61
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
   120
*}
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
   121
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
   122
method_setup prop_oracle = {*
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
   123
  Method.ctxt_args (fn ctxt => Method.SIMPLE_METHOD' (prop_oracle_tac ctxt))
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
   124
*} "Oracle-based decision procedure for propositional logic"
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
   125
79
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
   126
text {* (FIXME What does @{ML Method.SIMPLE_METHOD'} do?) *}
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
   127
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
   128
lemma test: "p \<or> \<not>p"
61
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
   129
  by prop_oracle
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
   130
83
0fb5f91d5109 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 81
diff changeset
   131
text {* (FIXME: say something about what the proof of the oracle is)
0fb5f91d5109 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 81
diff changeset
   132
0fb5f91d5109 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 81
diff changeset
   133
  @{ML_response_fake [display,gray] "Thm.proof_of @{thm test}" "???"}
79
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
   134
83
0fb5f91d5109 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 81
diff changeset
   135
*} 
0fb5f91d5109 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 81
diff changeset
   136
0fb5f91d5109 fixed typos
Christian Urban <urbanc@in.tum.de>
parents: 81
diff changeset
   137
79
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
   138
61
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
   139
lemma "((p \<longrightarrow> q) \<longrightarrow> p) \<longrightarrow> p"
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
   140
  by prop_oracle
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
   141
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
   142
lemma "\<forall>x::nat. x \<ge> 0"
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
   143
  sorry
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
   144
79
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
   145
text {*
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
   146
  (FIXME: proof reconstruction)
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
   147
*}
61
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
   148
79
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
   149
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
   150
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
   151
text {*
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
   152
  For communication with external programs, there are the primitives
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
   153
  @{ML_text system} and @{ML_text system_out}, the latter of which captures
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
   154
  the invoked program's output. For simplicity, here, we will use metis, an
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
   155
  external solver included in the Isabelle destribution. Since it is written
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
   156
  in ML, we can call it directly without the detour of invoking an external
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
   157
  program.
a53c7810e38b some tuning
Christian Urban <urbanc@in.tum.de>
parents: 61
diff changeset
   158
*}
61
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
   159
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
   160
64c9540f2f84 Added four recipes.
boehmes
parents:
diff changeset
   161
end