src/HOL/Tools/SMT/smt_normalize.ML
author boehmes
Tue, 07 Dec 2010 14:53:12 +0100
changeset 41307 d2b1fc1b8e19
parent 40934 4725ed462387
child 41320 9f9bc1bdacef
permissions -rw-r--r--
centralized handling of built-in types and constants;
also store types and constants which are rewritten during preprocessing;
interfaces are identified by classes (supporting inheritance, at least on the level of built-in symbols);
removed term_eq in favor of type replacements: term-level occurrences of type bool are replaced by type term_bool (only for the translation)
boehmes@36890
     1
(*  Title:      HOL/Tools/SMT/smt_normalize.ML
boehmes@36890
     2
    Author:     Sascha Boehme, TU Muenchen
boehmes@36890
     3
boehmes@36890
     4
Normalization steps on theorems required by SMT solvers:
boehmes@36890
     5
  * simplify trivial distincts (those with less than three elements),
boehmes@36890
     6
  * rewrite bool case expressions as if expressions,
boehmes@36890
     7
  * normalize numerals (e.g. replace negative numerals by negated positive
boehmes@36890
     8
    numerals),
boehmes@36890
     9
  * embed natural numbers into integers,
boehmes@36890
    10
  * add extra rules specifying types and constants which occur frequently,
boehmes@36890
    11
  * fully translate into object logic, add universal closure,
boehmes@39723
    12
  * monomorphize (create instances of schematic rules),
boehmes@36890
    13
  * lift lambda terms,
boehmes@36890
    14
  * make applications explicit for functions with varying number of arguments.
boehmes@39723
    15
  * add (hypothetical definitions for) missing datatype selectors,
boehmes@36890
    16
*)
boehmes@36890
    17
boehmes@36890
    18
signature SMT_NORMALIZE =
boehmes@36890
    19
sig
boehmes@40403
    20
  type extra_norm = bool -> (int * thm) list -> Proof.context ->
boehmes@40402
    21
    (int * thm) list * Proof.context
boehmes@40651
    22
  val normalize: extra_norm -> bool -> (int * thm) list -> Proof.context ->
boehmes@40402
    23
    (int * thm) list * Proof.context
boehmes@36891
    24
  val atomize_conv: Proof.context -> conv
boehmes@36890
    25
  val eta_expand_conv: (Proof.context -> conv) -> Proof.context -> conv
boehmes@41307
    26
  val setup: theory -> theory
boehmes@36890
    27
end
boehmes@36890
    28
boehmes@36890
    29
structure SMT_Normalize: SMT_NORMALIZE =
boehmes@36890
    30
struct
boehmes@36890
    31
boehmes@40911
    32
structure U = SMT_Utils
boehmes@41307
    33
structure B = SMT_Builtin
boehmes@40911
    34
boehmes@36890
    35
infix 2 ??
boehmes@36890
    36
fun (test ?? f) x = if test x then f x else x
boehmes@36890
    37
boehmes@36890
    38
boehmes@36890
    39
boehmes@40933
    40
(* instantiate elimination rules *)
boehmes@40933
    41
 
boehmes@40933
    42
