src/HOL/Tools/Argo/argo_tactic.ML
author haftmann
Thu, 19 Jun 2025 17:15:40 +0200
changeset 82734 89347c0cc6a3
parent 81942 da3c3948a39c
permissions -rw-r--r--
treat map_filter similar to list_all, list_ex, list_ex1
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
63992
3aa9837d05c7 updated headers;
wenzelm
parents: 63960
diff changeset
     1
(*  Title:      HOL/Tools/Argo/argo_tactic.ML
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
     2
    Author:     Sascha Boehme
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
     3
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
     4
HOL method and tactic for the Argo solver.
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
     5
*)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
     6
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
     7
signature ARGO_TACTIC =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
     8
sig
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
     9
  val trace: string Config.T
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    10
  val timeout: real Config.T
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    11
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    12
  (* extending the tactic *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    13
  type trans_context =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    14
    Name.context * Argo_Expr.typ Typtab.table * (string * Argo_Expr.typ) Termtab.table
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    15
  type ('a, 'b) trans = 'a -> trans_context -> 'b * trans_context
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    16
  type ('a, 'b) trans' = 'a -> trans_context -> ('b * trans_context) option
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    17
  type extension = {
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    18
    trans_type: (typ, Argo_Expr.typ) trans -> (typ, Argo_Expr.typ) trans',
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    19
    trans_term: (term, Argo_Expr.expr) trans -> (term, Argo_Expr.expr) trans',
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    20
    term_of: (Argo_Expr.expr -> term) -> Argo_Expr.expr -> term option,
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    21
    replay_rewr: Proof.context -> Argo_Proof.rewrite -> conv,
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    22
    replay: (Argo_Expr.expr -> cterm) -> Proof.context -> Argo_Proof.rule -> thm list -> thm option}
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    23
  val add_extension: extension -> theory -> theory
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    24
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    25
  (* proof utilities *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    26
  val discharges: thm -> thm list -> thm list
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    27
  val flatten_conv: conv -> thm -> conv
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    28
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    29
  (* interface to the tactic as well as the underlying checker and prover *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    30
  datatype result = Satisfiable of term -> bool option | Unsatisfiable
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    31
  val check: term list -> Proof.context -> result * Proof.context
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    32
  val prove: thm list -> Proof.context -> thm option * Proof.context
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    33
  val argo_tac: Proof.context -> thm list -> int -> tactic
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    34
end
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    35
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    36
structure Argo_Tactic: ARGO_TACTIC =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    37
struct
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    38
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    39
(* readable fresh names for terms *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    40
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    41
fun fresh_name n = Name.variant (case Long_Name.base_name n of "" => "x" | n' => n')
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    42
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    43
fun fresh_type_name (Type (n, _)) = fresh_name n
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    44
  | fresh_type_name (TFree (n, _)) = fresh_name n
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    45
  | fresh_type_name (TVar ((n, i), _)) = fresh_name (n ^ "." ^ string_of_int i)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    46
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    47
fun fresh_term_name (Const (n, _)) = fresh_name n
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    48
  | fresh_term_name (Free (n, _)) = fresh_name n
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    49
  | fresh_term_name (Var ((n, i), _)) = fresh_name (n ^ "." ^ string_of_int i)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    50
  | fresh_term_name _ = fresh_name ""
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    51
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    52
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    53
(* tracing *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    54
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    55
datatype mode = None | Basic | Full
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    56
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    57
fun string_of_mode None = "none"
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    58
  | string_of_mode Basic = "basic"
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    59
  | string_of_mode Full = "full"
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    60
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    61
fun requires_mode None = []
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    62
  | requires_mode Basic = [Basic, Full]
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    63
  | requires_mode Full = [Full]
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    64
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
    65
val trace = Attrib.setup_config_string \<^binding>\<open>argo_trace\<close> (K (string_of_mode None))
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    66
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    67
fun allows_mode ctxt = exists (equal (Config.get ctxt trace) o string_of_mode) o requires_mode
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    68
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    69
fun output mode ctxt msg = if allows_mode ctxt mode then Output.tracing ("Argo: " ^ msg) else ()
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    70
val tracing = output Basic
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    71
val full_tracing = output Full
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    72
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    73
fun with_mode mode ctxt f = if allows_mode ctxt mode then f ctxt else ()
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    74
val with_tracing = with_mode Basic
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    75
val with_full_tracing = with_mode Full
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    76
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    77
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    78
(* timeout *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    79
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
    80
val timeout = Attrib.setup_config_real \<^binding>\<open>argo_timeout\<close> (K 10.0)
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    81
69578
wenzelm
parents: 69576
diff changeset
    82
val timeout_seconds = seconds o Config.apply timeout
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    83
74646
wenzelm
parents: 74561
diff changeset
    84
fun with_timeout f ctxt = Timeout.apply (timeout_seconds ctxt) f ctxt
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    85
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    86
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    87
(* extending the tactic *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    88
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    89
type trans_context =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    90
  Name.context * Argo_Expr.typ Typtab.table * (string * Argo_Expr.typ) Termtab.table
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    91
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    92
type ('a, 'b) trans = 'a -> trans_context -> 'b * trans_context
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    93
type ('a, 'b) trans' = 'a -> trans_context -> ('b * trans_context) option
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    94
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    95
type extension = {
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    96
  trans_type: (typ, Argo_Expr.typ) trans -> (typ, Argo_Expr.typ) trans',
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    97
  trans_term: (term, Argo_Expr.expr) trans -> (term, Argo_Expr.expr) trans',
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    98
  term_of: (Argo_Expr.expr -> term) -> Argo_Expr.expr -> term option,
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
    99
  replay_rewr: Proof.context -> Argo_Proof.rewrite -> conv,
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   100
  replay: (Argo_Expr.expr -> cterm) -> Proof.context -> Argo_Proof.rule -> thm list -> thm option}
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   101
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   102
fun eq_extension ((serial1, _), (serial2, _)) = (serial1 = serial2)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   103
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   104
structure Extensions = Theory_Data
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   105
(
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   106
  type T = (serial * extension) list
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   107
  val empty = []
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   108
  val merge = merge eq_extension 
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   109
)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   110
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   111
fun add_extension ext = Extensions.map (insert eq_extension (serial (), ext))
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   112
fun get_extensions ctxt = Extensions.get (Proof_Context.theory_of ctxt)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   113
fun apply_first ctxt f = get_first (fn (_, e) => f e) (get_extensions ctxt)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   114
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   115
fun ext_trans sel ctxt f x tcx = apply_first ctxt (fn ext => sel ext f x tcx)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   116
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   117
val ext_trans_type = ext_trans (fn {trans_type, ...}: extension => trans_type)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   118
val ext_trans_term = ext_trans (fn {trans_term, ...}: extension => trans_term)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   119
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   120
fun ext_term_of ctxt f e = apply_first ctxt (fn {term_of, ...}: extension => term_of f e)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   121
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   122
fun ext_replay_rewr ctxt r = 
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   123
  get_extensions ctxt
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   124
  |> map (fn (_, {replay_rewr, ...}: extension) => replay_rewr ctxt r)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   125
  |> Conv.first_conv
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   126
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   127
fun ext_replay cprop_of ctxt rule prems =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   128
  (case apply_first ctxt (fn {replay, ...}: extension => replay cprop_of ctxt rule prems) of
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   129
    SOME thm => thm
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   130
  | NONE => raise THM ("failed to replay " ^ quote (Argo_Proof.string_of_rule rule), 0, []))
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   131
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   132
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   133
(* translating input terms *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   134
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   135
fun add_new_type T (names, types, terms) =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   136
  let
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   137
    val (n, names') = fresh_type_name T names
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   138
    val ty = Argo_Expr.Type n
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   139
  in (ty, (names', Typtab.update (T, ty) types, terms)) end
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   140
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   141
fun add_type T (tcx as (_, types, _)) =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   142
  (case Typtab.lookup types T of
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   143
    SOME ty => (ty, tcx)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   144
  | NONE => add_new_type T tcx)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   145
74323
5c452041fe83 clarified antiquotations;
wenzelm
parents: 74282
diff changeset
   146
fun trans_type _ \<^Type>\<open>bool\<close> = pair Argo_Expr.Bool
5c452041fe83 clarified antiquotations;
wenzelm
parents: 74282
diff changeset
   147
  | trans_type ctxt \<^Type>\<open>fun T1 T2\<close> =
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   148
      trans_type ctxt T1 ##>> trans_type ctxt T2 #>> Argo_Expr.Func
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   149
  | trans_type ctxt T = (fn tcx =>
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   150
      (case ext_trans_type ctxt (trans_type ctxt) T tcx of
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   151
        SOME result => result
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   152
      | NONE => add_type T tcx))
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   153
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   154
fun add_new_term ctxt t T tcx =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   155
  let
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   156
    val (ty, (names, types, terms)) = trans_type ctxt T tcx
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   157
    val (n, names') = fresh_term_name t names
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   158
    val c = (n, ty)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   159
  in (Argo_Expr.mk_con c, (names', types, Termtab.update (t, c) terms)) end
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   160
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   161
fun add_term ctxt t (tcx as (_, _, terms)) =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   162
  (case Termtab.lookup terms t of
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   163
    SOME c => (Argo_Expr.mk_con c, tcx)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   164
  | NONE => add_new_term ctxt t (Term.fastype_of t) tcx)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   165
74323
5c452041fe83 clarified antiquotations;
wenzelm
parents: 74282
diff changeset
   166
fun mk_eq \<^Type>\<open>bool\<close> = Argo_Expr.mk_iff
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   167
  | mk_eq _ = Argo_Expr.mk_eq
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   168
74323
5c452041fe83 clarified antiquotations;
wenzelm
parents: 74282
diff changeset
   169
fun trans_term _ \<^Const_>\<open>True\<close> = pair Argo_Expr.true_expr
5c452041fe83 clarified antiquotations;
wenzelm
parents: 74282
diff changeset
   170
  | trans_term _ \<^Const_>\<open>False\<close> = pair Argo_Expr.false_expr
5c452041fe83 clarified antiquotations;
wenzelm
parents: 74282
diff changeset
   171
  | trans_term ctxt \<^Const_>\<open>Not for t\<close> = trans_term ctxt t #>> Argo_Expr.mk_not
5c452041fe83 clarified antiquotations;
wenzelm
parents: 74282
diff changeset
   172
  | trans_term ctxt \<^Const_>\<open>conj for t1 t2\<close> =
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   173
      trans_term ctxt t1 ##>> trans_term ctxt t2 #>> uncurry Argo_Expr.mk_and2
74323
5c452041fe83 clarified antiquotations;
wenzelm
parents: 74282
diff changeset
   174
  | trans_term ctxt \<^Const_>\<open>disj for t1 t2\<close> =
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   175
      trans_term ctxt t1 ##>> trans_term ctxt t2 #>> uncurry Argo_Expr.mk_or2
74323
5c452041fe83 clarified antiquotations;
wenzelm
parents: 74282
diff changeset
   176
  | trans_term ctxt \<^Const_>\<open>implies for t1 t2\<close> =
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   177
      trans_term ctxt t1 ##>> trans_term ctxt t2 #>> uncurry Argo_Expr.mk_imp
74323
5c452041fe83 clarified antiquotations;
wenzelm
parents: 74282
diff changeset
   178
  | trans_term ctxt \<^Const_>\<open>If _ for t1 t2 t3\<close> =
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   179
      trans_term ctxt t1 ##>> trans_term ctxt t2 ##>> trans_term ctxt t3 #>>
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   180
      (fn ((u1, u2), u3) => Argo_Expr.mk_ite u1 u2 u3)
74323
5c452041fe83 clarified antiquotations;
wenzelm
parents: 74282
diff changeset
   181
  | trans_term ctxt \<^Const_>\<open>HOL.eq T for t1 t2\<close> =
5c452041fe83 clarified antiquotations;
wenzelm
parents: 74282
diff changeset
   182
      trans_term ctxt t1 ##>> trans_term ctxt t2 #>> uncurry (mk_eq T)
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   183
  | trans_term ctxt (t as (t1 $ t2)) = (fn tcx =>
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   184
      (case ext_trans_term ctxt (trans_term ctxt) t tcx of
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   185
        SOME result => result
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   186
      | NONE => tcx |> trans_term ctxt t1 ||>> trans_term ctxt t2 |>> uncurry Argo_Expr.mk_app))
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   187
  | trans_term ctxt t = (fn tcx =>
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   188
      (case ext_trans_term ctxt (trans_term ctxt) t tcx of
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   189
        SOME result => result
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   190
      | NONE => add_term ctxt t tcx))
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   191
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   192
fun translate ctxt prop = trans_term ctxt (HOLogic.dest_Trueprop prop)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   193
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   194
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   195
(* invoking the solver *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   196
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   197
type data = {
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   198
  names: Name.context,
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   199
  types: Argo_Expr.typ Typtab.table,
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   200
  terms: (string * Argo_Expr.typ) Termtab.table,
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   201
  argo: Argo_Solver.context}
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   202
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   203
fun mk_data names types terms argo: data = {names=names, types=types, terms=terms, argo=argo}
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   204
val empty_data = mk_data Name.context Typtab.empty Termtab.empty Argo_Solver.context
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   205
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   206
structure Solver_Data = Proof_Data
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   207
(
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   208
  type T = data option
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   209
  fun init _ = SOME empty_data
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   210
)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   211
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   212
datatype ('m, 'p) solver_result = Model of 'm | Proof of 'p
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   213
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   214
fun raw_solve es argo = Model (Argo_Solver.assert es argo)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   215
  handle Argo_Proof.UNSAT proof => Proof proof
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   216
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   217
fun value_of terms model t =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   218
  (case Termtab.lookup terms t of
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   219
    SOME c => model c
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   220
  | _ => NONE)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   221
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   222
fun trace_props props ctxt =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   223
  tracing ctxt (Pretty.string_of (Pretty.big_list "using these propositions:"
74648
d1117655110c tuned message;
wenzelm
parents: 74647
diff changeset
   224
    (map (Pretty.item o single o Syntax.pretty_term ctxt) props)))
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   225
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   226
fun trace_result ctxt ({elapsed, ...}: Timing.timing) msg =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   227
  tracing ctxt ("found a " ^ msg ^ " in " ^ string_of_int (Time.toMilliseconds elapsed) ^ " ms")
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   228
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   229
fun solve props ctxt =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   230
  (case Solver_Data.get ctxt of
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   231
    NONE => error "bad Argo solver context"
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   232
  | SOME {names, types, terms, argo} =>
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   233
      let
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   234
        val _ = with_tracing ctxt (trace_props props)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   235
        val (es, (names', types', terms')) = fold_map (translate ctxt) props (names, types, terms)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   236
        val _ = tracing ctxt "starting the prover"
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   237
      in
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   238
        (case Timing.timing (raw_solve es) argo of
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   239
          (time, Proof proof) =>
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   240
            let val _ = trace_result ctxt time "proof"
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   241
            in (Proof (terms', proof), Solver_Data.put NONE ctxt) end
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   242
        | (time, Model argo') =>
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   243
            let
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   244
              val _ = trace_result ctxt time "model"
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   245
              val model = value_of terms' (Argo_Solver.model_of argo')
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   246
            in (Model model, Solver_Data.put (SOME (mk_data names' types' terms' argo')) ctxt) end)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   247
      end)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   248
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   249
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   250
(* reverse translation *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   251
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   252
structure Contab = Table(type key = string * Argo_Expr.typ val ord = Argo_Expr.con_ord)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   253
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   254
fun mk_nary f ts = uncurry (fold_rev (curry f)) (split_last ts)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   255
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   256
fun mk_nary' d _ [] = d
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   257
  | mk_nary' _ f ts = mk_nary f ts
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   258
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   259
fun mk_ite t1 t2 t3 =
74323
5c452041fe83 clarified antiquotations;
wenzelm
parents: 74282
diff changeset
   260
  let val T = Term.fastype_of t2
5c452041fe83 clarified antiquotations;
wenzelm
parents: 74282
diff changeset
   261
  in \<^Const>\<open>If T for t1 t2 t3\<close> end
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   262
74323
5c452041fe83 clarified antiquotations;
wenzelm
parents: 74282
diff changeset
   263
fun term_of _ (Argo_Expr.E (Argo_Expr.True, _)) = \<^Const>\<open>True\<close>
5c452041fe83 clarified antiquotations;
wenzelm
parents: 74282
diff changeset
   264
  | term_of _ (Argo_Expr.E (Argo_Expr.False, _)) = \<^Const>\<open>False\<close>
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   265
  | term_of cx (Argo_Expr.E (Argo_Expr.Not, [e])) = HOLogic.mk_not (term_of cx e)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   266
  | term_of cx (Argo_Expr.E (Argo_Expr.And, es)) =
74323
5c452041fe83 clarified antiquotations;
wenzelm
parents: 74282
diff changeset
   267
      mk_nary' \<^Const>\<open>True\<close> HOLogic.mk_conj (map (term_of cx) es)
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   268
  | term_of cx (Argo_Expr.E (Argo_Expr.Or, es)) =
74323
5c452041fe83 clarified antiquotations;
wenzelm
parents: 74282
diff changeset
   269
      mk_nary' \<^Const>\<open>False\<close> HOLogic.mk_disj (map (term_of cx) es)
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   270
  | term_of cx (Argo_Expr.E (Argo_Expr.Imp, [e1, e2])) =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   271
      HOLogic.mk_imp (term_of cx e1, term_of cx e2)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   272
  | term_of cx (Argo_Expr.E (Argo_Expr.Iff, [e1, e2])) =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   273
      HOLogic.mk_eq (term_of cx e1, term_of cx e2)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   274
  | term_of cx (Argo_Expr.E (Argo_Expr.Ite, [e1, e2, e3])) =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   275
      mk_ite (term_of cx e1) (term_of cx e2) (term_of cx e3)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   276
  | term_of cx (Argo_Expr.E (Argo_Expr.Eq, [e1, e2])) =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   277
      HOLogic.mk_eq (term_of cx e1, term_of cx e2)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   278
  | term_of cx (Argo_Expr.E (Argo_Expr.App, [e1, e2])) =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   279
      term_of cx e1 $ term_of cx e2
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   280
  | term_of (_, cons) (Argo_Expr.E (Argo_Expr.Con (c as (n, _)), _)) =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   281
      (case Contab.lookup cons c of
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   282
        SOME t => t
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   283
      | NONE => error ("Unexpected expression named " ^ quote n))
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   284
  | term_of (cx as (ctxt, _)) e =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   285
      (case ext_term_of ctxt (term_of cx) e of
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   286
        SOME t => t
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   287
      | NONE => raise Fail "bad expression")
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   288
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   289
fun cterm_of ctxt cons e = Thm.cterm_of ctxt (term_of (ctxt, cons) e)
81942
da3c3948a39c clarified signature: more uniform cterm operations, without context;
wenzelm
parents: 80910
diff changeset
   290
fun cprop_of ctxt cons e = HOLogic.mk_judgment (cterm_of ctxt cons e)
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   291
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   292
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   293
(* generic proof tools *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   294
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   295
fun discharge thm rule = thm INCR_COMP rule
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   296
fun discharge2 thm1 thm2 rule = discharge thm2 (discharge thm1 rule)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   297
fun discharges thm rules = [thm] RL rules
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   298
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   299
fun under_assumption f ct =
81942
da3c3948a39c clarified signature: more uniform cterm operations, without context;
wenzelm
parents: 80910
diff changeset
   300
  let val cprop = HOLogic.mk_judgment ct
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   301
  in Thm.implies_intr cprop (f (Thm.assume cprop)) end
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   302
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   303
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   304
(* proof replay for tautologies *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   305
74646
wenzelm
parents: 74561
diff changeset
   306
fun prove_taut ctxt ns t =
wenzelm
parents: 74561
diff changeset
   307
  Goal.prove ctxt ns [] (HOLogic.mk_Trueprop t)
wenzelm
parents: 74561
diff changeset
   308
    (fn {context = ctxt', ...} => HEADGOAL (Classical.fast_tac ctxt'))
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   309
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   310
fun with_frees ctxt n mk =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   311
  let
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   312
    val ns = map (fn i => "P" ^ string_of_int i) (0 upto (n - 1))
74323
5c452041fe83 clarified antiquotations;
wenzelm
parents: 74282
diff changeset
   313
    val ts = map (Free o rpair \<^Type>\<open>bool\<close>) ns
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   314
    val t = mk_nary HOLogic.mk_disj (mk ts)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   315
  in prove_taut ctxt ns t end
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   316
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   317
fun taut_and1_term ts = mk_nary HOLogic.mk_conj ts :: map HOLogic.mk_not ts
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   318
fun taut_and2_term i ts = [HOLogic.mk_not (mk_nary HOLogic.mk_conj ts), nth ts i]
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   319
fun taut_or1_term i ts = [mk_nary HOLogic.mk_disj ts, HOLogic.mk_not (nth ts i)]
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   320
fun taut_or2_term ts = HOLogic.mk_not (mk_nary HOLogic.mk_disj ts) :: ts
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   321
67091
1393c2340eec more symbols;
wenzelm
parents: 66301
diff changeset
   322
val iff_1_taut = @{lemma "P = Q \<or> P \<or> Q" by fast}
1393c2340eec more symbols;
wenzelm
parents: 66301
diff changeset
   323
val iff_2_taut = @{lemma "P = Q \<or> (\<not>P) \<or> (\<not>Q)" by fast}
1393c2340eec more symbols;
wenzelm
parents: 66301
diff changeset
   324
val iff_3_taut = @{lemma "\<not>(P = Q) \<or> (\<not>P) \<or> Q" by fast}
1393c2340eec more symbols;
wenzelm
parents: 66301
diff changeset
   325
val iff_4_taut = @{lemma "\<not>(P = Q) \<or> P \<or> (\<not>Q)" by fast}
1393c2340eec more symbols;
wenzelm
parents: 66301
diff changeset
   326
val ite_then_taut = @{lemma "\<not>P \<or> (if P then t else u) = t" by auto}
1393c2340eec more symbols;
wenzelm
parents: 66301
diff changeset
   327
val ite_else_taut = @{lemma "P \<or> (if P then t else u) = u" by auto}
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   328
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   329
fun taut_rule_of ctxt (Argo_Proof.Taut_And_1 n) = with_frees ctxt n taut_and1_term
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   330
  | taut_rule_of ctxt (Argo_Proof.Taut_And_2 (i, n)) = with_frees ctxt n (taut_and2_term i)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   331
  | taut_rule_of ctxt (Argo_Proof.Taut_Or_1 (i, n)) = with_frees ctxt n (taut_or1_term i)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   332
  | taut_rule_of ctxt (Argo_Proof.Taut_Or_2 n) = with_frees ctxt n taut_or2_term
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   333
  | taut_rule_of _ Argo_Proof.Taut_Iff_1 = iff_1_taut
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   334
  | taut_rule_of _ Argo_Proof.Taut_Iff_2 = iff_2_taut
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   335
  | taut_rule_of _ Argo_Proof.Taut_Iff_3 = iff_3_taut
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   336
  | taut_rule_of _ Argo_Proof.Taut_Iff_4 = iff_4_taut
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   337
  | taut_rule_of _ Argo_Proof.Taut_Ite_Then = ite_then_taut
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   338
  | taut_rule_of _ Argo_Proof.Taut_Ite_Else = ite_else_taut
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   339
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   340
fun replay_taut ctxt k ct =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   341
  let val rule = taut_rule_of ctxt k
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   342
  in Thm.instantiate (Thm.match (Thm.cprop_of rule, ct)) rule end
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   343
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   344
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   345
(* proof replay for conjunct extraction *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   346
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   347
fun replay_conjunct 0 1 thm = thm
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   348
  | replay_conjunct 0 _ thm = discharge thm @{thm conjunct1}
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   349
  | replay_conjunct 1 2 thm = discharge thm @{thm conjunct2}
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   350
  | replay_conjunct i n thm = replay_conjunct (i - 1) (n - 1) (discharge thm @{thm conjunct2})
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   351
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   352
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   353
(* proof replay for rewrite steps *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   354
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   355
fun mk_rewr thm = thm RS @{thm eq_reflection}
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   356
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   357
fun not_nary_conv rule i ct =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   358
  if i > 1 then (Conv.rewr_conv rule then_conv Conv.arg_conv (not_nary_conv rule (i - 1))) ct
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   359
  else Conv.all_conv ct
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   360
67091
1393c2340eec more symbols;
wenzelm
parents: 66301
diff changeset
   361
val flatten_and_thm = @{lemma "(P1 \<and> P2) \<and> P3 \<equiv> P1 \<and> P2 \<and> P3" by simp}
1393c2340eec more symbols;
wenzelm
parents: 66301
diff changeset
   362
val flatten_or_thm = @{lemma "(P1 \<or> P2) \<or> P3 \<equiv> P1 \<or> P2 \<or> P3" by simp}
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   363
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   364
fun flatten_conv cv rule ct = (
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   365
  Conv.try_conv (Conv.arg_conv cv) then_conv
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   366
  Conv.try_conv (Conv.rewr_conv rule then_conv cv)) ct
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   367
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   368
fun flat_conj_conv ct =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   369
  (case Thm.term_of ct of
74323
5c452041fe83 clarified antiquotations;
wenzelm
parents: 74282
diff changeset
   370
    \<^Const_>\<open>conj for _ _\<close> => flatten_conv flat_conj_conv flatten_and_thm ct
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   371
  | _ => Conv.all_conv ct)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   372
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   373
fun flat_disj_conv ct =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   374
  (case Thm.term_of ct of
74323
5c452041fe83 clarified antiquotations;
wenzelm
parents: 74282
diff changeset
   375
    \<^Const_>\<open>disj for _ _\<close> => flatten_conv flat_disj_conv flatten_or_thm ct
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   376
  | _ => Conv.all_conv ct)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   377
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   378
fun explode rule1 rule2 thm =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   379
  explode rule1 rule2 (thm RS rule1) @ explode rule1 rule2 (thm RS rule2) handle THM _ => [thm]
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   380
val explode_conj = explode @{thm conjunct1} @{thm conjunct2}
67091
1393c2340eec more symbols;
wenzelm
parents: 66301
diff changeset
   381
val explode_ndis = explode @{lemma "\<not>(P \<or> Q) \<Longrightarrow> \<not>P" by auto} @{lemma "\<not>(P \<or> Q) \<Longrightarrow> \<not>Q" by auto}
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   382
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   383
fun pick_false i thms = nth thms i
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   384
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   385
fun pick_dual rule (i1, i2) thms =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   386
  rule OF [nth thms i1, nth thms i2] handle THM _ => rule OF [nth thms i2, nth thms i1]
67091
1393c2340eec more symbols;
wenzelm
parents: 66301
diff changeset
   387
val pick_dual_conj = pick_dual @{lemma "\<not>P \<Longrightarrow> P \<Longrightarrow> False" by auto}
1393c2340eec more symbols;
wenzelm
parents: 66301
diff changeset
   388
val pick_dual_ndis = pick_dual @{lemma "\<not>P \<Longrightarrow> P \<Longrightarrow> \<not>True" by auto}
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   389
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   390
fun join thm0 rule is thms =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   391
  let
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   392
    val l = length thms
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   393
    val thms' = fold (fn i => cons (if 0 <= i andalso i < l then nth thms i else thm0)) is []
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   394
  in fold (fn thm => fn thm' => discharge2 thm thm' rule) (tl thms') (hd thms') end
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   395
67091
1393c2340eec more symbols;
wenzelm
parents: 66301
diff changeset
   396
val join_conj = join @{lemma "True" by auto} @{lemma "P \<Longrightarrow> Q \<Longrightarrow> P \<and> Q" by auto}
1393c2340eec more symbols;
wenzelm
parents: 66301
diff changeset
   397
val join_ndis = join @{lemma "\<not>False" by auto} @{lemma "\<not>P \<Longrightarrow> \<not>Q \<Longrightarrow> \<not>(P \<or> Q)" by auto}
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   398
67091
1393c2340eec more symbols;
wenzelm
parents: 66301
diff changeset
   399
val false_thm = @{lemma "False \<Longrightarrow> P" by auto}
1393c2340eec more symbols;
wenzelm
parents: 66301
diff changeset
   400
val ntrue_thm = @{lemma "\<not>True \<Longrightarrow> P" by auto}
1393c2340eec more symbols;
wenzelm
parents: 66301
diff changeset
   401
val iff_conj_thm = @{lemma "(P \<Longrightarrow> Q) \<Longrightarrow> (Q \<Longrightarrow> P) \<Longrightarrow> P = Q" by auto}
1393c2340eec more symbols;
wenzelm
parents: 66301
diff changeset
   402
val iff_ndis_thm = @{lemma "(\<not>P \<Longrightarrow> \<not>Q) \<Longrightarrow> (\<not>Q \<Longrightarrow> \<not>P) \<Longrightarrow> P = Q" by auto}
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   403
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   404
fun iff_intro rule lf rf ct =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   405
  let
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   406
    val lhs = under_assumption lf ct
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   407
    val rhs = rf (Thm.dest_arg (snd (Thm.dest_implies (Thm.cprop_of lhs))))
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   408
  in mk_rewr (discharge2 lhs rhs rule) end
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   409
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   410
fun with_conj f g ct = iff_intro iff_conj_thm (f o explode_conj) g ct
74323
5c452041fe83 clarified antiquotations;
wenzelm
parents: 74282
diff changeset
   411
fun with_ndis f g ct = iff_intro iff_ndis_thm (f o explode_ndis) g (Thm.apply \<^cterm>\<open>Not\<close> ct)
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   412
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   413
fun swap_indices n iss = map (fn i => find_index (fn is => member (op =) is i) iss) (0 upto (n - 1))
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   414
fun sort_nary w f g (n, iss) = w (f (map hd iss)) (under_assumption (f (swap_indices n iss) o g))
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   415
val sort_conj = sort_nary with_conj join_conj explode_conj
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   416
val sort_ndis = sort_nary with_ndis join_ndis explode_ndis 
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   417
67091
1393c2340eec more symbols;
wenzelm
parents: 66301
diff changeset
   418
val not_true_thm = mk_rewr @{lemma "(\<not>True) = False" by auto}
1393c2340eec more symbols;
wenzelm
parents: 66301
diff changeset
   419
val not_false_thm = mk_rewr @{lemma "(\<not>False) = True" by auto}
1393c2340eec more symbols;
wenzelm
parents: 66301
diff changeset
   420
val not_not_thm = mk_rewr @{lemma "(\<not>\<not>P) = P" by auto}
1393c2340eec more symbols;
wenzelm
parents: 66301
diff changeset
   421
val not_and_thm = mk_rewr @{lemma "(\<not>(P \<and> Q)) = (\<not>P \<or> \<not>Q)" by auto}
1393c2340eec more symbols;
wenzelm
parents: 66301
diff changeset
   422
val not_or_thm = mk_rewr @{lemma "(\<not>(P \<or> Q)) = (\<not>P \<and> \<not>Q)" by auto}
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   423
val not_iff_thms = map mk_rewr
67091
1393c2340eec more symbols;
wenzelm
parents: 66301
diff changeset
   424
  @{lemma "(\<not>((\<not>P) = Q)) = (P = Q)" "(\<not>(P = (\<not>Q))) = (P = Q)" "(\<not>(P = Q)) = ((\<not>P) = Q)" by auto}
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   425
val iff_true_thms = map mk_rewr @{lemma "(True = P) = P" "(P = True) = P" by auto}
67091
1393c2340eec more symbols;
wenzelm
parents: 66301
diff changeset
   426
val iff_false_thms = map mk_rewr @{lemma "(False = P) = (\<not>P)" "(P = False) = (\<not>P)" by auto}
1393c2340eec more symbols;
wenzelm
parents: 66301
diff changeset
   427
val iff_not_not_thm = mk_rewr @{lemma "((\<not>P) = (\<not>Q)) = (P = Q)" by auto}
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   428
val iff_refl_thm = mk_rewr @{lemma "(P = P) = True" by auto}
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   429
val iff_symm_thm = mk_rewr @{lemma "(P = Q) = (Q = P)" by auto}
67091
1393c2340eec more symbols;
wenzelm
parents: 66301
diff changeset
   430
val iff_dual_thms = map mk_rewr @{lemma "(P = (\<not>P)) = False" "((\<not>P) = P) = False" by auto}
1393c2340eec more symbols;
wenzelm
parents: 66301
diff changeset
   431
val imp_thm = mk_rewr @{lemma "(P \<longrightarrow> Q) = (\<not>P \<or> Q)" by auto}
1393c2340eec more symbols;
wenzelm
parents: 66301
diff changeset
   432
val ite_prop_thm = mk_rewr @{lemma "(If P Q R) = ((\<not>P \<or> Q) \<and> (P \<or> R) \<and> (Q \<or> R))" by auto}
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   433
val ite_true_thm = mk_rewr @{lemma "(If True t u) = t" by auto}
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   434
val ite_false_thm = mk_rewr @{lemma "(If False t u) = u" by auto}
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   435
val ite_eq_thm = mk_rewr @{lemma "(If P t t) = t" by auto}
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   436
val eq_refl_thm = mk_rewr @{lemma "(t = t) = True" by auto}
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   437
val eq_symm_thm = mk_rewr @{lemma "(t1 = t2) = (t2 = t1)" by auto}
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   438
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   439
fun replay_rewr _ Argo_Proof.Rewr_Not_True = Conv.rewr_conv not_true_thm
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   440
  | replay_rewr _ Argo_Proof.Rewr_Not_False = Conv.rewr_conv not_false_thm
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   441
  | replay_rewr _ Argo_Proof.Rewr_Not_Not = Conv.rewr_conv not_not_thm
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   442
  | replay_rewr _ (Argo_Proof.Rewr_Not_And i) = not_nary_conv not_and_thm i
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   443
  | replay_rewr _ (Argo_Proof.Rewr_Not_Or i) = not_nary_conv not_or_thm i
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   444
  | replay_rewr _ Argo_Proof.Rewr_Not_Iff = Conv.rewrs_conv not_iff_thms
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   445
  | replay_rewr _ (Argo_Proof.Rewr_And_False i) = with_conj (pick_false i) (K false_thm)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   446
  | replay_rewr _ (Argo_Proof.Rewr_And_Dual ip) = with_conj (pick_dual_conj ip) (K false_thm)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   447
  | replay_rewr _ (Argo_Proof.Rewr_And_Sort is) = flat_conj_conv then_conv sort_conj is
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   448
  | replay_rewr _ (Argo_Proof.Rewr_Or_True i) = with_ndis (pick_false i) (K ntrue_thm)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   449
  | replay_rewr _ (Argo_Proof.Rewr_Or_Dual ip) = with_ndis (pick_dual_ndis ip) (K ntrue_thm)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   450
  | replay_rewr _ (Argo_Proof.Rewr_Or_Sort is) = flat_disj_conv then_conv sort_ndis is
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   451
  | replay_rewr _ Argo_Proof.Rewr_Iff_True = Conv.rewrs_conv iff_true_thms
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   452
  | replay_rewr _ Argo_Proof.Rewr_Iff_False = Conv.rewrs_conv iff_false_thms
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   453
  | replay_rewr _ Argo_Proof.Rewr_Iff_Not_Not = Conv.rewr_conv iff_not_not_thm
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   454
  | replay_rewr _ Argo_Proof.Rewr_Iff_Refl = Conv.rewr_conv iff_refl_thm
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   455
  | replay_rewr _ Argo_Proof.Rewr_Iff_Symm = Conv.rewr_conv iff_symm_thm
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   456
  | replay_rewr _ Argo_Proof.Rewr_Iff_Dual  = Conv.rewrs_conv iff_dual_thms
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   457
  | replay_rewr _ Argo_Proof.Rewr_Imp = Conv.rewr_conv imp_thm
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   458
  | replay_rewr _ Argo_Proof.Rewr_Ite_Prop = Conv.rewr_conv ite_prop_thm
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   459
  | replay_rewr _ Argo_Proof.Rewr_Ite_True = Conv.rewr_conv ite_true_thm
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   460
  | replay_rewr _ Argo_Proof.Rewr_Ite_False = Conv.rewr_conv ite_false_thm
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   461
  | replay_rewr _ Argo_Proof.Rewr_Ite_Eq = Conv.rewr_conv ite_eq_thm
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   462
  | replay_rewr _ Argo_Proof.Rewr_Eq_Refl = Conv.rewr_conv eq_refl_thm
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   463
  | replay_rewr _ Argo_Proof.Rewr_Eq_Symm = Conv.rewr_conv eq_symm_thm
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   464
  | replay_rewr ctxt r = ext_replay_rewr ctxt r
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   465
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   466
fun binop_conv cv1 cv2 = Conv.combination_conv (Conv.arg_conv cv1) cv2
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   467
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   468
fun replay_conv _ Argo_Proof.Keep_Conv ct = Conv.all_conv ct
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   469
  | replay_conv ctxt (Argo_Proof.Then_Conv (c1, c2)) ct = 
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   470
      (replay_conv ctxt c1 then_conv replay_conv ctxt c2) ct
66301
8a6a89d6cf2b more explicit Argo proof traces; more correct proof replay for term applications
boehmes
parents: 63992
diff changeset
   471
  | replay_conv ctxt (Argo_Proof.Args_Conv (Argo_Expr.App, [c1, c2])) ct =
8a6a89d6cf2b more explicit Argo proof traces; more correct proof replay for term applications
boehmes
parents: 63992
diff changeset
   472
      Conv.combination_conv (replay_conv ctxt c1) (replay_conv ctxt c2) ct
8a6a89d6cf2b more explicit Argo proof traces; more correct proof replay for term applications
boehmes
parents: 63992
diff changeset
   473
  | replay_conv ctxt (Argo_Proof.Args_Conv (_, cs)) ct = replay_args_conv ctxt cs ct
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   474
  | replay_conv ctxt (Argo_Proof.Rewr_Conv r) ct = replay_rewr ctxt r ct
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   475
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   476
and replay_args_conv _ [] ct = Conv.all_conv ct
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   477
  | replay_args_conv ctxt [c] ct = Conv.arg_conv (replay_conv ctxt c) ct
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   478
  | replay_args_conv ctxt [c1, c2] ct = binop_conv (replay_conv ctxt c1) (replay_conv ctxt c2) ct
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   479
  | replay_args_conv ctxt (c :: cs) ct =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   480
      (case Term.head_of (Thm.term_of ct) of
74323
5c452041fe83 clarified antiquotations;
wenzelm
parents: 74282
diff changeset
   481
        \<^Const_>\<open>If _\<close> =>
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   482
          let val (cs', c') = split_last cs
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   483
          in Conv.combination_conv (replay_args_conv ctxt (c :: cs')) (replay_conv ctxt c') ct end
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   484
      | _ => binop_conv (replay_conv ctxt c) (replay_args_conv ctxt cs) ct)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   485
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   486
fun replay_rewrite ctxt c thm = Conv.fconv_rule (HOLogic.Trueprop_conv (replay_conv ctxt c)) thm
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   487
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   488
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   489
(* proof replay for clauses *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   490
67091
1393c2340eec more symbols;
wenzelm
parents: 66301
diff changeset
   491
val prep_clause_rule = @{lemma "P \<Longrightarrow> \<not>P \<Longrightarrow> False" by fast}
1393c2340eec more symbols;
wenzelm
parents: 66301
diff changeset
   492
val extract_lit_rule = @{lemma "(\<not>(P \<or> Q) \<Longrightarrow> False) \<Longrightarrow> \<not>P \<Longrightarrow> \<not>Q \<Longrightarrow> False" by fast}
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   493
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   494
fun add_lit i thm lits =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   495
  let val ct = Thm.cprem_of thm 1
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   496
  in (Thm.implies_elim thm (Thm.assume ct), (i, ct) :: lits) end
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   497
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   498
fun extract_lits [] _ = error "Bad clause"
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   499
  | extract_lits [i] (thm, lits) = add_lit i thm lits
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   500
  | extract_lits (i :: is) (thm, lits) =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   501
      extract_lits is (add_lit i (discharge thm extract_lit_rule) lits)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   502
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   503
fun lit_ord ((l1, _), (l2, _)) = int_ord (abs l1, abs l2)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   504
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   505
fun replay_with_lits [] thm lits = (thm, lits)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   506
  | replay_with_lits is thm lits =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   507
      extract_lits is (discharge thm prep_clause_rule, lits)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   508
      ||> Ord_List.make lit_ord
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   509
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   510
fun replay_clause is thm = replay_with_lits is thm []
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   511
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   512
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   513
(* proof replay for unit resolution *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   514
74323
5c452041fe83 clarified antiquotations;
wenzelm
parents: 74282
diff changeset
   515
val bogus_ct = \<^cterm>\<open>True\<close>
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   516
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   517
fun replay_unit_res lit (pthm, plits) (nthm, nlits) =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   518
  let
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   519
    val plit = the (AList.lookup (op =) plits lit)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   520
    val nlit = the (AList.lookup (op =) nlits (~lit))
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   521
    val prune = Ord_List.remove lit_ord (lit, bogus_ct)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   522
  in
74647
b31683a544cf clarified antiquotations;
wenzelm
parents: 74646
diff changeset
   523
    \<^instantiate>\<open>P = \<open>Thm.dest_arg plit\<close> in
b31683a544cf clarified antiquotations;
wenzelm
parents: 74646
diff changeset
   524
      lemma "(P \<Longrightarrow> False) \<Longrightarrow> (\<not>P \<Longrightarrow> False) \<Longrightarrow> False" by fast\<close>
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   525
    |> Thm.elim_implies (Thm.implies_intr plit pthm)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   526
    |> Thm.elim_implies (Thm.implies_intr nlit nthm)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   527
    |> rpair (Ord_List.union lit_ord (prune nlits) (prune plits))
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   528
  end
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   529
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   530
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   531
(* proof replay for hypothesis *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   532
67091
1393c2340eec more symbols;
wenzelm
parents: 66301
diff changeset
   533
val dneg_rule = @{lemma "\<not>\<not>P \<Longrightarrow> P" by auto}
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   534
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   535
fun replay_hyp i ct =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   536
  if i < 0 then (Thm.assume ct, [(~i, ct)])
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   537
  else
81942
da3c3948a39c clarified signature: more uniform cterm operations, without context;
wenzelm
parents: 80910
diff changeset
   538
    let val cu = HOLogic.mk_judgment (Thm.apply \<^cterm>\<open>Not\<close> (Thm.apply \<^cterm>\<open>Not\<close> (Thm.dest_arg ct)))
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   539
    in (discharge (Thm.assume cu) dneg_rule, [(~i, cu)]) end
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   540
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   541
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   542
(* proof replay for lemma *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   543
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   544
fun replay_lemma is (thm, lits) = replay_with_lits is thm lits
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   545
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   546
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   547
(* proof replay for reflexivity *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   548
74647
b31683a544cf clarified antiquotations;
wenzelm
parents: 74646
diff changeset
   549
fun replay_refl x =
b31683a544cf clarified antiquotations;
wenzelm
parents: 74646
diff changeset
   550
  \<^instantiate>\<open>'a = \<open>Thm.ctyp_of_cterm x\<close> and x in lemma \<open>x = x\<close> by (rule refl)\<close>
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   551
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   552
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   553
(* proof replay for symmetry *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   554
67091
1393c2340eec more symbols;
wenzelm
parents: 66301
diff changeset
   555
val symm_rules = @{lemma "a = b ==> b = a" "\<not>(a = b) \<Longrightarrow> \<not>(b = a)" by simp_all}
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   556
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   557
fun replay_symm thm = hd (discharges thm symm_rules)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   558
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   559
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   560
(* proof replay for transitivity *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   561
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   562
val trans_rules = @{lemma
67091
1393c2340eec more symbols;
wenzelm
parents: 66301
diff changeset
   563
  "\<not>(a = b) \<Longrightarrow> b = c \<Longrightarrow> \<not>(a = c)"
1393c2340eec more symbols;
wenzelm
parents: 66301
diff changeset
   564
  "a = b \<Longrightarrow> \<not>(b = c) \<Longrightarrow> \<not>(a = c)"
1393c2340eec more symbols;
wenzelm
parents: 66301
diff changeset
   565
  "a = b \<Longrightarrow> b = c ==> a = c"
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   566
  by simp_all}
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   567
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   568
fun replay_trans thm1 thm2 = hd (discharges thm2 (discharges thm1 trans_rules))
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   569
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   570
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   571
(* proof replay for congruence *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   572
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   573
fun replay_cong thm1 thm2 = discharge thm2 (discharge thm1 @{thm cong})
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   574
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   575
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   576
(* proof replay for substitution *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   577
67091
1393c2340eec more symbols;
wenzelm
parents: 66301
diff changeset
   578
val subst_rule1 = @{lemma "\<not>(p a) \<Longrightarrow> p = q \<Longrightarrow> a = b \<Longrightarrow> \<not>(q b)" by simp}
1393c2340eec more symbols;
wenzelm
parents: 66301
diff changeset
   579
val subst_rule2 = @{lemma "p a \<Longrightarrow> p = q \<Longrightarrow> a = b \<Longrightarrow> q b" by simp}
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   580
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   581
fun replay_subst thm1 thm2 thm3 =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   582
  subst_rule1 OF [thm1, thm2, thm3] handle THM _ => subst_rule2 OF [thm1, thm2, thm3]
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   583
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   584
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   585
(* proof replay *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   586
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   587
structure Thm_Cache = Table(type key = Argo_Proof.proof_id val ord = Argo_Proof.proof_id_ord)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   588
67091
1393c2340eec more symbols;
wenzelm
parents: 66301
diff changeset
   589
val unclausify_rule1 = @{lemma "(\<not>P \<Longrightarrow> False) \<Longrightarrow> P" by auto}
1393c2340eec more symbols;
wenzelm
parents: 66301
diff changeset
   590
val unclausify_rule2 = @{lemma "(P \<Longrightarrow> False) \<Longrightarrow> \<not>P" by auto}
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   591
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   592
fun unclausify (thm, lits) ls =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   593
  (case (Thm.prop_of thm, lits) of
74375
ba880f3a4e52 tuned antiquotations;
wenzelm
parents: 74323
diff changeset
   594
    (\<^Const_>\<open>Trueprop for \<^Const_>\<open>False\<close>\<close>, [(_, ct)]) =>
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   595
      let val thm = Thm.implies_intr ct thm
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   596
      in (discharge thm unclausify_rule1 handle THM _ => discharge thm unclausify_rule2, ls) end
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   597
  | _ => (thm, Ord_List.union lit_ord lits ls))
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   598
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   599
fun with_thms f tps = fold_map unclausify tps [] |>> f
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   600
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   601
fun bad_premises () = raise Fail "bad number of premises"
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   602
fun with_thms1 f = with_thms (fn [thm] => f thm | _ => bad_premises ())
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   603
fun with_thms2 f = with_thms (fn [thm1, thm2] => f thm1 thm2 | _ => bad_premises ())
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   604
fun with_thms3 f = with_thms (fn [thm1, thm2, thm3] => f thm1 thm2 thm3 | _ => bad_premises ())
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   605
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   606
fun replay_rule (ctxt, cons, facts) prems rule =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   607
  (case rule of
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   608
    Argo_Proof.Axiom i => (nth facts i, [])
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   609
  | Argo_Proof.Taut (k, concl) => (replay_taut ctxt k (cprop_of ctxt cons concl), [])
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   610
  | Argo_Proof.Conjunct (i, n) => with_thms1 (replay_conjunct i n) prems
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   611
  | Argo_Proof.Rewrite c => with_thms1 (replay_rewrite ctxt c) prems
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   612
  | Argo_Proof.Clause is => replay_clause is (fst (hd prems))
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   613
  | Argo_Proof.Unit_Res i => replay_unit_res i (hd prems) (hd (tl prems))
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   614
  | Argo_Proof.Hyp (i, concl) => replay_hyp i (cprop_of ctxt cons concl)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   615
  | Argo_Proof.Lemma is => replay_lemma is (hd prems)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   616
  | Argo_Proof.Refl concl => (replay_refl (cterm_of ctxt cons concl), [])
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   617
  | Argo_Proof.Symm => with_thms1 replay_symm prems
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   618
  | Argo_Proof.Trans => with_thms2 replay_trans prems
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   619
  | Argo_Proof.Cong => with_thms2 replay_cong prems
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   620
  | Argo_Proof.Subst => with_thms3 replay_subst prems
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   621
  | _ => with_thms (ext_replay (cprop_of ctxt cons) ctxt rule) prems)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   622
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   623
fun with_cache f proof thm_cache =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   624
  (case Thm_Cache.lookup thm_cache (Argo_Proof.id_of proof) of
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   625
    SOME thm => (thm, thm_cache)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   626
  | NONE =>
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   627
      let val (thm, thm_cache') = f proof thm_cache
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   628
      in (thm, Thm_Cache.update (Argo_Proof.id_of proof, thm) thm_cache') end)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   629
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   630
fun trace_step ctxt proof_id rule proofs = with_full_tracing ctxt (fn ctxt' =>
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   631
  let
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   632
    val id = Argo_Proof.string_of_proof_id proof_id
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   633
    val ids = map (Argo_Proof.string_of_proof_id o Argo_Proof.id_of) proofs
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   634
    val rule_string = Argo_Proof.string_of_rule rule
80910
406a85a25189 clarified signature: more explicit operations;
wenzelm
parents: 77879
diff changeset
   635
  in full_tracing ctxt' ("  " ^ id ^ " <- " ^ implode_space ids ^ " . " ^ rule_string) end)
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   636
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   637
fun replay_bottom_up (env as (ctxt, _, _)) proof thm_cache =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   638
  let
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   639
    val (proof_id, rule, proofs) = Argo_Proof.dest proof
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   640
    val (prems, thm_cache) = fold_map (with_cache (replay_bottom_up env)) proofs thm_cache
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   641
    val _ = trace_step ctxt proof_id rule proofs
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   642
  in (replay_rule env prems rule, thm_cache) end
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   643
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   644
fun replay_proof env proof = with_cache (replay_bottom_up env) proof Thm_Cache.empty
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   645
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   646
fun replay ctxt terms facts proof =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   647
  let
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   648
    val env = (ctxt, Termtab.fold (Contab.update o swap) terms Contab.empty, facts)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   649
    val _ = tracing ctxt "replaying the proof"
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   650
    val ({elapsed=t, ...}, ((thm, _), _)) = Timing.timing (replay_proof env) proof
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   651
    val _ = tracing ctxt ("replayed the proof in " ^ string_of_int (Time.toMilliseconds t) ^ " ms")
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   652
  in thm end
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   653
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   654
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   655
(* normalizing goals *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   656
77879
wenzelm
parents: 74648
diff changeset
   657
fun instantiate v ct = Thm.instantiate (TVars.empty, Vars.make1 (v, ct))
74647
b31683a544cf clarified antiquotations;
wenzelm
parents: 74646
diff changeset
   658
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   659
fun instantiate_elim_rule thm =
74647
b31683a544cf clarified antiquotations;
wenzelm
parents: 74646
diff changeset
   660
  let
b31683a544cf clarified antiquotations;
wenzelm
parents: 74646
diff changeset
   661
    val ct = Drule.strip_imp_concl (Thm.cprop_of thm)
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   662
  in
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   663
    (case Thm.term_of ct of
74647
b31683a544cf clarified antiquotations;
wenzelm
parents: 74646
diff changeset
   664
      \<^Const_>\<open>Trueprop for \<open>Var (v as (_, \<^Type>\<open>bool\<close>))\<close>\<close> => instantiate v \<^cterm>\<open>False\<close> thm
b31683a544cf clarified antiquotations;
wenzelm
parents: 74646
diff changeset
   665
    | Var v => instantiate v \<^cprop>\<open>False\<close> thm
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   666
    | _ => thm)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   667
  end
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   668
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   669
fun atomize_conv ctxt ct =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   670
  (case Thm.term_of ct of
74323
5c452041fe83 clarified antiquotations;
wenzelm
parents: 74282
diff changeset
   671
    \<^Const_>\<open>Trueprop for _\<close> => Conv.all_conv
5c452041fe83 clarified antiquotations;
wenzelm
parents: 74282
diff changeset
   672
  | \<^Const_>\<open>Pure.imp for _ _\<close> =>
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   673
      Conv.binop_conv (atomize_conv ctxt) then_conv
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   674
      Conv.rewr_conv @{thm atomize_imp}
74323
5c452041fe83 clarified antiquotations;
wenzelm
parents: 74282
diff changeset
   675
  | \<^Const>\<open>Pure.eq _ for _ _\<close> =>
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   676
      Conv.binop_conv (atomize_conv ctxt) then_conv
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   677
      Conv.rewr_conv @{thm atomize_eq}
74323
5c452041fe83 clarified antiquotations;
wenzelm
parents: 74282
diff changeset
   678
  | \<^Const>\<open>Pure.all _ for \<open>Abs _\<close>\<close> =>
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   679
      Conv.binder_conv (atomize_conv o snd) ctxt then_conv
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   680
      Conv.rewr_conv @{thm atomize_all}
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   681
  | _ => Conv.all_conv) ct
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   682
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   683
fun normalize ctxt thm =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   684
  thm
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   685
  |> instantiate_elim_rule
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   686
  |> Conv.fconv_rule (Thm.beta_conversion true then_conv Thm.eta_conversion)
74245
282cd3aa6cc6 clarified modules;
wenzelm
parents: 69578
diff changeset
   687
  |> Thm.forall_intr_vars
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   688
  |> Conv.fconv_rule (atomize_conv ctxt)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   689
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   690
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   691
(* prover, tactic and method *)
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   692
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   693
datatype result = Satisfiable of term -> bool option | Unsatisfiable
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   694
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   695
fun check props ctxt =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   696
  (case solve props ctxt of
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   697
    (Proof _, ctxt') => (Unsatisfiable, ctxt')
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   698
  | (Model model, ctxt') => (Satisfiable model, ctxt'))
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   699
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   700
fun prove thms ctxt =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   701
  let val thms' = map (normalize ctxt) thms
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   702
  in
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   703
    (case solve (map Thm.prop_of thms') ctxt of
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   704
      (Model _, ctxt') => (NONE, ctxt')
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   705
    | (Proof (terms, proof), ctxt') => (SOME (replay ctxt' terms thms' proof), ctxt'))
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   706
  end
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   707
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   708
fun argo_tac ctxt thms =
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   709
  CONVERSION (Conv.params_conv ~1 (K (Conv.concl_conv ~1
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   710
    (Conv.try_conv (Conv.rewr_conv @{thm atomize_eq})))) ctxt)
74646
wenzelm
parents: 74561
diff changeset
   711
  THEN' resolve_tac ctxt [@{thm ccontr}]
wenzelm
parents: 74561
diff changeset
   712
  THEN' Subgoal.FOCUS (fn {context = ctxt', prems, ...} =>
wenzelm
parents: 74561
diff changeset
   713
    (case with_timeout (prove (thms @ prems)) ctxt' of
wenzelm
parents: 74561
diff changeset
   714
      (SOME thm, _) => resolve_tac ctxt' [thm] 1
wenzelm
parents: 74561
diff changeset
   715
    | (NONE, _) => no_tac)) ctxt
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   716
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   717
val _ =
67149
e61557884799 prefer control symbol antiquotations;
wenzelm
parents: 67091
diff changeset
   718
  Theory.setup (Method.setup \<^binding>\<open>argo\<close>
63960
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   719
    (Scan.optional Attrib.thms [] >>
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   720
      (fn thms => fn ctxt => METHOD (fn facts =>
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   721
        HEADGOAL (argo_tac ctxt (thms @ facts)))))
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   722
    "Applies the Argo prover")
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   723
3daf02070be5 new proof method "argo" for a combination of quantifier-free propositional logic with equality and linear real arithmetic
boehmes
parents:
diff changeset
   724
end