src/HOL/SMT/Tools/smt_monomorph.ML
author haftmann
Fri, 07 May 2010 15:05:52 +0200
changeset 36751 7f1da69cacb3
parent 34010 ac78f5cdc430
child 36896 c030819254d3
permissions -rw-r--r--
split of semiring normalization from Groebner theory; moved field_comp_conv to Numeral_Simproces
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
32618
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
     1
(*  Title:      HOL/SMT/Tools/smt_monomorph.ML
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
     2
    Author:     Sascha Boehme, TU Muenchen
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
     3
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
     4
Monomorphization of terms, i.e., computation of all (necessary) instances.
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
     5
*)
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
     6
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
     7
signature SMT_MONOMORPH =
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
     8
sig
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
     9
  val monomorph: theory -> term list -> term list
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    10
end
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    11
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    12
structure SMT_Monomorph: SMT_MONOMORPH =
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    13
struct
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    14
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    15
fun selection [] = []
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    16
  | selection (x :: xs) = (x, xs) :: map (apsnd (cons x)) (selection xs)
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    17
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    18
fun permute [] = []
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    19
  | permute [x] = [[x]]
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    20
  | permute xs = maps (fn (y, ys) => map (cons y) (permute ys)) (selection xs)
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    21
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    22
fun fold_all f = fold (fn x => maps (f x))
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    23
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    24
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    25
val typ_has_tvars = Term.exists_subtype (fn TVar _ => true | _ => false)
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    26
val term_has_tvars = Term.exists_type typ_has_tvars
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    27
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    28
val ignored = member (op =) [
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    29
  @{const_name All}, @{const_name Ex}, @{const_name Let}, @{const_name If},
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    30
  @{const_name "op ="}, @{const_name zero_class.zero},
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    31
  @{const_name one_class.one}, @{const_name number_of}]
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    32
fun consts_of ts = AList.group (op =) (fold Term.add_consts ts [])
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    33
  |> filter_out (ignored o fst)
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    34
33047
69780aef0531 proper handling of single literal case,
boehmes
parents: 33038
diff changeset
    35
fun join_consts cs ds = AList.join (op =) (K (merge (op =))) (cs, ds)
32618
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    36
fun diff_consts cs ds = 
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    37
  let fun diff (n, Ts) =
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    38
    (case AList.lookup (op =) cs n of
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    39
      NONE => SOME (n, Ts)
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    40
    | SOME Us =>
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    41
        let val Ts' = fold (remove (op =)) Us Ts
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    42
        in if null Ts' then NONE else SOME (n, Ts') end)
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    43
  in map_filter diff ds end
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    44
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    45
fun instances thy is (n, Ts) env =
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    46
  let
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    47
    val Us = these (AList.lookup (op =) is n)
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    48
    val Ts' = filter typ_has_tvars (map (Envir.subst_type env) Ts)
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    49
  in
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    50
    (case map_product pair Ts' Us of
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    51
      [] => [env]
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    52
    | TUs => map_filter (try (fn TU => Sign.typ_match thy TU env)) TUs)
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    53
  end
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    54
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    55
fun proper_match ps env =
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    56
  forall (forall (not o typ_has_tvars o Envir.subst_type env) o snd) ps
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    57
33047
69780aef0531 proper handling of single literal case,
boehmes
parents: 33038
diff changeset
    58