local
boehmes@40933
    43
  val (cpfalse, cfalse) = `U.mk_cprop (Thm.cterm_of @{theory} @{const False})
boehmes@40933
    44
boehmes@40933
    45
  fun inst f ct thm =
boehmes@40933
    46
    let val cv = f (Drule.strip_imp_concl (Thm.cprop_of thm))
boehmes@40933
    47
    in Thm.instantiate ([], [(cv, ct)]) thm end
boehmes@40933
    48
in
boehmes@40933
    49
boehmes@40933
    50
fun instantiate_elim thm =
boehmes@40933
    51
  (case Thm.concl_of thm of
boehmes@40933
    52
    @{const Trueprop} $ Var (_, @{typ bool}) => inst Thm.dest_arg cfalse thm
boehmes@40933
    53
  | Var _ => inst I cpfalse thm
boehmes@40933
    54
  | _ => thm)
boehmes@40933
    55
boehmes@40933
    56
end
boehmes@40933
    57
boehmes@40933
    58
boehmes@40933
    59
boehmes@36890
    60
(* simplification of trivial distincts (distinct should have at least
boehmes@36890
    61
   three elements in the argument list) *)
boehmes@36890
    62
boehmes@36890
    63
local
boehmes@40929
    64
  fun is_trivial_distinct (Const (@{const_name distinct}, _) $ t) =
boehmes@40929
    65
        (case try HOLogic.dest_list t of
boehmes@40929
    66
          SOME [] => true
boehmes@40929
    67
        | SOME [_] => true
boehmes@40929
    68
        | _ => false)
boehmes@36890
    69
    | is_trivial_distinct _ = false
boehmes@36890
    70
boehmes@37786
    71
  val thms = map mk_meta_eq @{lemma
boehmes@40929
    72
    "distinct [] = True"
boehmes@40929
    73
    "distinct [x] = True"
boehmes@40929
    74
    "distinct [x, y] = (x ~= y)"
boehmes@40929
    75
    by simp_all}
boehmes@36890
    76
  fun distinct_conv _ =
boehmes@40911
    77
    U.if_true_conv is_trivial_distinct (Conv.rewrs_conv thms)
boehmes@36890
    78
in
boehmes@36890
    79
fun trivial_distinct ctxt =
boehmes@40402
    80
  map (apsnd ((Term.exists_subterm is_trivial_distinct o Thm.prop_of) ??
boehmes@40402
    81
    Conv.fconv_rule (Conv.top_conv distinct_conv ctxt)))
boehmes@36890
    82
end
boehmes@36890
    83
boehmes@36890
    84
boehmes@36890
    85
boehmes@36890
    86
(* rewrite bool case expressions as if expressions *)
boehmes@36890
    87
boehmes@36890
    88
local
boehmes@36890
    89
  val is_bool_case = (fn
boehmes@36890
    90
      Const (@{const_name "bool.bool_case"}, _) $ _ $ _ $ _ => true
boehmes@36890
    91
    | _ => false)
boehmes@36890
    92
boehmes@40521
    93
  val thm = mk_meta_eq @{lemma
boehmes@40521
    94
    "(case P of True => x | False => y) = (if P then x else y)" by simp}
boehmes@40911
    95
  val unfold_conv = U.if_true_conv is_bool_case (Conv.rewr_conv thm)
boehmes@36890
    96
in
boehmes@36890
    97
fun rewrite_bool_cases ctxt =
boehmes@40402
    98
  map (apsnd ((Term.exists_subterm is_bool_case o Thm.prop_of) ??
boehmes@40402
    99
    Conv.fconv_rule (Conv.top_conv (K unfold_conv) ctxt)))
boehmes@41307
   100
boehmes@41307
   101
val setup_bool_case = B.add_builtin_fun_ext'' @{const_name "bool.bool_case"}
boehmes@41307
   102
boehmes@36890
   103
end
boehmes@36890
   104
boehmes@36890
   105
boehmes@36890
   106
boehmes@36890
   107
(* normalization of numerals: rewriting of negative integer numerals into
boehmes@36890
   108
   positive numerals, Numeral0 into 0, Numeral1 into 1 *)
boehmes@36890
   109
boehmes@36890
   110
local
boehmes@36890
   111
  fun is_number_sort ctxt T =
boehmes@36890
   112
    Sign.of_sort (ProofContext.theory_of ctxt) (T, @{sort number_ring})
boehmes@36890
   113
boehmes@36890
   114
  fun is_strange_number ctxt (t as Const (@{const_name number_of}, _) $ _) =
boehmes@36890
   115
        (case try HOLogic.dest_number t of
boehmes@36890
   116
          SOME (T, i) => is_number_sort ctxt T andalso i < 2
boehmes@36890
   117
        | NONE => false)
boehmes@36890
   118
    | is_strange_number _ _ = false
boehmes@36890
   119
boehmes@36890
   120
  val pos_numeral_ss = HOL_ss
boehmes@36890
   121
    addsimps [@{thm Int.number_of_minus}, @{thm Int.number_of_Min}]
boehmes@36890
   122
    addsimps [@{thm Int.number_of_Pls}, @{thm Int.numeral_1_eq_1}]
boehmes@36890
   123
    addsimps @{thms Int.pred_bin_simps}
boehmes@36890
   124
    addsimps @{thms Int.normalize_bin_simps}
boehmes@36890
   125
    addsimps @{lemma
boehmes@36890
   126
      "Int.Min = - Int.Bit1 Int.Pls"
boehmes@36890
   127
      "Int.Bit0 (- Int.Pls) = - Int.Pls"
boehmes@36890
   128
      "Int.Bit0 (- k) = - Int.Bit0 k"
boehmes@36890
   129
      "Int.Bit1 (- k) = - Int.Bit1 (Int.pred k)"
boehmes@36890
   130
      by simp_all (simp add: pred_def)}
boehmes@36890
   131
boehmes@40911
   132
  fun pos_conv ctxt = U.if_conv (is_strange_number ctxt)
boehmes@36890
   133
    (Simplifier.rewrite (Simplifier.context ctxt pos_numeral_ss))
boehmes@36890
   134
    Conv.no_conv
boehmes@36890
   135
in
boehmes@36890
   136
fun normalize_numerals ctxt =
boehmes@40402
   137
  map (apsnd ((Term.exists_subterm (is_strange_number ctxt) o Thm.prop_of) ??
boehmes@40402
   138
    Conv.fconv_rule (Conv.top_sweep_conv pos_conv ctxt)))
boehmes@36890
   139
end
boehmes@36890
   140
boehmes@36890
   141
boehmes@36890
   142
boehmes@36890
   143
(* embedding of standard natural number operations into integer operations *)
boehmes@36890
   144
boehmes@36890
   145
local
boehmes@40402
   146
  val nat_embedding = map (pair ~1) @{lemma
boehmes@36890
   147
    "nat (int n) = n"
boehmes@36890
   148
    "i >= 0 --> int (nat i) = i"
boehmes@36890
   149
    "i < 0 --> int (nat i) = 0"
boehmes@36890
   150
    by simp_all}
boehmes@36890
   151
boehmes@36890
   152
  val nat_rewriting = @{lemma
boehmes@36890
   153
    "0 = nat 0"
boehmes@36890
   154
    "1 = nat 1"
boehmes@40525
   155
    "(number_of :: int => nat) = (%i. nat (number_of i))"
boehmes@36890
   156
    "int (nat 0) = 0"
boehmes@36890
   157
    "int (nat 1) = 1"
boehmes@40525
   158
    "op < = (%a b. int a < int b)"
boehmes@40525
   159
    "op <= = (%a b. int a <= int b)"
boehmes@40525
   160
    "Suc = (%a. nat (int a + 1))"
boehmes@40525
   161
    "op + = (%a b. nat (int a + int b))"
boehmes@40525
   162
    "op - = (%a b. nat (int a - int b))"
boehmes@40525
   163
    "op * = (%a b. nat (int a * int b))"
boehmes@40525
   164
    "op div = (%a b. nat (int a div int b))"
boehmes@40525
   165
    "op mod = (%a b. nat (int a mod int b))"
boehmes@40525
   166
    "min = (%a b. nat (min (int a) (int b)))"
boehmes@40525
   167
    "max = (%a b. nat (max (int a) (int b)))"
boehmes@36890
   168
    "int (nat (int a + int b)) = int a + int b"
boehmes@40525
   169
    "int (nat (int a + 1)) = int a + 1"  (* special rule due to Suc above *)
boehmes@36890
   170
    "int (nat (int a * int b)) = int a * int b"
boehmes@36890
   171
    "int (nat (int a div int b)) = int a div int b"
boehmes@36890
   172
    "int (nat (int a mod int b)) = int a mod int b"
boehmes@36890
   173
    "int (nat (min (int a) (int b))) = min (int a) (int b)"
boehmes@36890
   174
    "int (nat (max (int a) (int b))) = max (int a) (int b)"
boehmes@40525
   175
    by (auto intro!: ext simp add: nat_mult_distrib nat_div_distrib
boehmes@40525
   176
      nat_mod_distrib int_mult[symmetric] zdiv_int[symmetric]
boehmes@40525
   177
      zmod_int[symmetric])}
boehmes@36890
   178
boehmes@36890
   179
  fun on_positive num f x = 
boehmes@36890
   180
    (case try HOLogic.dest_number (Thm.term_of num) of
boehmes@36890
   181
      SOME (_, i) => if i >= 0 then SOME (f x) else NONE
boehmes@36890
   182
    | NONE => NONE)
boehmes@36890
   183
boehmes@36890
   184
  val cancel_int_nat_ss = HOL_ss
boehmes@36890
   185
    addsimps [@{thm Nat_Numeral.nat_number_of}]
boehmes@36890
   186
    addsimps [@{thm Nat_Numeral.int_nat_number_of}]
boehmes@36890
   187
    addsimps @{thms neg_simps}
boehmes@36890
   188
boehmes@40817
   189
  val int_eq = Thm.cterm_of @{theory} @{const "==" (int)}
boehmes@40817
   190
boehmes@36890
   191
  fun cancel_int_nat_simproc _ ss ct = 
boehmes@36890
   192
    let
boehmes@36890
   193
      val num = Thm.dest_arg (Thm.dest_arg ct)
boehmes@40817
   194
      val goal = Thm.mk_binop int_eq ct num
boehmes@36890
   195
      val simpset = Simplifier.inherit_context ss cancel_int_nat_ss
boehmes@36890
   196
      fun tac _ = Simplifier.simp_tac simpset 1
boehmes@36890
   197
    in on_positive num (Goal.prove_internal [] goal) tac end
boehmes@36890
   198
boehmes@36890
   199
  val nat_ss = HOL_ss
boehmes@36890
   200
    addsimps nat_rewriting
boehmes@40525
   201
    addsimprocs [
boehmes@40525
   202
      Simplifier.make_simproc {
boehmes@40525
   203
        name = "cancel_int_nat_num", lhss = [@{cpat "int (nat _)"}],
boehmes@40525
   204
        proc = cancel_int_nat_simproc, identifier = [] }]
boehmes@36890
   205
boehmes@36890
   206
  fun conv ctxt = Simplifier.rewrite (Simplifier.context ctxt nat_ss)
boehmes@36890
   207
boehmes@36890
   208
  val uses_nat_type = Term.exists_type (Term.exists_subtype (equal @{typ nat}))
boehmes@40817
   209
  val uses_nat_int = Term.exists_subterm (member (op aconv)
boehmes@40817
   210
    [@{const of_nat (int)}, @{const nat}])
boehmes@41307
   211
boehmes@41307
   212
  val nat_ops = [
boehmes@41307
   213
    @{const less (nat)}, @{const less_eq (nat)},
boehmes@41307
   214
    @{const Suc}, @{const plus (nat)}, @{const minus (nat)},
boehmes@41307
   215
    @{const times (nat)}, @{const div (nat)}, @{const mod (nat)}]
boehmes@41307
   216
  val nat_ops' = @{const of_nat (int)} :: @{const nat} :: nat_ops
boehmes@36890
   217
in
boehmes@36890
   218
fun nat_as_int ctxt =
boehmes@40402
   219
  map (apsnd ((uses_nat_type o Thm.prop_of) ?? Conv.fconv_rule (conv ctxt))) #>
boehmes@40402
   220
  exists (uses_nat_int o Thm.prop_of o snd) ?? append nat_embedding
boehmes@41307
   221
boehmes@41307
   222
val setup_nat_as_int =
boehmes@41307
   223
  B.add_builtin_typ_ext (@{typ nat}, K true) #>
boehmes@41307
   224
  fold (B.add_builtin_fun_ext' o Term.dest_Const) nat_ops'
boehmes@36890
   225
end
boehmes@36890
   226
boehmes@36890
   227
boehmes@36890
   228
boehmes@36890
   229
(* further normalizations: beta/eta, universal closure, atomize *)
boehmes@36890
   230
boehmes@36890
   231
val eta_expand_eq = @{lemma "f == (%x. f x)" by (rule reflexive)}
boehmes@36890
   232
boehmes@36890
   233
fun eta_expand_conv cv ctxt =
boehmes@36890
   234
  Conv.rewr_conv eta_expand_eq then_conv Conv.abs_conv (cv o snd) ctxt
boehmes@36890
   235
boehmes@36890
   236
local
boehmes@36890
   237
  val eta_conv = eta_expand_conv
boehmes@36890
   238
boehmes@40525
   239
  fun args_conv cv ct =
boehmes@40525
   240
    (case Thm.term_of ct of
boehmes@40525
   241
      _ $ _ => Conv.combination_conv (args_conv cv) cv
boehmes@40525
   242
    | _ => Conv.all_conv) ct
boehmes@40525
   243
boehmes@40525
   244
  fun eta_args_conv cv 0 = args_conv o cv
boehmes@40525
   245
    | eta_args_conv cv i = eta_conv (eta_args_conv cv (i-1))
boehmes@40525
   246
wenzelm@36938
   247
  fun keep_conv ctxt = Conv.binder_conv (norm_conv o snd) ctxt
boehmes@36890
   248
  and eta_binder_conv ctxt = Conv.arg_conv (eta_conv norm_conv ctxt)
boehmes@36890
   249
  and keep_let_conv ctxt = Conv.combination_conv
boehmes@36890
   250
    (Conv.arg_conv (norm_conv ctxt)) (Conv.abs_conv (norm_conv o snd) ctxt)
boehmes@36890
   251
  and unfold_let_conv ctxt = Conv.combination_conv
boehmes@36890
   252
    (Conv.arg_conv (norm_conv ctxt)) (eta_conv norm_conv ctxt)
boehmes@36890
   253
  and unfold_conv thm ctxt = Conv.rewr_conv thm then_conv keep_conv ctxt
boehmes@36890
   254
  and unfold_ex1_conv ctxt = unfold_conv @{thm Ex1_def} ctxt
boehmes@37786
   255
  and unfold_ball_conv ctxt = unfold_conv (mk_meta_eq @{thm Ball_def}) ctxt
boehmes@37786
   256
  and unfold_bex_conv ctxt = unfold_conv (mk_meta_eq @{thm Bex_def}) ctxt
boehmes@36890
   257
  and norm_conv ctxt ct =
boehmes@36890
   258
    (case Thm.term_of ct of
boehmes@36890
   259
      Const (@{const_name All}, _) $ Abs _ => keep_conv
boehmes@36890
   260
    | Const (@{const_name All}, _) $ _ => eta_binder_conv
boehmes@36890
   261
    | Const (@{const_name All}, _) => eta_conv eta_binder_conv
boehmes@36890
   262
    | Const (@{const_name Ex}, _) $ Abs _ => keep_conv
boehmes@36890
   263
    | Const (@{const_name Ex}, _) $ _ => eta_binder_conv
boehmes@36890
   264
    | Const (@{const_name Ex}, _) => eta_conv eta_binder_conv
boehmes@36890
   265
    | Const (@{const_name Let}, _) $ _ $ Abs _ => keep_let_conv
boehmes@36890
   266
    | Const (@{const_name Let}, _) $ _ $ _ => unfold_let_conv
boehmes@36890
   267
    | Const (@{const_name Let}, _) $ _ => eta_conv unfold_let_conv
boehmes@36890
   268
    | Const (@{const_name Let}, _) => eta_conv (eta_conv unfold_let_conv)
boehmes@36890
   269
    | Const (@{const_name Ex1}, _) $ _ => unfold_ex1_conv
boehmes@36890
   270
    | Const (@{const_name Ex1}, _) => eta_conv unfold_ex1_conv 
boehmes@36890
   271
    | Const (@{const_name Ball}, _) $ _ $ _ => unfold_ball_conv
boehmes@36890
   272
    | Const (@{const_name Ball}, _) $ _ => eta_conv unfold_ball_conv
boehmes@36890
   273
    | Const (@{const_name Ball}, _) => eta_conv (eta_conv unfold_ball_conv)
boehmes@36890
   274
    | Const (@{const_name Bex}, _) $ _ $ _ => unfold_bex_conv
boehmes@36890
   275
    | Const (@{const_name Bex}, _) $ _ => eta_conv unfold_bex_conv
boehmes@36890
   276
    | Const (@{const_name Bex}, _) => eta_conv (eta_conv unfold_bex_conv)
boehmes@36890
   277
    | Abs _ => Conv.abs_conv (norm_conv o snd)
boehmes@40525
   278
    | _ =>
boehmes@40525
   279
        (case Term.strip_comb (Thm.term_of ct) of
boehmes@40525
   280
          (Const (c as (_, T)), ts) =>
boehmes@41307
   281
            if SMT_Builtin.is_builtin_fun ctxt c ts
boehmes@40525
   282
            then eta_args_conv norm_conv
boehmes@40525
   283
              (length (Term.binder_types T) - length ts)
boehmes@40525
   284
            else args_conv o norm_conv
boehmes@40651
   285
        | _ => args_conv o norm_conv)) ctxt ct
boehmes@36890
   286
boehmes@40525
   287
  fun is_normed ctxt t =
boehmes@36890
   288
    (case t of
boehmes@40525
   289
      Const (@{const_name All}, _) $ Abs (_, _, u) => is_normed ctxt u
boehmes@36890
   290
    | Const (@{const_name All}, _) $ _ => false
boehmes@36890
   291
    | Const (@{const_name All}, _) => false
boehmes@40525
   292
    | Const (@{const_name Ex}, _) $ Abs (_, _, u) => is_normed ctxt u
boehmes@36890
   293
    | Const (@{const_name Ex}, _) $ _ => false
boehmes@36890
   294
    | Const (@{const_name Ex}, _) => false
boehmes@36890
   295
    | Const (@{const_name Let}, _) $ u1 $ Abs (_, _, u2) =>
boehmes@40525
   296
        is_normed ctxt u1 andalso is_normed ctxt u2
boehmes@36890
   297
    | Const (@{const_name Let}, _) $ _ $ _ => false
boehmes@36890
   298
    | Const (@{const_name Let}, _) $ _ => false
boehmes@36890
   299
    | Const (@{const_name Let}, _) => false
boehmes@40525
   300
    | Const (@{const_name Ex1}, _) $ _ => false
boehmes@36890
   301
    | Const (@{const_name Ex1}, _) => false
boehmes@40525
   302
    | Const (@{const_name Ball}, _) $ _ $ _ => false
boehmes@40525
   303
    | Const (@{const_name Ball}, _) $ _ => false
boehmes@36890
   304
    | Const (@{const_name Ball}, _) => false
boehmes@40525
   305
    | Const (@{const_name Bex}, _) $ _ $ _ => false
boehmes@40525
   306
    | Const (@{const_name Bex}, _) $ _ => false
boehmes@36890
   307
    | Const (@{const_name Bex}, _) => false
boehmes@40525
   308
    | Abs (_, _, u) => is_normed ctxt u
boehmes@40525
   309
    | _ =>
boehmes@40525
   310
        (case Term.strip_comb t of
boehmes@40525
   311
          (Const (c as (_, T)), ts) =>
boehmes@41307
   312
            if SMT_Builtin.is_builtin_fun ctxt c ts
boehmes@40525
   313
            then length (Term.binder_types T) = length ts andalso
boehmes@40525
   314
              forall (is_normed ctxt) ts
boehmes@40525
   315
            else forall (is_normed ctxt) ts
boehmes@40525
   316
        | (_, ts) => forall (is_normed ctxt) ts))
boehmes@36890
   317
in
boehmes@40525
   318
fun norm_binder_conv ctxt =
boehmes@40911
   319
  U.if_conv (is_normed ctxt) Conv.all_conv (norm_conv ctxt)
boehmes@41307
   320
boehmes@41307
   321
val setup_unfolded_quants =
boehmes@41307
   322
  fold B.add_builtin_fun_ext'' [@{const_name Ball}, @{const_name Bex},
boehmes@41307
   323
    @{const_name Ex1}]
boehmes@41307
   324
boehmes@36890
   325
end
boehmes@36890
   326
boehmes@36890
   327
fun norm_def ctxt thm =
boehmes@36890
   328
  (case Thm.prop_of thm of
boehmes@40817
   329
    @{const Trueprop} $ (Const (@{const_name HOL.eq}, _) $ _ $ Abs _) =>
boehmes@36890
   330
      norm_def ctxt (thm RS @{thm fun_cong})
boehmes@36890
   331
  | Const (@{const_name "=="}, _) $ _ $ Abs _ =>
boehmes@36890
   332
      norm_def ctxt (thm RS @{thm meta_eq_to_obj_eq})
boehmes@36890
   333
  | _ => thm)
boehmes@36890
   334
boehmes@36890
   335
fun atomize_conv ctxt ct =
boehmes@36890
   336
  (case Thm.term_of ct of
boehmes@40817
   337
    @{const "==>"} $ _ $ _ =>
boehmes@36890
   338
      Conv.binop_conv (atomize_conv ctxt) then_conv
boehmes@36890
   339
      Conv.rewr_conv @{thm atomize_imp}
boehmes@36890
   340
  | Const (@{const_name "=="}, _) $ _ $ _ =>
boehmes@36890
   341
      Conv.binop_conv (atomize_conv ctxt) then_conv
boehmes@36890
   342
      Conv.rewr_conv @{thm atomize_eq}
boehmes@36890
   343
  | Const (@{const_name all}, _) $ Abs _ =>
wenzelm@36938
   344
      Conv.binder_conv (atomize_conv o snd) ctxt then_conv
boehmes@36890
   345
      Conv.rewr_conv @{thm atomize_all}
boehmes@36890
   346
  | _ => Conv.all_conv) ct
boehmes@36890
   347
boehmes@41307
   348
val setup_atomize =
boehmes@41307
   349
  fold B.add_builtin_fun_ext'' [@{const_name "==>"}, @{const_name "=="},
boehmes@41307
   350
    @{const_name all}, @{const_name Trueprop}]
boehmes@41307
   351
boehmes@36890
   352
fun normalize_rule ctxt =
boehmes@36890
   353
  Conv.fconv_rule (
boehmes@36890
   354
    (* reduce lambda abstractions, except at known binders: *)
boehmes@36890
   355
    Thm.beta_conversion true then_conv
boehmes@36890
   356
    Thm.eta_conversion then_conv
boehmes@36890
   357
    norm_binder_conv ctxt) #>
boehmes@36890
   358
  norm_def ctxt #>
boehmes@36890
   359
  Drule.forall_intr_vars #>
boehmes@36890
   360
  Conv.fconv_rule (atomize_conv ctxt)
boehmes@36890
   361
boehmes@36890
   362
boehmes@36890
   363
boehmes@36890
   364
(* lift lambda terms into additional rules *)
boehmes@36890
   365
boehmes@36890
   366
local
boehmes@36890
   367
  fun used_vars cvs ct =
boehmes@36890
   368
    let
boehmes@36890
   369
      val lookup = AList.lookup (op aconv) (map (` Thm.term_of) cvs)
