src/HOL/Tools/SMT/smt_utils.ML
author boehmes
Mon, 22 Nov 2010 15:45:42 +0100
changeset 40662 798aad2229c0
child 40663 e080c9e68752
permissions -rw-r--r--
added prove reconstruction for injective functions; added SMT_Utils to collect frequently used functions
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
40662
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
     1
(*  Title:      HOL/Tools/SMT/smt_utils.ML
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
     2
    Author:     Sascha Boehme, TU Muenchen
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
     3
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
     4
General utility functions.
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
     5
*)
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
     6
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
     7
signature SMT_UTILS =
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
     8
sig
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
     9
  val repeat: ('a -> 'a option) -> 'a -> 'a
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    10
  val repeat_yield: ('a -> 'b -> ('a * 'b) option) -> 'a -> 'b -> 'a * 'b
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    11
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    12
  (* terms *)
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    13
  val dest_conj: term -> term * term
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    14
  val dest_disj: term -> term * term
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    15
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    16
  (* patterns and instantiations *)
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    17
  val mk_const_pat: theory -> string -> (ctyp -> 'a) -> 'a * cterm
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    18
  val destT1: ctyp -> ctyp
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    19
  val destT2: ctyp -> ctyp
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    20
  val instTs: ctyp list -> ctyp list * cterm -> cterm
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    21
  val instT: ctyp -> ctyp * cterm -> cterm
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    22
  val instT': cterm -> ctyp * cterm -> cterm
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    23
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    24
  (* certified terms *)
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    25
  val certify: Proof.context -> term -> cterm
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    26
  val dest_cabs: cterm -> Proof.context -> cterm * Proof.context
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    27
  val dest_all_cabs: cterm -> Proof.context -> cterm * Proof.context
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    28
  val dest_cbinder: cterm -> Proof.context -> cterm * Proof.context
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    29
  val dest_all_cbinders: cterm -> Proof.context -> cterm * Proof.context
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    30
  val mk_cprop: cterm -> cterm
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    31
  val dest_cprop: cterm -> cterm
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    32
  val mk_cequals: cterm -> cterm -> cterm
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    33
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    34
  (* conversions *)
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    35
  val if_conv: (term -> bool) -> conv -> conv -> conv
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    36
  val if_true_conv: (term -> bool) -> conv -> conv
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    37
  val binders_conv: (Proof.context -> conv) -> Proof.context -> conv
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    38
  val prop_conv: conv -> conv
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    39
end
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    40
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    41
structure SMT_Utils: SMT_UTILS =
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    42
struct
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    43
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    44
fun repeat f =
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    45
  let fun rep x = (case f x of SOME y => rep y | NONE => x)
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    46
  in rep end
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    47
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    48
fun repeat_yield f =
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    49
  let fun rep x y = (case f x y of SOME (x', y') => rep x' y' | NONE => (x, y))
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    50
  in rep end
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    51
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    52
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    53
(* terms *)
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    54
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    55
fun dest_conj (@{const HOL.conj} $ t $ u) = (t, u)
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    56
  | dest_conj t = raise TERM ("not a conjunction", [t])
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    57
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    58
fun dest_disj (@{const HOL.disj} $ t $ u) = (t, u)
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    59
  | dest_disj t = raise TERM ("not a disjunction", [t])
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    60
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    61
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    62
(* patterns and instantiations *)
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    63
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    64
fun mk_const_pat thy name destT =
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    65
  let val cpat = Thm.cterm_of thy (Const (name, Sign.the_const_type thy name))
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    66
  in (destT (Thm.ctyp_of_term cpat), cpat) end
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    67
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    68
val destT1 = hd o Thm.dest_ctyp
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    69
val destT2 = hd o tl o Thm.dest_ctyp
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    70
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    71
fun instTs cUs (cTs, ct) = Thm.instantiate_cterm (cTs ~~ cUs, []) ct
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    72
fun instT cU (cT, ct) = instTs [cU] ([cT], ct)
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    73
fun instT' ct = instT (Thm.ctyp_of_term ct)
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    74
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    75
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    76
(* certified terms *)
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    77
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    78
fun certify ctxt = Thm.cterm_of (ProofContext.theory_of ctxt)
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    79
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    80
fun dest_cabs ct ctxt =
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    81
  (case Thm.term_of ct of
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    82
    Abs _ =>
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    83
      let val (n, ctxt') = yield_singleton Variable.variant_fixes Name.uu ctxt
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    84
      in (snd (Thm.dest_abs (SOME n) ct), ctxt') end
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    85
  | _ => raise CTERM ("no abstraction", [ct]))
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    86
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    87
val dest_all_cabs = repeat_yield (try o dest_cabs) 
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    88
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    89
fun dest_cbinder ct ctxt =
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    90
  (case Thm.term_of ct of
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    91
    Const _ $ Abs _ => dest_cabs (Thm.dest_arg ct) ctxt
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    92
  | _ => raise CTERM ("not a binder", [ct]))
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    93
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    94
val dest_all_cbinders = repeat_yield (try o dest_cbinder)
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    95
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    96
val mk_cprop = Thm.capply @{cterm Trueprop}
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    97
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    98
fun dest_cprop ct =
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
    99
  (case Thm.term_of ct of
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
   100
    @{const Trueprop} $ _ => Thm.dest_arg ct
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
   101
  | _ => raise CTERM ("not a property", [ct]))
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
   102
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
   103
val equals = mk_const_pat @{theory} @{const_name "=="} destT1
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
   104
fun mk_cequals ct cu = Thm.mk_binop (instT' ct equals) ct cu
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
   105
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
   106
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
   107
(* conversions *)
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
   108
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
   109
fun if_conv f cv1 cv2 ct = if f (Thm.term_of ct) then cv1 ct else cv2 ct
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
   110
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
   111
fun if_true_conv f cv = if_conv f cv Conv.all_conv
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
   112
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
   113
fun binders_conv cv ctxt =
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
   114
  Conv.binder_conv (binders_conv cv o snd) ctxt else_conv cv ctxt
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
   115
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
   116
fun prop_conv cv ct =
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
   117
  (case Thm.term_of ct of
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
   118
    @{const Trueprop} $ _ => Conv.arg_conv cv ct
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
   119
  | _ => raise CTERM ("not a property", [ct]))
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
   120
798aad2229c0 added prove reconstruction for injective functions;
boehmes
parents:
diff changeset
   121
end