fun eq_tab (tab1, tab2) = eq_set (op =) (Vartab.dest tab1, Vartab.dest tab2)
32618
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    59
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    60
fun specialize thy cs is ((r, ps), ces) (ts, ns) =
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    61
  let
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    62
    val ps' = filter (AList.defined (op =) is o fst) ps
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    63
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    64
    val envs = permute ps'
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    65
      |> maps (fn ps => fold_all (instances thy is) ps [Vartab.empty])
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    66
      |> filter (proper_match ps')
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    67
      |> filter_out (member eq_tab ces)
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    68
      |> distinct eq_tab
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    69
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    70
    val us = map (fn env => Envir.subst_term_types env r) envs
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    71
    val ns' = join_consts (diff_consts is (diff_consts cs (consts_of us))) ns
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    72
  in (envs @ ces, (fold (insert (op aconv)) us ts, ns')) end
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    73
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    74
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    75
fun incr_tvar_indices i t =
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    76
  let
34010
ac78f5cdc430 faster preprocessing: before applying a step, test if it is applicable (normalization of binders, unfolding of abs/min/max definitions, lambda lifting, explicit application, monomorphization),
boehmes
parents: 33047
diff changeset
    77
    val incrT = Logic.incr_tvar_same i
32618
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    78
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    79
    fun incr t =
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    80
      (case t of
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    81
        Const (n, T) => Const (n, incrT T)
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    82
      | Free (n, T) => Free (n, incrT T)
34010
ac78f5cdc430 faster preprocessing: before applying a step, test if it is applicable (normalization of binders, unfolding of abs/min/max definitions, lambda lifting, explicit application, monomorphization),
boehmes
parents: 33047
diff changeset
    83
      | Abs (n, T, t1) => (Abs (n, incrT T, incr t1 handle Same.SAME => t1)
ac78f5cdc430 faster preprocessing: before applying a step, test if it is applicable (normalization of binders, unfolding of abs/min/max definitions, lambda lifting, explicit application, monomorphization),
boehmes
parents: 33047
diff changeset
    84
          handle Same.SAME => Abs (n, T, incr t1))
ac78f5cdc430 faster preprocessing: before applying a step, test if it is applicable (normalization of binders, unfolding of abs/min/max definitions, lambda lifting, explicit application, monomorphization),
boehmes
parents: 33047
diff changeset
    85
      | t1 $ t2 => (incr t1 $ (incr t2 handle Same.SAME => t2)
ac78f5cdc430 faster preprocessing: before applying a step, test if it is applicable (normalization of binders, unfolding of abs/min/max definitions, lambda lifting, explicit application, monomorphization),
boehmes
parents: 33047
diff changeset
    86
          handle Same.SAME => t1 $ incr t2)
ac78f5cdc430 faster preprocessing: before applying a step, test if it is applicable (normalization of binders, unfolding of abs/min/max definitions, lambda lifting, explicit application, monomorphization),
boehmes
parents: 33047
diff changeset
    87
      | _ => Same.same t)
ac78f5cdc430 faster preprocessing: before applying a step, test if it is applicable (normalization of binders, unfolding of abs/min/max definitions, lambda lifting, explicit application, monomorphization),
boehmes
parents: 33047
diff changeset
    88
  in incr t handle Same.SAME => t end
32618
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    89
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    90
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    91
val monomorph_limit = 10
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    92
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    93
(* Instantiate all polymorphic constants (i.e., constants occurring both with
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    94
   ground types and type variables) with all (necessary) ground types; thereby
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    95
   create copies of terms containing those constants.
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    96
   To prevent non-termination, there is an upper limit for the number of
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    97
   recursions involved in the fixpoint construction. *)
34010
ac78f5cdc430 faster preprocessing: before applying a step, test if it is applicable (normalization of binders, unfolding of abs/min/max definitions, lambda lifting, explicit application, monomorphization),
boehmes
parents: 33047
diff changeset
    98
fun monomorph thy =
32618
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
    99
  let
34010
ac78f5cdc430 faster preprocessing: before applying a step, test if it is applicable (normalization of binders, unfolding of abs/min/max definitions, lambda lifting, explicit application, monomorphization),
boehmes
parents: 33047
diff changeset
   100
    fun incr t idx = (incr_tvar_indices idx t, idx + Term.maxidx_of_term t + 1)
ac78f5cdc430 faster preprocessing: before applying a step, test if it is applicable (normalization of binders, unfolding of abs/min/max definitions, lambda lifting, explicit application, monomorphization),
boehmes
parents: 33047
diff changeset
   101
    fun incr_indices ts = fst (fold_map incr ts 0)
32618
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   102
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   103
    fun with_tvar (n, Ts) =
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   104
      let val Ts' = filter typ_has_tvars Ts
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   105
      in if null Ts' then NONE else SOME (n, Ts') end
34010
ac78f5cdc430 faster preprocessing: before applying a step, test if it is applicable (normalization of binders, unfolding of abs/min/max definitions, lambda lifting, explicit application, monomorphization),
boehmes
parents: 33047
diff changeset
   106
    fun extract_consts_with_tvar t = (t, map_filter with_tvar (consts_of [t]))
32618
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   107
34010
ac78f5cdc430 faster preprocessing: before applying a step, test if it is applicable (normalization of binders, unfolding of abs/min/max definitions, lambda lifting, explicit application, monomorphization),
boehmes
parents: 33047
diff changeset
   108
    fun mono rps count is ces cs ts =
32618
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   109
      let
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   110
        val spec = specialize thy cs is
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   111
        val (ces', (ts', is')) = fold_map spec (rps ~~ ces) (ts, [])
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   112
        val cs' = join_consts is cs
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   113
      in
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   114
        if null is' then ts'
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   115
        else if count > monomorph_limit then
32950
5d5e123443b3 normalized aliases of Output operations;
wenzelm
parents: 32618
diff changeset
   116
          (warning "monomorphization limit reached"; ts')
34010
ac78f5cdc430 faster preprocessing: before applying a step, test if it is applicable (normalization of binders, unfolding of abs/min/max definitions, lambda lifting, explicit application, monomorphization),
boehmes
parents: 33047
diff changeset
   117
        else mono rps (count + 1) is' ces' cs' ts'
32618
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   118
      end
34010
ac78f5cdc430 faster preprocessing: before applying a step, test if it is applicable (normalization of binders, unfolding of abs/min/max definitions, lambda lifting, explicit application, monomorphization),
boehmes
parents: 33047
diff changeset
   119
    fun mono_all rps ms = if null rps then ms
ac78f5cdc430 faster preprocessing: before applying a step, test if it is applicable (normalization of binders, unfolding of abs/min/max definitions, lambda lifting, explicit application, monomorphization),
boehmes
parents: 33047
diff changeset
   120
      else mono rps 0 (consts_of ms) (map (K []) rps) [] ms
ac78f5cdc430 faster preprocessing: before applying a step, test if it is applicable (normalization of binders, unfolding of abs/min/max definitions, lambda lifting, explicit application, monomorphization),
boehmes
parents: 33047
diff changeset
   121
  in
ac78f5cdc430 faster preprocessing: before applying a step, test if it is applicable (normalization of binders, unfolding of abs/min/max definitions, lambda lifting, explicit application, monomorphization),
boehmes
parents: 33047
diff changeset
   122
    List.partition term_has_tvars
ac78f5cdc430 faster preprocessing: before applying a step, test if it is applicable (normalization of binders, unfolding of abs/min/max definitions, lambda lifting, explicit application, monomorphization),
boehmes
parents: 33047
diff changeset
   123
    #>> incr_indices
ac78f5cdc430 faster preprocessing: before applying a step, test if it is applicable (normalization of binders, unfolding of abs/min/max definitions, lambda lifting, explicit application, monomorphization),
boehmes
parents: 33047
diff changeset
   124
    #>> map extract_consts_with_tvar
ac78f5cdc430 faster preprocessing: before applying a step, test if it is applicable (normalization of binders, unfolding of abs/min/max definitions, lambda lifting, explicit application, monomorphization),
boehmes
parents: 33047
diff changeset
   125
    #-> mono_all
ac78f5cdc430 faster preprocessing: before applying a step, test if it is applicable (normalization of binders, unfolding of abs/min/max definitions, lambda lifting, explicit application, monomorphization),
boehmes
parents: 33047
diff changeset
   126
  end
32618
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   127
42865636d006 added new method "smt": an oracle-based connection to external SMT solvers
boehmes
parents:
diff changeset
   128
end