boehmes@36890
   370
      val add = (fn SOME ct => insert (op aconvc) ct | _ => I)
boehmes@36890
   371
    in Term.fold_aterms (add o lookup) (Thm.term_of ct) [] end
boehmes@36890
   372
boehmes@36890
   373
  fun apply cv thm = 
boehmes@36890
   374
    let val thm' = Thm.combination thm (Thm.reflexive cv)
boehmes@36890
   375
    in Thm.transitive thm' (Thm.beta_conversion false (Thm.rhs_of thm')) end
boehmes@36890
   376
  fun apply_def cvs eq = Thm.symmetric (fold apply cvs eq)
boehmes@36890
   377
boehmes@36890
   378
  fun replace_lambda cvs ct (cx as (ctxt, defs)) =
boehmes@36890
   379
    let
boehmes@36890
   380
      val cvs' = used_vars cvs ct
boehmes@36890
   381
      val ct' = fold_rev Thm.cabs cvs' ct
boehmes@36890
   382
    in
boehmes@36890
   383
      (case Termtab.lookup defs (Thm.term_of ct') of
boehmes@36890
   384
        SOME eq => (apply_def cvs' eq, cx)
boehmes@36890
   385
      | NONE =>
boehmes@36890
   386
          let
boehmes@36890
   387
            val {T, ...} = Thm.rep_cterm ct' and n = Name.uu
boehmes@36890
   388
            val (n', ctxt') = yield_singleton Variable.variant_fixes n ctxt
boehmes@40911
   389
            val cu = U.mk_cequals (U.certify ctxt (Free (n', T))) ct'
boehmes@36890
   390
            val (eq, ctxt'') = yield_singleton Assumption.add_assumes cu ctxt'
boehmes@36890
   391
            val defs' = Termtab.update (Thm.term_of ct', eq) defs
boehmes@36890
   392
          in (apply_def cvs' eq, (ctxt'', defs')) end)
boehmes@36890
   393
    end
boehmes@36890
   394
boehmes@36890
   395
  fun none ct cx = (Thm.reflexive ct, cx)
boehmes@36890
   396
  fun in_comb f g ct cx =
boehmes@36890
   397
    let val (cu1, cu2) = Thm.dest_comb ct
boehmes@36890
   398
    in cx |> f cu1 ||>> g cu2 |>> uncurry Thm.combination end
boehmes@36890
   399
  fun in_arg f = in_comb none f
boehmes@36890
   400
  fun in_abs f cvs ct (ctxt, defs) =
boehmes@36890
   401
    let
boehmes@36890
   402
      val (n, ctxt') = yield_singleton Variable.variant_fixes Name.uu ctxt
boehmes@36890
   403
      val (cv, cu) = Thm.dest_abs (SOME n) ct
boehmes@36890
   404
    in  (ctxt', defs) |> f (cv :: cvs) cu |>> Thm.abstract_rule n cv end
boehmes@36890
   405
boehmes@36890
   406
  fun traverse cvs ct =
boehmes@36890
   407
    (case Thm.term_of ct of
boehmes@36890
   408
      Const (@{const_name All}, _) $ Abs _ => in_arg (in_abs traverse cvs)
boehmes@36890
   409
    | Const (@{const_name Ex}, _) $ Abs _ => in_arg (in_abs traverse cvs)
boehmes@36890
   410
    | Const (@{const_name Let}, _) $ _ $ Abs _ =>
boehmes@36890
   411
        in_comb (in_arg (traverse cvs)) (in_abs traverse cvs)
boehmes@36890
   412
    | Abs _ => at_lambda cvs
boehmes@36890
   413
    | _ $ _ => in_comb (traverse cvs) (traverse cvs)
boehmes@36890
   414
    | _ => none) ct
boehmes@36890
   415
boehmes@36890
   416
  and at_lambda cvs ct =
boehmes@36890
   417
    in_abs traverse cvs ct #-> (fn thm =>
boehmes@36890
   418
    replace_lambda cvs (Thm.rhs_of thm) #>> Thm.transitive thm)
boehmes@36890
   419
boehmes@36890
   420
  fun has_free_lambdas t =
boehmes@36890
   421
    (case t of
boehmes@36890
   422
      Const (@{const_name All}, _) $ Abs (_, _, u) => has_free_lambdas u
boehmes@36890
   423
    | Const (@{const_name Ex}, _) $ Abs (_, _, u) => has_free_lambdas u
boehmes@36890
   424
    | Const (@{const_name Let}, _) $ u1 $ Abs (_, _, u2) =>
boehmes@36890
   425
        has_free_lambdas u1 orelse has_free_lambdas u2
boehmes@36890
   426
    | Abs _ => true
boehmes@36890
   427
    | u1 $ u2 => has_free_lambdas u1 orelse has_free_lambdas u2
boehmes@36890
   428
    | _ => false)
boehmes@36890
   429
boehmes@36890
   430
  fun lift_lm f thm cx =
boehmes@36890
   431
    if not (has_free_lambdas (Thm.prop_of thm)) then (thm, cx)
boehmes@36890
   432
    else cx |> f (Thm.cprop_of thm) |>> (fn thm' => Thm.equal_elim thm' thm)
boehmes@36890
   433
in
boehmes@40402
   434
fun lift_lambdas irules ctxt =
boehmes@36890
   435
  let
boehmes@36890
   436
    val cx = (ctxt, Termtab.empty)
boehmes@40402
   437
    val (idxs, thms) = split_list irules
boehmes@36890
   438
    val (thms', (ctxt', defs)) = fold_map (lift_lm (traverse [])) thms cx
boehmes@36890
   439
    val eqs = Termtab.fold (cons o normalize_rule ctxt' o snd) defs []
boehmes@40402
   440
  in (map (pair ~1) eqs @ (idxs ~~ thms'), ctxt') end
boehmes@36890
   441
end
boehmes@36890
   442
boehmes@36890
   443
boehmes@36890
   444
boehmes@36890
   445
(* make application explicit for functions with varying number of arguments *)
boehmes@36890
   446
boehmes@36890
   447
local
boehmes@36890
   448
  val const = prefix "c" and free = prefix "f"
boehmes@36890
   449
  fun min i (e as (_, j)) = if i <> j then (true, Int.min (i, j)) else e
boehmes@36890
   450
  fun add t i = Symtab.map_default (t, (false, i)) (min i)
boehmes@36890
   451
  fun traverse t =
boehmes@36890
   452
    (case Term.strip_comb t of
boehmes@36890
   453
      (Const (n, _), ts) => add (const n) (length ts) #> fold traverse ts 
boehmes@36890
   454
    | (Free (n, _), ts) => add (free n) (length ts) #> fold traverse ts
boehmes@36890
   455
    | (Abs (_, _, u), ts) => fold traverse (u :: ts)
boehmes@36890
   456
    | (_, ts) => fold traverse ts)
boehmes@40402
   457
  fun prune tab = Symtab.fold (fn (n, (true, i)) =>
boehmes@40402
   458
    Symtab.update (n, i) | _ => I) tab Symtab.empty
boehmes@36890
   459
boehmes@36890
   460
  fun binop_conv cv1 cv2 = Conv.combination_conv (Conv.arg_conv cv1) cv2
boehmes@36890
   461
  fun nary_conv conv1 conv2 ct =
boehmes@36890
   462
    (Conv.combination_conv (nary_conv conv1 conv2) conv2 else_conv conv1) ct
boehmes@36890
   463
  fun abs_conv conv tb = Conv.abs_conv (fn (cv, cx) =>
boehmes@36890
   464
    let val n = fst (Term.dest_Free (Thm.term_of cv))
boehmes@36890
   465
    in conv (Symtab.update (free n, 0) tb) cx end)
boehmes@37148
   466
  val fun_app_rule = @{lemma "f x == fun_app f x" by (simp add: fun_app_def)}
boehmes@36890
   467
in
boehmes@40402
   468
fun explicit_application ctxt irules =
boehmes@36890
   469
  let
boehmes@36890
   470
    fun sub_conv tb ctxt ct =
boehmes@36890
   471
      (case Term.strip_comb (Thm.term_of ct) of
boehmes@36890
   472
        (Const (n, _), ts) => app_conv tb (const n) (length ts) ctxt
boehmes@36890
   473
      | (Free (n, _), ts) => app_conv tb (free n) (length ts) ctxt
boehmes@36890
   474
      | (Abs _, _) => nary_conv (abs_conv sub_conv tb ctxt) (sub_conv tb ctxt)
boehmes@36890
   475
      | (_, _) => nary_conv Conv.all_conv (sub_conv tb ctxt)) ct
boehmes@36890
   476
    and app_conv tb n i ctxt =
boehmes@36890
   477
      (case Symtab.lookup tb n of
boehmes@36890
   478
        NONE => nary_conv Conv.all_conv (sub_conv tb ctxt)
boehmes@37148
   479
      | SOME j => fun_app_conv tb ctxt (i - j))
boehmes@37148
   480
    and fun_app_conv tb ctxt i ct = (
boehmes@36890
   481
      if i = 0 then nary_conv Conv.all_conv (sub_conv tb ctxt)
boehmes@36890
   482
      else
boehmes@37148
   483
        Conv.rewr_conv fun_app_rule then_conv
boehmes@37148
   484
        binop_conv (fun_app_conv tb ctxt (i-1)) (sub_conv tb ctxt)) ct
boehmes@36890
   485
boehmes@36890
   486
    fun needs_exp_app tab = Term.exists_subterm (fn
boehmes@36890
   487
        Bound _ $ _ => true
boehmes@36890
   488
      | Const (n, _) => Symtab.defined tab (const n)
boehmes@36890
   489
      | Free (n, _) => Symtab.defined tab (free n)
boehmes@36890
   490
      | _ => false)
boehmes@36890
   491
boehmes@36890
   492
    fun rewrite tab ctxt thm =
boehmes@36890
   493
      if not (needs_exp_app tab (Thm.prop_of thm)) then thm
boehmes@36890
   494
      else Conv.fconv_rule (sub_conv tab ctxt) thm
boehmes@36890
   495
boehmes@40402
   496
    val tab = prune (fold (traverse o Thm.prop_of o snd) irules Symtab.empty)
boehmes@40402
   497
  in map (apsnd (rewrite tab ctxt)) irules end
boehmes@36890
   498
end
boehmes@36890
   499
boehmes@36890
   500
boehmes@36890
   501
boehmes@39723
   502
(* add missing datatype selectors via hypothetical definitions *)
boehmes@39723
   503
boehmes@39723
   504
local
boehmes@39723
   505
  val add = (fn Type (n, _) => Symtab.update (n, ()) | _ => I)
boehmes@39723
   506
boehmes@39723
   507
  fun collect t =
boehmes@39723
   508
    (case Term.strip_comb t of
boehmes@39723
   509
      (Abs (_, T, t), _) => add T #> collect t
boehmes@39723
   510
    | (Const (_, T), ts) => collects T ts
boehmes@39723
   511
    | (Free (_, T), ts) => collects T ts
boehmes@39723
   512
    | _ => I)
boehmes@39723
   513
  and collects T ts =
boehmes@39723
   514
    let val ((Ts, Us), U) = Term.strip_type T |> apfst (chop (length ts))
boehmes@39723
   515
    in fold add Ts #> add (Us ---> U) #> fold collect ts end
boehmes@39723
   516
boehmes@39723
   517
  fun add_constructors thy n =
boehmes@39723
   518
    (case Datatype.get_info thy n of
boehmes@39723
   519
      NONE => I
boehmes@39723
   520
    | SOME {descr, ...} => fold (fn (_, (_, _, cs)) => fold (fn (n, ds) =>
boehmes@39723
   521
        fold (insert (op =) o pair n) (1 upto length ds)) cs) descr)
boehmes@39723
   522
boehmes@39723
   523
  fun add_selector (c as (n, i)) ctxt =
boehmes@39723
   524
    (case Datatype_Selectors.lookup_selector ctxt c of
boehmes@39723
   525
      SOME _ => ctxt
boehmes@39723
   526
    | NONE =>
boehmes@39723
   527
        let
boehmes@39723
   528
          val T = Sign.the_const_type (ProofContext.theory_of ctxt) n
boehmes@39723
   529
          val U = Term.body_type T --> nth (Term.binder_types T) (i-1)
boehmes@39723
   530
        in
boehmes@39723
   531
          ctxt
boehmes@39723
   532
          |> yield_singleton Variable.variant_fixes Name.uu
boehmes@39723
   533
          |>> pair ((n, T), i) o rpair U
boehmes@39723
   534
          |-> Context.proof_map o Datatype_Selectors.add_selector
boehmes@39723
   535
        end)
boehmes@39723
   536
in
boehmes@39723
   537
boehmes@40402
   538
fun datatype_selectors irules ctxt =
boehmes@39723
   539
  let
boehmes@40402
   540
    val ns = Symtab.keys (fold (collect o Thm.prop_of o snd) irules Symtab.empty)
boehmes@39723
   541
    val cs = fold (add_constructors (ProofContext.theory_of ctxt)) ns []
boehmes@40402
   542
  in (irules, fold add_selector cs ctxt) end
boehmes@39723
   543
    (* FIXME: also generate hypothetical definitions for the selectors *)
boehmes@39723
   544
boehmes@39723
   545
end
boehmes@39723
   546
boehmes@39723
   547
boehmes@39723
   548
boehmes@36890
   549
(* combined normalization *)
boehmes@36890
   550
boehmes@40403
   551
type extra_norm = bool -> (int * thm) list -> Proof.context ->
boehmes@40402
   552
  (int * thm) list * Proof.context
boehmes@36890
   553
boehmes@40402
   554
fun with_context f irules ctxt = (f ctxt irules, ctxt)
boehmes@36890
   555
boehmes@40651
   556
fun normalize extra_norm with_datatypes irules ctxt =
boehmes@40524
   557
  let
boehmes@40524
   558
    fun norm f ctxt' (i, thm) =
boehmes@40651
   559
      if Config.get ctxt' SMT_Config.drop_bad_facts then
boehmes@40524
   560
        (case try (f ctxt') thm of
boehmes@40524
   561
          SOME thm' => SOME (i, thm')
boehmes@40651
   562
        | NONE => (SMT_Config.verbose_msg ctxt' (prefix ("Warning: " ^
boehmes@40524
   563
            "dropping assumption: ") o Display.string_of_thm ctxt') thm; NONE))
boehmes@40651
   564
      else SOME (i, f ctxt' thm)
boehmes@40524
   565
  in
boehmes@40524
   566
    irules
boehmes@40933
   567
    |> map (apsnd instantiate_elim)
boehmes@40524
   568
    |> trivial_distinct ctxt
boehmes@40524
   569
    |> rewrite_bool_cases ctxt
boehmes@40524
   570
    |> normalize_numerals ctxt
boehmes@40524
   571
    |> nat_as_int ctxt
boehmes@40524
   572
    |> rpair ctxt
boehmes@40524
   573
    |-> extra_norm with_datatypes
boehmes@40524
   574
    |-> with_context (map_filter o norm normalize_rule)
boehmes@40524
   575
    |-> SMT_Monomorph.monomorph
boehmes@40524
   576
    |-> lift_lambdas
boehmes@40524
   577
    |-> with_context explicit_application
boehmes@40524
   578
    |-> (if with_datatypes then datatype_selectors else pair)
boehmes@40524
   579
  end
boehmes@36890
   580
boehmes@41307
   581
boehmes@41307
   582
boehmes@41307
   583
(* setup *)
boehmes@41307
   584
boehmes@41307
   585
val setup =
boehmes@41307
   586
  setup_bool_case #>
boehmes@41307
   587
  setup_nat_as_int #>
boehmes@41307
   588
  setup_unfolded_quants #>
boehmes@41307
   589
  setup_atomize
boehmes@41307
   590
boehmes@36890
   591
end