src/HOL/Tools/refute.ML
changeset 49985 5b4b0e4e5205
parent 49984 9f655a6bffd8
child 49986 90e7be285b49
equal deleted inserted replaced
49984:9f655a6bffd8 49985:5b4b0e4e5205
     1 (*  Title:      HOL/Tools/refute.ML
       
     2     Author:     Tjark Weber, TU Muenchen
       
     3 
       
     4 Finite model generation for HOL formulas, using a SAT solver.
       
     5 *)
       
     6 
       
     7 (* ------------------------------------------------------------------------- *)
       
     8 (* Declares the 'REFUTE' signature as well as a structure 'Refute'.          *)
       
     9 (* Documentation is available in the Isabelle/Isar theory 'HOL/Refute.thy'.  *)
       
    10 (* ------------------------------------------------------------------------- *)
       
    11 
       
    12 signature REFUTE =
       
    13 sig
       
    14 
       
    15   exception REFUTE of string * string
       
    16 
       
    17 (* ------------------------------------------------------------------------- *)
       
    18 (* Model/interpretation related code (translation HOL -> propositional logic *)
       
    19 (* ------------------------------------------------------------------------- *)
       
    20 
       
    21   type params
       
    22   type interpretation
       
    23   type model
       
    24   type arguments
       
    25 
       
    26   exception MAXVARS_EXCEEDED
       
    27 
       
    28   val add_interpreter : string -> (Proof.context -> model -> arguments -> term ->
       
    29     (interpretation * model * arguments) option) -> theory -> theory
       
    30   val add_printer : string -> (Proof.context -> model -> typ ->
       
    31     interpretation -> (int -> bool) -> term option) -> theory -> theory
       
    32 
       
    33   val interpret : Proof.context -> model -> arguments -> term ->
       
    34     (interpretation * model * arguments)
       
    35 
       
    36   val print : Proof.context -> model -> typ -> interpretation -> (int -> bool) -> term
       
    37   val print_model : Proof.context -> model -> (int -> bool) -> string
       
    38 
       
    39 (* ------------------------------------------------------------------------- *)
       
    40 (* Interface                                                                 *)
       
    41 (* ------------------------------------------------------------------------- *)
       
    42 
       
    43   val set_default_param  : (string * string) -> theory -> theory
       
    44   val get_default_param  : Proof.context -> string -> string option
       
    45   val get_default_params : Proof.context -> (string * string) list
       
    46   val actual_params      : Proof.context -> (string * string) list -> params
       
    47 
       
    48   val find_model :
       
    49     Proof.context -> params -> term list -> term -> bool -> string
       
    50 
       
    51   (* tries to find a model for a formula: *)
       
    52   val satisfy_term :
       
    53     Proof.context -> (string * string) list -> term list -> term -> string
       
    54   (* tries to find a model that refutes a formula: *)
       
    55   val refute_term :
       
    56     Proof.context -> (string * string) list -> term list -> term -> string
       
    57   val refute_goal :
       
    58     Proof.context -> (string * string) list -> thm -> int -> string
       
    59 
       
    60   val setup : theory -> theory
       
    61 
       
    62 (* ------------------------------------------------------------------------- *)
       
    63 (* Additional functions used by Nitpick (to be factored out)                 *)
       
    64 (* ------------------------------------------------------------------------- *)
       
    65 
       
    66   val get_classdef : theory -> string -> (string * term) option
       
    67   val norm_rhs : term -> term
       
    68   val get_def : theory -> string * typ -> (string * term) option
       
    69   val get_typedef : theory -> typ -> (string * term) option
       
    70   val is_IDT_constructor : theory -> string * typ -> bool
       
    71   val is_IDT_recursor : theory -> string * typ -> bool
       
    72   val is_const_of_class: theory -> string * typ -> bool
       
    73   val string_of_typ : typ -> string
       
    74 end;
       
    75 
       
    76 structure Refute : REFUTE =
       
    77 struct
       
    78 
       
    79 open Prop_Logic;
       
    80 
       
    81 (* We use 'REFUTE' only for internal error conditions that should    *)
       
    82 (* never occur in the first place (i.e. errors caused by bugs in our *)
       
    83 (* code).  Otherwise (e.g. to indicate invalid input data) we use    *)
       
    84 (* 'error'.                                                          *)
       
    85 exception REFUTE of string * string;  (* ("in function", "cause") *)
       
    86 
       
    87 (* should be raised by an interpreter when more variables would be *)
       
    88 (* required than allowed by 'maxvars'                              *)
       
    89 exception MAXVARS_EXCEEDED;
       
    90 
       
    91 
       
    92 (* ------------------------------------------------------------------------- *)
       
    93 (* TREES                                                                     *)
       
    94 (* ------------------------------------------------------------------------- *)
       
    95 
       
    96 (* ------------------------------------------------------------------------- *)
       
    97 (* tree: implements an arbitrarily (but finitely) branching tree as a list   *)
       
    98 (*       of (lists of ...) elements                                          *)
       
    99 (* ------------------------------------------------------------------------- *)
       
   100 
       
   101 datatype 'a tree =
       
   102     Leaf of 'a
       
   103   | Node of ('a tree) list;
       
   104 
       
   105 (* ('a -> 'b) -> 'a tree -> 'b tree *)
       
   106 
       
   107 fun tree_map f tr =
       
   108   case tr of
       
   109     Leaf x  => Leaf (f x)
       
   110   | Node xs => Node (map (tree_map f) xs);
       
   111 
       
   112 (* ('a * 'b -> 'a) -> 'a * ('b tree) -> 'a *)
       
   113 
       
   114 fun tree_foldl f =
       
   115   let
       
   116     fun itl (e, Leaf x)  = f(e,x)
       
   117       | itl (e, Node xs) = Library.foldl (tree_foldl f) (e,xs)
       
   118   in
       
   119     itl
       
   120   end;
       
   121 
       
   122 (* 'a tree * 'b tree -> ('a * 'b) tree *)
       
   123 
       
   124 fun tree_pair (t1, t2) =
       
   125   case t1 of
       
   126     Leaf x =>
       
   127       (case t2 of
       
   128           Leaf y => Leaf (x,y)
       
   129         | Node _ => raise REFUTE ("tree_pair",
       
   130             "trees are of different height (second tree is higher)"))
       
   131   | Node xs =>
       
   132       (case t2 of
       
   133           (* '~~' will raise an exception if the number of branches in   *)
       
   134           (* both trees is different at the current node                 *)
       
   135           Node ys => Node (map tree_pair (xs ~~ ys))
       
   136         | Leaf _  => raise REFUTE ("tree_pair",
       
   137             "trees are of different height (first tree is higher)"));
       
   138 
       
   139 (* ------------------------------------------------------------------------- *)
       
   140 (* params: parameters that control the translation into a propositional      *)
       
   141 (*         formula/model generation                                          *)
       
   142 (*                                                                           *)
       
   143 (* The following parameters are supported (and required (!), except for      *)
       
   144 (* "sizes" and "expect"):                                                    *)
       
   145 (*                                                                           *)
       
   146 (* Name          Type    Description                                         *)
       
   147 (*                                                                           *)
       
   148 (* "sizes"       (string * int) list                                         *)
       
   149 (*                       Size of ground types (e.g. 'a=2), or depth of IDTs. *)
       
   150 (* "minsize"     int     If >0, minimal size of each ground type/IDT depth.  *)
       
   151 (* "maxsize"     int     If >0, maximal size of each ground type/IDT depth.  *)
       
   152 (* "maxvars"     int     If >0, use at most 'maxvars' Boolean variables      *)
       
   153 (*                       when transforming the term into a propositional     *)
       
   154 (*                       formula.                                            *)
       
   155 (* "maxtime"     int     If >0, terminate after at most 'maxtime' seconds.   *)
       
   156 (* "satsolver"   string  SAT solver to be used.                              *)
       
   157 (* "no_assms"    bool    If "true", assumptions in structured proofs are     *)
       
   158 (*                       not considered.                                     *)
       
   159 (* "expect"      string  Expected result ("genuine", "potential", "none", or *)
       
   160 (*                       "unknown").                                         *)
       
   161 (* ------------------------------------------------------------------------- *)
       
   162 
       
   163 type params =
       
   164   {
       
   165     sizes    : (string * int) list,
       
   166     minsize  : int,
       
   167     maxsize  : int,
       
   168     maxvars  : int,
       
   169     maxtime  : int,
       
   170     satsolver: string,
       
   171     no_assms : bool,
       
   172     expect   : string
       
   173   };
       
   174 
       
   175 (* ------------------------------------------------------------------------- *)
       
   176 (* interpretation: a term's interpretation is given by a variable of type    *)
       
   177 (*                 'interpretation'                                          *)
       
   178 (* ------------------------------------------------------------------------- *)
       
   179 
       
   180 type interpretation =
       
   181   prop_formula list tree;
       
   182 
       
   183 (* ------------------------------------------------------------------------- *)
       
   184 (* model: a model specifies the size of types and the interpretation of      *)
       
   185 (*        terms                                                              *)
       
   186 (* ------------------------------------------------------------------------- *)
       
   187 
       
   188 type model =
       
   189   (typ * int) list * (term * interpretation) list;
       
   190 
       
   191 (* ------------------------------------------------------------------------- *)
       
   192 (* arguments: additional arguments required during interpretation of terms   *)
       
   193 (* ------------------------------------------------------------------------- *)
       
   194 
       
   195 type arguments =
       
   196   {
       
   197     (* just passed unchanged from 'params': *)
       
   198     maxvars   : int,
       
   199     (* whether to use 'make_equality' or 'make_def_equality': *)
       
   200     def_eq    : bool,
       
   201     (* the following may change during the translation: *)
       
   202     next_idx  : int,
       
   203     bounds    : interpretation list,
       
   204     wellformed: prop_formula
       
   205   };
       
   206 
       
   207 structure Data = Theory_Data
       
   208 (
       
   209   type T =
       
   210     {interpreters: (string * (Proof.context -> model -> arguments -> term ->
       
   211       (interpretation * model * arguments) option)) list,
       
   212      printers: (string * (Proof.context -> model -> typ -> interpretation ->
       
   213       (int -> bool) -> term option)) list,
       
   214      parameters: string Symtab.table};
       
   215   val empty = {interpreters = [], printers = [], parameters = Symtab.empty};
       
   216   val extend = I;
       
   217   fun merge
       
   218     ({interpreters = in1, printers = pr1, parameters = pa1},
       
   219      {interpreters = in2, printers = pr2, parameters = pa2}) : T =
       
   220     {interpreters = AList.merge (op =) (K true) (in1, in2),
       
   221      printers = AList.merge (op =) (K true) (pr1, pr2),
       
   222      parameters = Symtab.merge (op =) (pa1, pa2)};
       
   223 );
       
   224 
       
   225 val get_data = Data.get o Proof_Context.theory_of;
       
   226 
       
   227 
       
   228 (* ------------------------------------------------------------------------- *)
       
   229 (* interpret: interprets the term 't' using a suitable interpreter; returns  *)
       
   230 (*            the interpretation and a (possibly extended) model that keeps  *)
       
   231 (*            track of the interpretation of subterms                        *)
       
   232 (* ------------------------------------------------------------------------- *)
       
   233 
       
   234 fun interpret ctxt model args t =
       
   235   case get_first (fn (_, f) => f ctxt model args t)
       
   236       (#interpreters (get_data ctxt)) of
       
   237     NONE => raise REFUTE ("interpret",
       
   238       "no interpreter for term " ^ quote (Syntax.string_of_term ctxt t))
       
   239   | SOME x => x;
       
   240 
       
   241 (* ------------------------------------------------------------------------- *)
       
   242 (* print: converts the interpretation 'intr', which must denote a term of    *)
       
   243 (*        type 'T', into a term using a suitable printer                     *)
       
   244 (* ------------------------------------------------------------------------- *)
       
   245 
       
   246 fun print ctxt model T intr assignment =
       
   247   case get_first (fn (_, f) => f ctxt model T intr assignment)
       
   248       (#printers (get_data ctxt)) of
       
   249     NONE => raise REFUTE ("print",
       
   250       "no printer for type " ^ quote (Syntax.string_of_typ ctxt T))
       
   251   | SOME x => x;
       
   252 
       
   253 (* ------------------------------------------------------------------------- *)
       
   254 (* print_model: turns the model into a string, using a fixed interpretation  *)
       
   255 (*              (given by an assignment for Boolean variables) and suitable  *)
       
   256 (*              printers                                                     *)
       
   257 (* ------------------------------------------------------------------------- *)
       
   258 
       
   259 fun print_model ctxt model assignment =
       
   260   let
       
   261     val (typs, terms) = model
       
   262     val typs_msg =
       
   263       if null typs then
       
   264         "empty universe (no type variables in term)\n"
       
   265       else
       
   266         "Size of types: " ^ commas (map (fn (T, i) =>
       
   267           Syntax.string_of_typ ctxt T ^ ": " ^ string_of_int i) typs) ^ "\n"
       
   268     val show_consts_msg =
       
   269       if not (Config.get ctxt show_consts) andalso Library.exists (is_Const o fst) terms then
       
   270         "enable \"show_consts\" to show the interpretation of constants\n"
       
   271       else
       
   272         ""
       
   273     val terms_msg =
       
   274       if null terms then
       
   275         "empty interpretation (no free variables in term)\n"
       
   276       else
       
   277         cat_lines (map_filter (fn (t, intr) =>
       
   278           (* print constants only if 'show_consts' is true *)
       
   279           if Config.get ctxt show_consts orelse not (is_Const t) then
       
   280             SOME (Syntax.string_of_term ctxt t ^ ": " ^
       
   281               Syntax.string_of_term ctxt
       
   282                 (print ctxt model (Term.type_of t) intr assignment))
       
   283           else
       
   284             NONE) terms) ^ "\n"
       
   285   in
       
   286     typs_msg ^ show_consts_msg ^ terms_msg
       
   287   end;
       
   288 
       
   289 
       
   290 (* ------------------------------------------------------------------------- *)
       
   291 (* PARAMETER MANAGEMENT                                                      *)
       
   292 (* ------------------------------------------------------------------------- *)
       
   293 
       
   294 fun add_interpreter name f = Data.map (fn {interpreters, printers, parameters} =>
       
   295   case AList.lookup (op =) interpreters name of
       
   296     NONE => {interpreters = (name, f) :: interpreters,
       
   297       printers = printers, parameters = parameters}
       
   298   | SOME _ => error ("Interpreter " ^ name ^ " already declared"));
       
   299 
       
   300 fun add_printer name f = Data.map (fn {interpreters, printers, parameters} =>
       
   301   case AList.lookup (op =) printers name of
       
   302     NONE => {interpreters = interpreters,
       
   303       printers = (name, f) :: printers, parameters = parameters}
       
   304   | SOME _ => error ("Printer " ^ name ^ " already declared"));
       
   305 
       
   306 (* ------------------------------------------------------------------------- *)
       
   307 (* set_default_param: stores the '(name, value)' pair in Data's              *)
       
   308 (*                    parameter table                                        *)
       
   309 (* ------------------------------------------------------------------------- *)
       
   310 
       
   311 fun set_default_param (name, value) = Data.map
       
   312   (fn {interpreters, printers, parameters} =>
       
   313     {interpreters = interpreters, printers = printers,
       
   314       parameters = Symtab.update (name, value) parameters});
       
   315 
       
   316 (* ------------------------------------------------------------------------- *)
       
   317 (* get_default_param: retrieves the value associated with 'name' from        *)
       
   318 (*                    Data's parameter table                                 *)
       
   319 (* ------------------------------------------------------------------------- *)
       
   320 
       
   321 val get_default_param = Symtab.lookup o #parameters o get_data;
       
   322 
       
   323 (* ------------------------------------------------------------------------- *)
       
   324 (* get_default_params: returns a list of all '(name, value)' pairs that are  *)
       
   325 (*                     stored in Data's parameter table                      *)
       
   326 (* ------------------------------------------------------------------------- *)
       
   327 
       
   328 val get_default_params = Symtab.dest o #parameters o get_data;
       
   329 
       
   330 (* ------------------------------------------------------------------------- *)
       
   331 (* actual_params: takes a (possibly empty) list 'params' of parameters that  *)
       
   332 (*      override the default parameters currently specified, and             *)
       
   333 (*      returns a record that can be passed to 'find_model'.                 *)
       
   334 (* ------------------------------------------------------------------------- *)
       
   335 
       
   336 fun actual_params ctxt override =
       
   337   let
       
   338     (* (string * string) list * string -> bool *)
       
   339     fun read_bool (parms, name) =
       
   340       case AList.lookup (op =) parms name of
       
   341         SOME "true" => true
       
   342       | SOME "false" => false
       
   343       | SOME s => error ("parameter " ^ quote name ^
       
   344           " (value is " ^ quote s ^ ") must be \"true\" or \"false\"")
       
   345       | NONE   => error ("parameter " ^ quote name ^
       
   346           " must be assigned a value")
       
   347     (* (string * string) list * string -> int *)
       
   348     fun read_int (parms, name) =
       
   349       case AList.lookup (op =) parms name of
       
   350         SOME s =>
       
   351           (case Int.fromString s of
       
   352             SOME i => i
       
   353           | NONE   => error ("parameter " ^ quote name ^
       
   354             " (value is " ^ quote s ^ ") must be an integer value"))
       
   355       | NONE => error ("parameter " ^ quote name ^
       
   356           " must be assigned a value")
       
   357     (* (string * string) list * string -> string *)
       
   358     fun read_string (parms, name) =
       
   359       case AList.lookup (op =) parms name of
       
   360         SOME s => s
       
   361       | NONE => error ("parameter " ^ quote name ^
       
   362         " must be assigned a value")
       
   363     (* 'override' first, defaults last: *)
       
   364     (* (string * string) list *)
       
   365     val allparams = override @ get_default_params ctxt
       
   366     (* int *)
       
   367     val minsize = read_int (allparams, "minsize")
       
   368     val maxsize = read_int (allparams, "maxsize")
       
   369     val maxvars = read_int (allparams, "maxvars")
       
   370     val maxtime = read_int (allparams, "maxtime")
       
   371     (* string *)
       
   372     val satsolver = read_string (allparams, "satsolver")
       
   373     val no_assms = read_bool (allparams, "no_assms")
       
   374     val expect = the_default "" (AList.lookup (op =) allparams "expect")
       
   375     (* all remaining parameters of the form "string=int" are collected in *)
       
   376     (* 'sizes'                                                            *)
       
   377     (* TODO: it is currently not possible to specify a size for a type    *)
       
   378     (*       whose name is one of the other parameters (e.g. 'maxvars')   *)
       
   379     (* (string * int) list *)
       
   380     val sizes = map_filter
       
   381       (fn (name, value) => Option.map (pair name) (Int.fromString value))
       
   382       (filter (fn (name, _) => name<>"minsize" andalso name<>"maxsize"
       
   383         andalso name<>"maxvars" andalso name<>"maxtime"
       
   384         andalso name<>"satsolver" andalso name<>"no_assms") allparams)
       
   385   in
       
   386     {sizes=sizes, minsize=minsize, maxsize=maxsize, maxvars=maxvars,
       
   387       maxtime=maxtime, satsolver=satsolver, no_assms=no_assms, expect=expect}
       
   388   end;
       
   389 
       
   390 
       
   391 (* ------------------------------------------------------------------------- *)
       
   392 (* TRANSLATION HOL -> PROPOSITIONAL LOGIC, BOOLEAN ASSIGNMENT -> MODEL       *)
       
   393 (* ------------------------------------------------------------------------- *)
       
   394 
       
   395 val typ_of_dtyp = ATP_Util.typ_of_dtyp
       
   396 
       
   397 (* ------------------------------------------------------------------------- *)
       
   398 (* close_form: universal closure over schematic variables in 't'             *)
       
   399 (* ------------------------------------------------------------------------- *)
       
   400 
       
   401 (* Term.term -> Term.term *)
       
   402 
       
   403 fun close_form t =
       
   404   let
       
   405     val vars = sort_wrt (fst o fst) (Term.add_vars t [])
       
   406   in
       
   407     fold (fn ((x, i), T) => fn t' =>
       
   408       Logic.all_const T $ Abs (x, T, abstract_over (Var ((x, i), T), t'))) vars t
       
   409   end;
       
   410 
       
   411 val monomorphic_term = ATP_Util.monomorphic_term
       
   412 val specialize_type = ATP_Util.specialize_type
       
   413 
       
   414 (* ------------------------------------------------------------------------- *)
       
   415 (* is_const_of_class: returns 'true' iff 'Const (s, T)' is a constant that   *)
       
   416 (*                    denotes membership to an axiomatic type class          *)
       
   417 (* ------------------------------------------------------------------------- *)
       
   418 
       
   419 fun is_const_of_class thy (s, _) =
       
   420   let
       
   421     val class_const_names = map Logic.const_of_class (Sign.all_classes thy)
       
   422   in
       
   423     (* I'm not quite sure if checking the name 's' is sufficient, *)
       
   424     (* or if we should also check the type 'T'.                   *)
       
   425     member (op =) class_const_names s
       
   426   end;
       
   427 
       
   428 (* ------------------------------------------------------------------------- *)
       
   429 (* is_IDT_constructor: returns 'true' iff 'Const (s, T)' is the constructor  *)
       
   430 (*                     of an inductive datatype in 'thy'                     *)
       
   431 (* ------------------------------------------------------------------------- *)
       
   432 
       
   433 fun is_IDT_constructor thy (s, T) =
       
   434   (case body_type T of
       
   435     Type (s', _) =>
       
   436       (case Datatype.get_constrs thy s' of
       
   437         SOME constrs =>
       
   438           List.exists (fn (cname, cty) =>
       
   439             cname = s andalso Sign.typ_instance thy (T, cty)) constrs
       
   440       | NONE => false)
       
   441   | _  => false);
       
   442 
       
   443 (* ------------------------------------------------------------------------- *)
       
   444 (* is_IDT_recursor: returns 'true' iff 'Const (s, T)' is the recursion       *)
       
   445 (*                  operator of an inductive datatype in 'thy'               *)
       
   446 (* ------------------------------------------------------------------------- *)
       
   447 
       
   448 fun is_IDT_recursor thy (s, _) =
       
   449   let
       
   450     val rec_names = Symtab.fold (append o #rec_names o snd)
       
   451       (Datatype.get_all thy) []
       
   452   in
       
   453     (* I'm not quite sure if checking the name 's' is sufficient, *)
       
   454     (* or if we should also check the type 'T'.                   *)
       
   455     member (op =) rec_names s
       
   456   end;
       
   457 
       
   458 (* ------------------------------------------------------------------------- *)
       
   459 (* norm_rhs: maps  f ?t1 ... ?tn == rhs  to  %t1...tn. rhs                   *)
       
   460 (* ------------------------------------------------------------------------- *)
       
   461 
       
   462 fun norm_rhs eqn =
       
   463   let
       
   464     fun lambda (v as Var ((x, _), T)) t = Abs (x, T, abstract_over (v, t))
       
   465       | lambda v t = raise TERM ("lambda", [v, t])
       
   466     val (lhs, rhs) = Logic.dest_equals eqn
       
   467     val (_, args) = Term.strip_comb lhs
       
   468   in
       
   469     fold lambda (rev args) rhs
       
   470   end
       
   471 
       
   472 (* ------------------------------------------------------------------------- *)
       
   473 (* get_def: looks up the definition of a constant                            *)
       
   474 (* ------------------------------------------------------------------------- *)
       
   475 
       
   476 fun get_def thy (s, T) =
       
   477   let
       
   478     (* (string * Term.term) list -> (string * Term.term) option *)
       
   479     fun get_def_ax [] = NONE
       
   480       | get_def_ax ((axname, ax) :: axioms) =
       
   481           (let
       
   482             val (lhs, _) = Logic.dest_equals ax  (* equations only *)
       
   483             val c        = Term.head_of lhs
       
   484             val (s', T') = Term.dest_Const c
       
   485           in
       
   486             if s=s' then
       
   487               let
       
   488                 val typeSubs = Sign.typ_match thy (T', T) Vartab.empty
       
   489                 val ax'      = monomorphic_term typeSubs ax
       
   490                 val rhs      = norm_rhs ax'
       
   491               in
       
   492                 SOME (axname, rhs)
       
   493               end
       
   494             else
       
   495               get_def_ax axioms
       
   496           end handle ERROR _         => get_def_ax axioms
       
   497                    | TERM _          => get_def_ax axioms
       
   498                    | Type.TYPE_MATCH => get_def_ax axioms)
       
   499   in
       
   500     get_def_ax (Theory.all_axioms_of thy)
       
   501   end;
       
   502 
       
   503 (* ------------------------------------------------------------------------- *)
       
   504 (* get_typedef: looks up the definition of a type, as created by "typedef"   *)
       
   505 (* ------------------------------------------------------------------------- *)
       
   506 
       
   507 fun get_typedef thy T =
       
   508   let
       
   509     (* (string * Term.term) list -> (string * Term.term) option *)
       
   510     fun get_typedef_ax [] = NONE
       
   511       | get_typedef_ax ((axname, ax) :: axioms) =
       
   512           (let
       
   513             (* Term.term -> Term.typ option *)
       
   514             fun type_of_type_definition (Const (s', T')) =
       
   515                   if s'= @{const_name type_definition} then
       
   516                     SOME T'
       
   517                   else
       
   518                     NONE
       
   519               | type_of_type_definition (Free _) = NONE
       
   520               | type_of_type_definition (Var _) = NONE
       
   521               | type_of_type_definition (Bound _) = NONE
       
   522               | type_of_type_definition (Abs (_, _, body)) =
       
   523                   type_of_type_definition body
       
   524               | type_of_type_definition (t1 $ t2) =
       
   525                   (case type_of_type_definition t1 of
       
   526                     SOME x => SOME x
       
   527                   | NONE => type_of_type_definition t2)
       
   528           in
       
   529             case type_of_type_definition ax of
       
   530               SOME T' =>
       
   531                 let
       
   532                   val T'' = domain_type (domain_type T')
       
   533                   val typeSubs = Sign.typ_match thy (T'', T) Vartab.empty
       
   534                 in
       
   535                   SOME (axname, monomorphic_term typeSubs ax)
       
   536                 end
       
   537             | NONE => get_typedef_ax axioms
       
   538           end handle ERROR _         => get_typedef_ax axioms
       
   539                    | TERM _          => get_typedef_ax axioms
       
   540                    | Type.TYPE_MATCH => get_typedef_ax axioms)
       
   541   in
       
   542     get_typedef_ax (Theory.all_axioms_of thy)
       
   543   end;
       
   544 
       
   545 (* ------------------------------------------------------------------------- *)
       
   546 (* get_classdef: looks up the defining axiom for an axiomatic type class, as *)
       
   547 (*               created by the "axclass" command                            *)
       
   548 (* ------------------------------------------------------------------------- *)
       
   549 
       
   550 fun get_classdef thy class =
       
   551   let
       
   552     val axname = class ^ "_class_def"
       
   553   in
       
   554     Option.map (pair axname)
       
   555       (AList.lookup (op =) (Theory.all_axioms_of thy) axname)
       
   556   end;
       
   557 
       
   558 (* ------------------------------------------------------------------------- *)
       
   559 (* unfold_defs: unfolds all defined constants in a term 't', beta-eta        *)
       
   560 (*              normalizes the result term; certain constants are not        *)
       
   561 (*              unfolded (cf. 'collect_axioms' and the various interpreters  *)
       
   562 (*              below): if the interpretation respects a definition anyway,  *)
       
   563 (*              that definition does not need to be unfolded                 *)
       
   564 (* ------------------------------------------------------------------------- *)
       
   565 
       
   566 (* Note: we could intertwine unfolding of constants and beta-(eta-)       *)
       
   567 (*       normalization; this would save some unfolding for terms where    *)
       
   568 (*       constants are eliminated by beta-reduction (e.g. 'K c1 c2').  On *)
       
   569 (*       the other hand, this would cause additional work for terms where *)
       
   570 (*       constants are duplicated by beta-reduction (e.g. 'S c1 c2 c3').  *)
       
   571 
       
   572 fun unfold_defs thy t =
       
   573   let
       
   574     (* Term.term -> Term.term *)
       
   575     fun unfold_loop t =
       
   576       case t of
       
   577       (* Pure *)
       
   578         Const (@{const_name all}, _) => t
       
   579       | Const (@{const_name "=="}, _) => t
       
   580       | Const (@{const_name "==>"}, _) => t
       
   581       | Const (@{const_name TYPE}, _) => t  (* axiomatic type classes *)
       
   582       (* HOL *)
       
   583       | Const (@{const_name Trueprop}, _) => t
       
   584       | Const (@{const_name Not}, _) => t
       
   585       | (* redundant, since 'True' is also an IDT constructor *)
       
   586         Const (@{const_name True}, _) => t
       
   587       | (* redundant, since 'False' is also an IDT constructor *)
       
   588         Const (@{const_name False}, _) => t
       
   589       | Const (@{const_name undefined}, _) => t
       
   590       | Const (@{const_name The}, _) => t
       
   591       | Const (@{const_name Hilbert_Choice.Eps}, _) => t
       
   592       | Const (@{const_name All}, _) => t
       
   593       | Const (@{const_name Ex}, _) => t
       
   594       | Const (@{const_name HOL.eq}, _) => t
       
   595       | Const (@{const_name HOL.conj}, _) => t
       
   596       | Const (@{const_name HOL.disj}, _) => t
       
   597       | Const (@{const_name HOL.implies}, _) => t
       
   598       (* sets *)
       
   599       | Const (@{const_name Collect}, _) => t
       
   600       | Const (@{const_name Set.member}, _) => t
       
   601       (* other optimizations *)
       
   602       | Const (@{const_name Finite_Set.card}, _) => t
       
   603       | Const (@{const_name Finite_Set.finite}, _) => t
       
   604       | Const (@{const_name Orderings.less}, Type ("fun", [@{typ nat},
       
   605           Type ("fun", [@{typ nat}, @{typ bool}])])) => t
       
   606       | Const (@{const_name Groups.plus}, Type ("fun", [@{typ nat},
       
   607           Type ("fun", [@{typ nat}, @{typ nat}])])) => t
       
   608       | Const (@{const_name Groups.minus}, Type ("fun", [@{typ nat},
       
   609           Type ("fun", [@{typ nat}, @{typ nat}])])) => t
       
   610       | Const (@{const_name Groups.times}, Type ("fun", [@{typ nat},
       
   611           Type ("fun", [@{typ nat}, @{typ nat}])])) => t
       
   612       | Const (@{const_name List.append}, _) => t
       
   613 (* UNSOUND
       
   614       | Const (@{const_name lfp}, _) => t
       
   615       | Const (@{const_name gfp}, _) => t
       
   616 *)
       
   617       | Const (@{const_name fst}, _) => t
       
   618       | Const (@{const_name snd}, _) => t
       
   619       (* simply-typed lambda calculus *)
       
   620       | Const (s, T) =>
       
   621           (if is_IDT_constructor thy (s, T)
       
   622             orelse is_IDT_recursor thy (s, T) then
       
   623             t  (* do not unfold IDT constructors/recursors *)
       
   624           (* unfold the constant if there is a defining equation *)
       
   625           else
       
   626             case get_def thy (s, T) of
       
   627               SOME ((*axname*) _, rhs) =>
       
   628               (* Note: if the term to be unfolded (i.e. 'Const (s, T)')  *)
       
   629               (* occurs on the right-hand side of the equation, i.e. in  *)
       
   630               (* 'rhs', we must not use this equation to unfold, because *)
       
   631               (* that would loop.  Here would be the right place to      *)
       
   632               (* check this.  However, getting this really right seems   *)
       
   633               (* difficult because the user may state arbitrary axioms,  *)
       
   634               (* which could interact with overloading to create loops.  *)
       
   635               ((*tracing (" unfolding: " ^ axname);*)
       
   636                unfold_loop rhs)
       
   637             | NONE => t)
       
   638       | Free _ => t
       
   639       | Var _ => t
       
   640       | Bound _ => t
       
   641       | Abs (s, T, body) => Abs (s, T, unfold_loop body)
       
   642       | t1 $ t2 => (unfold_loop t1) $ (unfold_loop t2)
       
   643     val result = Envir.beta_eta_contract (unfold_loop t)
       
   644   in
       
   645     result
       
   646   end;
       
   647 
       
   648 (* ------------------------------------------------------------------------- *)
       
   649 (* collect_axioms: collects (monomorphic, universally quantified, unfolded   *)
       
   650 (*                 versions of) all HOL axioms that are relevant w.r.t 't'   *)
       
   651 (* ------------------------------------------------------------------------- *)
       
   652 
       
   653 (* Note: to make the collection of axioms more easily extensible, this    *)
       
   654 (*       function could be based on user-supplied "axiom collectors",     *)
       
   655 (*       similar to 'interpret'/interpreters or 'print'/printers          *)
       
   656 
       
   657 (* Note: currently we use "inverse" functions to the definitional         *)
       
   658 (*       mechanisms provided by Isabelle/HOL, e.g. for "axclass",         *)
       
   659 (*       "typedef", "definition".  A more general approach could consider *)
       
   660 (*       *every* axiom of the theory and collect it if it has a constant/ *)
       
   661 (*       type/typeclass in common with the term 't'.                      *)
       
   662 
       
   663 (* Which axioms are "relevant" for a particular term/type goes hand in    *)
       
   664 (* hand with the interpretation of that term/type by its interpreter (see *)
       
   665 (* way below): if the interpretation respects an axiom anyway, the axiom  *)
       
   666 (* does not need to be added as a constraint here.                        *)
       
   667 
       
   668 (* To avoid collecting the same axiom multiple times, we use an           *)
       
   669 (* accumulator 'axs' which contains all axioms collected so far.          *)
       
   670 
       
   671 fun collect_axioms ctxt t =
       
   672   let
       
   673     val thy = Proof_Context.theory_of ctxt
       
   674     val _ = tracing "Adding axioms..."
       
   675     val axioms = Theory.all_axioms_of thy
       
   676     fun collect_this_axiom (axname, ax) axs =
       
   677       let
       
   678         val ax' = unfold_defs thy ax
       
   679       in
       
   680         if member (op aconv) axs ax' then axs
       
   681         else (tracing axname; collect_term_axioms ax' (ax' :: axs))
       
   682       end
       
   683     and collect_sort_axioms T axs =
       
   684       let
       
   685         val sort =
       
   686           (case T of
       
   687             TFree (_, sort) => sort
       
   688           | TVar (_, sort)  => sort
       
   689           | _ => raise REFUTE ("collect_axioms",
       
   690               "type " ^ Syntax.string_of_typ ctxt T ^ " is not a variable"))
       
   691         (* obtain axioms for all superclasses *)
       
   692         val superclasses = sort @ maps (Sign.super_classes thy) sort
       
   693         (* merely an optimization, because 'collect_this_axiom' disallows *)
       
   694         (* duplicate axioms anyway:                                       *)
       
   695         val superclasses = distinct (op =) superclasses
       
   696         val class_axioms = maps (fn class => map (fn ax =>
       
   697           ("<" ^ class ^ ">", Thm.prop_of ax))
       
   698           (#axioms (AxClass.get_info thy class) handle ERROR _ => []))
       
   699           superclasses
       
   700         (* replace the (at most one) schematic type variable in each axiom *)
       
   701         (* by the actual type 'T'                                          *)
       
   702         val monomorphic_class_axioms = map (fn (axname, ax) =>
       
   703           (case Term.add_tvars ax [] of
       
   704             [] => (axname, ax)
       
   705           | [(idx, S)] => (axname, monomorphic_term (Vartab.make [(idx, (S, T))]) ax)
       
   706           | _ =>
       
   707             raise REFUTE ("collect_axioms", "class axiom " ^ axname ^ " (" ^
       
   708               Syntax.string_of_term ctxt ax ^
       
   709               ") contains more than one type variable")))
       
   710           class_axioms
       
   711       in
       
   712         fold collect_this_axiom monomorphic_class_axioms axs
       
   713       end
       
   714     and collect_type_axioms T axs =
       
   715       case T of
       
   716       (* simple types *)
       
   717         Type ("prop", []) => axs
       
   718       | Type ("fun", [T1, T2]) => collect_type_axioms T2 (collect_type_axioms T1 axs)
       
   719       | Type (@{type_name set}, [T1]) => collect_type_axioms T1 axs
       
   720       (* axiomatic type classes *)
       
   721       | Type ("itself", [T1]) => collect_type_axioms T1 axs
       
   722       | Type (s, Ts) =>
       
   723         (case Datatype.get_info thy s of
       
   724           SOME _ =>  (* inductive datatype *)
       
   725             (* only collect relevant type axioms for the argument types *)
       
   726             fold collect_type_axioms Ts axs
       
   727         | NONE =>
       
   728           (case get_typedef thy T of
       
   729             SOME (axname, ax) =>
       
   730               collect_this_axiom (axname, ax) axs
       
   731           | NONE =>
       
   732             (* unspecified type, perhaps introduced with "typedecl" *)
       
   733             (* at least collect relevant type axioms for the argument types *)
       
   734             fold collect_type_axioms Ts axs))
       
   735       (* axiomatic type classes *)
       
   736       | TFree _ => collect_sort_axioms T axs
       
   737       (* axiomatic type classes *)
       
   738       | TVar _ => collect_sort_axioms T axs
       
   739     and collect_term_axioms t axs =
       
   740       case t of
       
   741       (* Pure *)
       
   742         Const (@{const_name all}, _) => axs
       
   743       | Const (@{const_name "=="}, _) => axs
       
   744       | Const (@{const_name "==>"}, _) => axs
       
   745       (* axiomatic type classes *)
       
   746       | Const (@{const_name TYPE}, T) => collect_type_axioms T axs
       
   747       (* HOL *)
       
   748       | Const (@{const_name Trueprop}, _) => axs
       
   749       | Const (@{const_name Not}, _) => axs
       
   750       (* redundant, since 'True' is also an IDT constructor *)
       
   751       | Const (@{const_name True}, _) => axs
       
   752       (* redundant, since 'False' is also an IDT constructor *)
       
   753       | Const (@{const_name False}, _) => axs
       
   754       | Const (@{const_name undefined}, T) => collect_type_axioms T axs
       
   755       | Const (@{const_name The}, T) =>
       
   756           let
       
   757             val ax = specialize_type thy (@{const_name The}, T)
       
   758               (the (AList.lookup (op =) axioms "HOL.the_eq_trivial"))
       
   759           in
       
   760             collect_this_axiom ("HOL.the_eq_trivial", ax) axs
       
   761           end
       
   762       | Const (@{const_name Hilbert_Choice.Eps}, T) =>
       
   763           let
       
   764             val ax = specialize_type thy (@{const_name Hilbert_Choice.Eps}, T)
       
   765               (the (AList.lookup (op =) axioms "Hilbert_Choice.someI"))
       
   766           in
       
   767             collect_this_axiom ("Hilbert_Choice.someI", ax) axs
       
   768           end
       
   769       | Const (@{const_name All}, T) => collect_type_axioms T axs
       
   770       | Const (@{const_name Ex}, T) => collect_type_axioms T axs
       
   771       | Const (@{const_name HOL.eq}, T) => collect_type_axioms T axs
       
   772       | Const (@{const_name HOL.conj}, _) => axs
       
   773       | Const (@{const_name HOL.disj}, _) => axs
       
   774       | Const (@{const_name HOL.implies}, _) => axs
       
   775       (* sets *)
       
   776       | Const (@{const_name Collect}, T) => collect_type_axioms T axs
       
   777       | Const (@{const_name Set.member}, T) => collect_type_axioms T axs
       
   778       (* other optimizations *)
       
   779       | Const (@{const_name Finite_Set.card}, T) => collect_type_axioms T axs
       
   780       | Const (@{const_name Finite_Set.finite}, T) =>
       
   781         collect_type_axioms T axs
       
   782       | Const (@{const_name Orderings.less}, T as Type ("fun", [@{typ nat},
       
   783         Type ("fun", [@{typ nat}, @{typ bool}])])) =>
       
   784           collect_type_axioms T axs
       
   785       | Const (@{const_name Groups.plus}, T as Type ("fun", [@{typ nat},
       
   786         Type ("fun", [@{typ nat}, @{typ nat}])])) =>
       
   787           collect_type_axioms T axs
       
   788       | Const (@{const_name Groups.minus}, T as Type ("fun", [@{typ nat},
       
   789         Type ("fun", [@{typ nat}, @{typ nat}])])) =>
       
   790           collect_type_axioms T axs
       
   791       | Const (@{const_name Groups.times}, T as Type ("fun", [@{typ nat},
       
   792         Type ("fun", [@{typ nat}, @{typ nat}])])) =>
       
   793           collect_type_axioms T axs
       
   794       | Const (@{const_name List.append}, T) => collect_type_axioms T axs
       
   795 (* UNSOUND
       
   796       | Const (@{const_name lfp}, T) => collect_type_axioms T axs
       
   797       | Const (@{const_name gfp}, T) => collect_type_axioms T axs
       
   798 *)
       
   799       | Const (@{const_name fst}, T) => collect_type_axioms T axs
       
   800       | Const (@{const_name snd}, T) => collect_type_axioms T axs
       
   801       (* simply-typed lambda calculus *)
       
   802       | Const (s, T) =>
       
   803           if is_const_of_class thy (s, T) then
       
   804             (* axiomatic type classes: add "OFCLASS(?'a::c, c_class)" *)
       
   805             (* and the class definition                               *)
       
   806             let
       
   807               val class = Logic.class_of_const s
       
   808               val of_class = Logic.mk_of_class (TVar (("'a", 0), [class]), class)
       
   809               val ax_in = SOME (specialize_type thy (s, T) of_class)
       
   810                 (* type match may fail due to sort constraints *)
       
   811                 handle Type.TYPE_MATCH => NONE
       
   812               val ax_1 = Option.map (fn ax => (Syntax.string_of_term ctxt ax, ax)) ax_in
       
   813               val ax_2 = Option.map (apsnd (specialize_type thy (s, T))) (get_classdef thy class)
       
   814             in
       
   815               collect_type_axioms T (fold collect_this_axiom (map_filter I [ax_1, ax_2]) axs)
       
   816             end
       
   817           else if is_IDT_constructor thy (s, T)
       
   818             orelse is_IDT_recursor thy (s, T)
       
   819           then
       
   820             (* only collect relevant type axioms *)
       
   821             collect_type_axioms T axs
       
   822           else
       
   823             (* other constants should have been unfolded, with some *)
       
   824             (* exceptions: e.g. Abs_xxx/Rep_xxx functions for       *)
       
   825             (* typedefs, or type-class related constants            *)
       
   826             (* only collect relevant type axioms *)
       
   827             collect_type_axioms T axs
       
   828       | Free (_, T) => collect_type_axioms T axs
       
   829       | Var (_, T) => collect_type_axioms T axs
       
   830       | Bound _ => axs
       
   831       | Abs (_, T, body) => collect_term_axioms body (collect_type_axioms T axs)
       
   832       | t1 $ t2 => collect_term_axioms t2 (collect_term_axioms t1 axs)
       
   833     val result = map close_form (collect_term_axioms t [])
       
   834     val _ = tracing " ...done."
       
   835   in
       
   836     result
       
   837   end;
       
   838 
       
   839 (* ------------------------------------------------------------------------- *)
       
   840 (* ground_types: collects all ground types in a term (including argument     *)
       
   841 (*               types of other types), suppressing duplicates.  Does not    *)
       
   842 (*               return function types, set types, non-recursive IDTs, or    *)
       
   843 (*               'propT'.  For IDTs, also the argument types of constructors *)
       
   844 (*               and all mutually recursive IDTs are considered.             *)
       
   845 (* ------------------------------------------------------------------------- *)
       
   846 
       
   847 fun ground_types ctxt t =
       
   848   let
       
   849     val thy = Proof_Context.theory_of ctxt
       
   850     fun collect_types T acc =
       
   851       (case T of
       
   852         Type ("fun", [T1, T2]) => collect_types T1 (collect_types T2 acc)
       
   853       | Type ("prop", []) => acc
       
   854       | Type (@{type_name set}, [T1]) => collect_types T1 acc
       
   855       | Type (s, Ts) =>
       
   856           (case Datatype.get_info thy s of
       
   857             SOME info =>  (* inductive datatype *)
       
   858               let
       
   859                 val index = #index info
       
   860                 val descr = #descr info
       
   861                 val (_, typs, _) = the (AList.lookup (op =) descr index)
       
   862                 val typ_assoc = typs ~~ Ts
       
   863                 (* sanity check: every element in 'dtyps' must be a *)
       
   864                 (* 'DtTFree'                                        *)
       
   865                 val _ = if Library.exists (fn d =>
       
   866                   case d of Datatype.DtTFree _ => false | _ => true) typs then
       
   867                   raise REFUTE ("ground_types", "datatype argument (for type "
       
   868                     ^ Syntax.string_of_typ ctxt T ^ ") is not a variable")
       
   869                 else ()
       
   870                 (* required for mutually recursive datatypes; those need to   *)
       
   871                 (* be added even if they are an instance of an otherwise non- *)
       
   872                 (* recursive datatype                                         *)
       
   873                 fun collect_dtyp d acc =
       
   874                   let
       
   875                     val dT = typ_of_dtyp descr typ_assoc d
       
   876                   in
       
   877                     case d of
       
   878                       Datatype.DtTFree _ =>
       
   879                       collect_types dT acc
       
   880                     | Datatype.DtType (_, ds) =>
       
   881                       collect_types dT (fold_rev collect_dtyp ds acc)
       
   882                     | Datatype.DtRec i =>
       
   883                       if member (op =) acc dT then
       
   884                         acc  (* prevent infinite recursion *)
       
   885                       else
       
   886                         let
       
   887                           val (_, dtyps, dconstrs) = the (AList.lookup (op =) descr i)
       
   888                           (* if the current type is a recursive IDT (i.e. a depth *)
       
   889                           (* is required), add it to 'acc'                        *)
       
   890                           val acc_dT = if Library.exists (fn (_, ds) =>
       
   891                             Library.exists Datatype_Aux.is_rec_type ds) dconstrs then
       
   892                               insert (op =) dT acc
       
   893                             else acc
       
   894                           (* collect argument types *)
       
   895                           val acc_dtyps = fold_rev collect_dtyp dtyps acc_dT
       
   896                           (* collect constructor types *)
       
   897                           val acc_dconstrs = fold_rev collect_dtyp (maps snd dconstrs) acc_dtyps
       
   898                         in
       
   899                           acc_dconstrs
       
   900                         end
       
   901                   end
       
   902               in
       
   903                 (* argument types 'Ts' could be added here, but they are also *)
       
   904                 (* added by 'collect_dtyp' automatically                      *)
       
   905                 collect_dtyp (Datatype.DtRec index) acc
       
   906               end
       
   907           | NONE =>
       
   908             (* not an inductive datatype, e.g. defined via "typedef" or *)
       
   909             (* "typedecl"                                               *)
       
   910             insert (op =) T (fold collect_types Ts acc))
       
   911       | TFree _ => insert (op =) T acc
       
   912       | TVar _ => insert (op =) T acc)
       
   913   in
       
   914     fold_types collect_types t []
       
   915   end;
       
   916 
       
   917 (* ------------------------------------------------------------------------- *)
       
   918 (* string_of_typ: (rather naive) conversion from types to strings, used to   *)
       
   919 (*                look up the size of a type in 'sizes'.  Parameterized      *)
       
   920 (*                types with different parameters (e.g. "'a list" vs. "bool  *)
       
   921 (*                list") are identified.                                     *)
       
   922 (* ------------------------------------------------------------------------- *)
       
   923 
       
   924 (* Term.typ -> string *)
       
   925 
       
   926 fun string_of_typ (Type (s, _))     = s
       
   927   | string_of_typ (TFree (s, _))    = s
       
   928   | string_of_typ (TVar ((s,_), _)) = s;
       
   929 
       
   930 (* ------------------------------------------------------------------------- *)
       
   931 (* first_universe: returns the "first" (i.e. smallest) universe by assigning *)
       
   932 (*                 'minsize' to every type for which no size is specified in *)
       
   933 (*                 'sizes'                                                   *)
       
   934 (* ------------------------------------------------------------------------- *)
       
   935 
       
   936 (* Term.typ list -> (string * int) list -> int -> (Term.typ * int) list *)
       
   937 
       
   938 fun first_universe xs sizes minsize =
       
   939   let
       
   940     fun size_of_typ T =
       
   941       case AList.lookup (op =) sizes (string_of_typ T) of
       
   942         SOME n => n
       
   943       | NONE => minsize
       
   944   in
       
   945     map (fn T => (T, size_of_typ T)) xs
       
   946   end;
       
   947 
       
   948 (* ------------------------------------------------------------------------- *)
       
   949 (* next_universe: enumerates all universes (i.e. assignments of sizes to     *)
       
   950 (*                types), where the minimal size of a type is given by       *)
       
   951 (*                'minsize', the maximal size is given by 'maxsize', and a   *)
       
   952 (*                type may have a fixed size given in 'sizes'                *)
       
   953 (* ------------------------------------------------------------------------- *)
       
   954 
       
   955 (* (Term.typ * int) list -> (string * int) list -> int -> int ->
       
   956   (Term.typ * int) list option *)
       
   957 
       
   958 fun next_universe xs sizes minsize maxsize =
       
   959   let
       
   960     (* creates the "first" list of length 'len', where the sum of all list *)
       
   961     (* elements is 'sum', and the length of the list is 'len'              *)
       
   962     (* int -> int -> int -> int list option *)
       
   963     fun make_first _ 0 sum =
       
   964           if sum = 0 then
       
   965             SOME []
       
   966           else
       
   967             NONE
       
   968       | make_first max len sum =
       
   969           if sum <= max orelse max < 0 then
       
   970             Option.map (fn xs' => sum :: xs') (make_first max (len-1) 0)
       
   971           else
       
   972             Option.map (fn xs' => max :: xs') (make_first max (len-1) (sum-max))
       
   973     (* enumerates all int lists with a fixed length, where 0<=x<='max' for *)
       
   974     (* all list elements x (unless 'max'<0)                                *)
       
   975     (* int -> int -> int -> int list -> int list option *)
       
   976     fun next _ _ _ [] =
       
   977           NONE
       
   978       | next max len sum [x] =
       
   979           (* we've reached the last list element, so there's no shift possible *)
       
   980           make_first max (len+1) (sum+x+1)  (* increment 'sum' by 1 *)
       
   981       | next max len sum (x1::x2::xs) =
       
   982           if x1>0 andalso (x2<max orelse max<0) then
       
   983             (* we can shift *)
       
   984             SOME (the (make_first max (len+1) (sum+x1-1)) @ (x2+1) :: xs)
       
   985           else
       
   986             (* continue search *)
       
   987             next max (len+1) (sum+x1) (x2::xs)
       
   988     (* only consider those types for which the size is not fixed *)
       
   989     val mutables = filter_out (AList.defined (op =) sizes o string_of_typ o fst) xs
       
   990     (* subtract 'minsize' from every size (will be added again at the end) *)
       
   991     val diffs = map (fn (_, n) => n-minsize) mutables
       
   992   in
       
   993     case next (maxsize-minsize) 0 0 diffs of
       
   994       SOME diffs' =>
       
   995         (* merge with those types for which the size is fixed *)
       
   996         SOME (fst (fold_map (fn (T, _) => fn ds =>
       
   997           case AList.lookup (op =) sizes (string_of_typ T) of
       
   998           (* return the fixed size *)
       
   999             SOME n => ((T, n), ds)
       
  1000           (* consume the head of 'ds', add 'minsize' *)
       
  1001           | NONE   => ((T, minsize + hd ds), tl ds))
       
  1002           xs diffs'))
       
  1003     | NONE => NONE
       
  1004   end;
       
  1005 
       
  1006 (* ------------------------------------------------------------------------- *)
       
  1007 (* toTrue: converts the interpretation of a Boolean value to a propositional *)
       
  1008 (*         formula that is true iff the interpretation denotes "true"        *)
       
  1009 (* ------------------------------------------------------------------------- *)
       
  1010 
       
  1011 (* interpretation -> prop_formula *)
       
  1012 
       
  1013 fun toTrue (Leaf [fm, _]) = fm
       
  1014   | toTrue _ = raise REFUTE ("toTrue", "interpretation does not denote a Boolean value");
       
  1015 
       
  1016 (* ------------------------------------------------------------------------- *)
       
  1017 (* toFalse: converts the interpretation of a Boolean value to a              *)
       
  1018 (*          propositional formula that is true iff the interpretation        *)
       
  1019 (*          denotes "false"                                                  *)
       
  1020 (* ------------------------------------------------------------------------- *)
       
  1021 
       
  1022 (* interpretation -> prop_formula *)
       
  1023 
       
  1024 fun toFalse (Leaf [_, fm]) = fm
       
  1025   | toFalse _ = raise REFUTE ("toFalse", "interpretation does not denote a Boolean value");
       
  1026 
       
  1027 (* ------------------------------------------------------------------------- *)
       
  1028 (* find_model: repeatedly calls 'interpret' with appropriate parameters,     *)
       
  1029 (*             applies a SAT solver, and (in case a model is found) displays *)
       
  1030 (*             the model to the user by calling 'print_model'                *)
       
  1031 (* {...}     : parameters that control the translation/model generation      *)
       
  1032 (* assm_ts   : assumptions to be considered unless "no_assms" is specified   *)
       
  1033 (* t         : term to be translated into a propositional formula            *)
       
  1034 (* negate    : if true, find a model that makes 't' false (rather than true) *)
       
  1035 (* ------------------------------------------------------------------------- *)
       
  1036 
       
  1037 fun find_model ctxt
       
  1038     {sizes, minsize, maxsize, maxvars, maxtime, satsolver, no_assms, expect}
       
  1039     assm_ts t negate =
       
  1040   let
       
  1041     val thy = Proof_Context.theory_of ctxt
       
  1042     (* string -> string *)
       
  1043     fun check_expect outcome_code =
       
  1044       if expect = "" orelse outcome_code = expect then outcome_code
       
  1045       else error ("Unexpected outcome: " ^ quote outcome_code ^ ".")
       
  1046     (* unit -> string *)
       
  1047     fun wrapper () =
       
  1048       let
       
  1049         val timer = Timer.startRealTimer ()
       
  1050         val t =
       
  1051           if no_assms then t
       
  1052           else if negate then Logic.list_implies (assm_ts, t)
       
  1053           else Logic.mk_conjunction_list (t :: assm_ts)
       
  1054         val u = unfold_defs thy t
       
  1055         val _ = tracing ("Unfolded term: " ^ Syntax.string_of_term ctxt u)
       
  1056         val axioms = collect_axioms ctxt u
       
  1057         (* Term.typ list *)
       
  1058         val types = fold (union (op =) o ground_types ctxt) (u :: axioms) []
       
  1059         val _ = tracing ("Ground types: "
       
  1060           ^ (if null types then "none."
       
  1061              else commas (map (Syntax.string_of_typ ctxt) types)))
       
  1062         (* we can only consider fragments of recursive IDTs, so we issue a  *)
       
  1063         (* warning if the formula contains a recursive IDT                  *)
       
  1064         (* TODO: no warning needed for /positive/ occurrences of IDTs       *)
       
  1065         val maybe_spurious = Library.exists (fn
       
  1066             Type (s, _) =>
       
  1067               (case Datatype.get_info thy s of
       
  1068                 SOME info =>  (* inductive datatype *)
       
  1069                   let
       
  1070                     val index           = #index info
       
  1071                     val descr           = #descr info
       
  1072                     val (_, _, constrs) = the (AList.lookup (op =) descr index)
       
  1073                   in
       
  1074                     (* recursive datatype? *)
       
  1075                     Library.exists (fn (_, ds) =>
       
  1076                       Library.exists Datatype_Aux.is_rec_type ds) constrs
       
  1077                   end
       
  1078               | NONE => false)
       
  1079           | _ => false) types
       
  1080         val _ =
       
  1081           if maybe_spurious then
       
  1082             warning ("Term contains a recursive datatype; "
       
  1083               ^ "countermodel(s) may be spurious!")
       
  1084           else
       
  1085             ()
       
  1086         (* (Term.typ * int) list -> string *)
       
  1087         fun find_model_loop universe =
       
  1088           let
       
  1089             val msecs_spent = Time.toMilliseconds (Timer.checkRealTimer timer)
       
  1090             val _ = maxtime = 0 orelse msecs_spent < 1000 * maxtime
       
  1091                     orelse raise TimeLimit.TimeOut
       
  1092             val init_model = (universe, [])
       
  1093             val init_args  = {maxvars = maxvars, def_eq = false, next_idx = 1,
       
  1094               bounds = [], wellformed = True}
       
  1095             val _ = tracing ("Translating term (sizes: "
       
  1096               ^ commas (map (fn (_, n) => string_of_int n) universe) ^ ") ...")
       
  1097             (* translate 'u' and all axioms *)
       
  1098             val (intrs, (model, args)) = fold_map (fn t' => fn (m, a) =>
       
  1099               let
       
  1100                 val (i, m', a') = interpret ctxt m a t'
       
  1101               in
       
  1102                 (* set 'def_eq' to 'true' *)
       
  1103                 (i, (m', {maxvars = #maxvars a', def_eq = true,
       
  1104                   next_idx = #next_idx a', bounds = #bounds a',
       
  1105                   wellformed = #wellformed a'}))
       
  1106               end) (u :: axioms) (init_model, init_args)
       
  1107             (* make 'u' either true or false, and make all axioms true, and *)
       
  1108             (* add the well-formedness side condition                       *)
       
  1109             val fm_u = (if negate then toFalse else toTrue) (hd intrs)
       
  1110             val fm_ax = Prop_Logic.all (map toTrue (tl intrs))
       
  1111             val fm = Prop_Logic.all [#wellformed args, fm_ax, fm_u]
       
  1112             val _ =
       
  1113               (if satsolver = "dpll" orelse satsolver = "enumerate" then
       
  1114                 warning ("Using SAT solver " ^ quote satsolver ^
       
  1115                          "; for better performance, consider installing an \
       
  1116                          \external solver.")
       
  1117                else ());
       
  1118             val solver =
       
  1119               SatSolver.invoke_solver satsolver
       
  1120               handle Option.Option =>
       
  1121                      error ("Unknown SAT solver: " ^ quote satsolver ^
       
  1122                             ". Available solvers: " ^
       
  1123                             commas (map (quote o fst) (!SatSolver.solvers)) ^ ".")
       
  1124           in
       
  1125             Output.urgent_message "Invoking SAT solver...";
       
  1126             (case solver fm of
       
  1127               SatSolver.SATISFIABLE assignment =>
       
  1128                 (Output.urgent_message ("Model found:\n" ^ print_model ctxt model
       
  1129                   (fn i => case assignment i of SOME b => b | NONE => true));
       
  1130                  if maybe_spurious then "potential" else "genuine")
       
  1131             | SatSolver.UNSATISFIABLE _ =>
       
  1132                 (Output.urgent_message "No model exists.";
       
  1133                 case next_universe universe sizes minsize maxsize of
       
  1134                   SOME universe' => find_model_loop universe'
       
  1135                 | NONE => (Output.urgent_message
       
  1136                     "Search terminated, no larger universe within the given limits.";
       
  1137                     "none"))
       
  1138             | SatSolver.UNKNOWN =>
       
  1139                 (Output.urgent_message "No model found.";
       
  1140                 case next_universe universe sizes minsize maxsize of
       
  1141                   SOME universe' => find_model_loop universe'
       
  1142                 | NONE => (Output.urgent_message
       
  1143                   "Search terminated, no larger universe within the given limits.";
       
  1144                   "unknown"))) handle SatSolver.NOT_CONFIGURED =>
       
  1145               (error ("SAT solver " ^ quote satsolver ^ " is not configured.");
       
  1146                "unknown")
       
  1147           end
       
  1148           handle MAXVARS_EXCEEDED =>
       
  1149             (Output.urgent_message ("Search terminated, number of Boolean variables ("
       
  1150               ^ string_of_int maxvars ^ " allowed) exceeded.");
       
  1151               "unknown")
       
  1152 
       
  1153         val outcome_code = find_model_loop (first_universe types sizes minsize)
       
  1154       in
       
  1155         check_expect outcome_code
       
  1156       end
       
  1157   in
       
  1158     (* some parameter sanity checks *)
       
  1159     minsize>=1 orelse
       
  1160       error ("\"minsize\" is " ^ string_of_int minsize ^ ", must be at least 1");
       
  1161     maxsize>=1 orelse
       
  1162       error ("\"maxsize\" is " ^ string_of_int maxsize ^ ", must be at least 1");
       
  1163     maxsize>=minsize orelse
       
  1164       error ("\"maxsize\" (=" ^ string_of_int maxsize ^
       
  1165       ") is less than \"minsize\" (=" ^ string_of_int minsize ^ ").");
       
  1166     maxvars>=0 orelse
       
  1167       error ("\"maxvars\" is " ^ string_of_int maxvars ^ ", must be at least 0");
       
  1168     maxtime>=0 orelse
       
  1169       error ("\"maxtime\" is " ^ string_of_int maxtime ^ ", must be at least 0");
       
  1170     (* enter loop with or without time limit *)
       
  1171     Output.urgent_message ("Trying to find a model that "
       
  1172       ^ (if negate then "refutes" else "satisfies") ^ ": "
       
  1173       ^ Syntax.string_of_term ctxt t);
       
  1174     if maxtime > 0 then (
       
  1175       TimeLimit.timeLimit (Time.fromSeconds maxtime)
       
  1176         wrapper ()
       
  1177       handle TimeLimit.TimeOut =>
       
  1178         (Output.urgent_message ("Search terminated, time limit (" ^
       
  1179             string_of_int maxtime
       
  1180             ^ (if maxtime=1 then " second" else " seconds") ^ ") exceeded.");
       
  1181          check_expect "unknown")
       
  1182     ) else wrapper ()
       
  1183   end;
       
  1184 
       
  1185 
       
  1186 (* ------------------------------------------------------------------------- *)
       
  1187 (* INTERFACE, PART 2: FINDING A MODEL                                        *)
       
  1188 (* ------------------------------------------------------------------------- *)
       
  1189 
       
  1190 (* ------------------------------------------------------------------------- *)
       
  1191 (* satisfy_term: calls 'find_model' to find a model that satisfies 't'       *)
       
  1192 (* params      : list of '(name, value)' pairs used to override default      *)
       
  1193 (*               parameters                                                  *)
       
  1194 (* ------------------------------------------------------------------------- *)
       
  1195 
       
  1196 fun satisfy_term ctxt params assm_ts t =
       
  1197   find_model ctxt (actual_params ctxt params) assm_ts t false;
       
  1198 
       
  1199 (* ------------------------------------------------------------------------- *)
       
  1200 (* refute_term: calls 'find_model' to find a model that refutes 't'          *)
       
  1201 (* params     : list of '(name, value)' pairs used to override default       *)
       
  1202 (*              parameters                                                   *)
       
  1203 (* ------------------------------------------------------------------------- *)
       
  1204 
       
  1205 fun refute_term ctxt params assm_ts t =
       
  1206   let
       
  1207     (* disallow schematic type variables, since we cannot properly negate  *)
       
  1208     (* terms containing them (their logical meaning is that there EXISTS a *)
       
  1209     (* type s.t. ...; to refute such a formula, we would have to show that *)
       
  1210     (* for ALL types, not ...)                                             *)
       
  1211     val _ = null (Term.add_tvars t []) orelse
       
  1212       error "Term to be refuted contains schematic type variables"
       
  1213 
       
  1214     (* existential closure over schematic variables *)
       
  1215     val vars = sort_wrt (fst o fst) (Term.add_vars t [])
       
  1216     (* Term.term *)
       
  1217     val ex_closure = fold (fn ((x, i), T) => fn t' =>
       
  1218       HOLogic.exists_const T $
       
  1219         Abs (x, T, abstract_over (Var ((x, i), T), t'))) vars t
       
  1220     (* Note: If 't' is of type 'propT' (rather than 'boolT'), applying   *)
       
  1221     (* 'HOLogic.exists_const' is not type-correct.  However, this is not *)
       
  1222     (* really a problem as long as 'find_model' still interprets the     *)
       
  1223     (* resulting term correctly, without checking its type.              *)
       
  1224 
       
  1225     (* replace outermost universally quantified variables by Free's:     *)
       
  1226     (* refuting a term with Free's is generally faster than refuting a   *)
       
  1227     (* term with (nested) quantifiers, because quantifiers are expanded, *)
       
  1228     (* while the SAT solver searches for an interpretation for Free's.   *)
       
  1229     (* Also we get more information back that way, namely an             *)
       
  1230     (* interpretation which includes values for the (formerly)           *)
       
  1231     (* quantified variables.                                             *)
       
  1232     (* maps  !!x1...xn. !xk...xm. t   to   t  *)
       
  1233     fun strip_all_body (Const (@{const_name all}, _) $ Abs (_, _, t)) =
       
  1234           strip_all_body t
       
  1235       | strip_all_body (Const (@{const_name Trueprop}, _) $ t) =
       
  1236           strip_all_body t
       
  1237       | strip_all_body (Const (@{const_name All}, _) $ Abs (_, _, t)) =
       
  1238           strip_all_body t
       
  1239       | strip_all_body t = t
       
  1240     (* maps  !!x1...xn. !xk...xm. t   to   [x1, ..., xn, xk, ..., xm]  *)
       
  1241     fun strip_all_vars (Const (@{const_name all}, _) $ Abs (a, T, t)) =
       
  1242           (a, T) :: strip_all_vars t
       
  1243       | strip_all_vars (Const (@{const_name Trueprop}, _) $ t) =
       
  1244           strip_all_vars t
       
  1245       | strip_all_vars (Const (@{const_name All}, _) $ Abs (a, T, t)) =
       
  1246           (a, T) :: strip_all_vars t
       
  1247       | strip_all_vars _ = [] : (string * typ) list
       
  1248     val strip_t = strip_all_body ex_closure
       
  1249     val frees = Term.rename_wrt_term strip_t (strip_all_vars ex_closure)
       
  1250     val subst_t = Term.subst_bounds (map Free frees, strip_t)
       
  1251   in
       
  1252     find_model ctxt (actual_params ctxt params) assm_ts subst_t true
       
  1253   end;
       
  1254 
       
  1255 (* ------------------------------------------------------------------------- *)
       
  1256 (* refute_goal                                                               *)
       
  1257 (* ------------------------------------------------------------------------- *)
       
  1258 
       
  1259 fun refute_goal ctxt params th i =
       
  1260   let
       
  1261     val t = th |> prop_of
       
  1262   in
       
  1263     if Logic.count_prems t = 0 then
       
  1264       (Output.urgent_message "No subgoal!"; "none")
       
  1265     else
       
  1266       let
       
  1267         val assms = map term_of (Assumption.all_assms_of ctxt)
       
  1268         val (t, frees) = Logic.goal_params t i
       
  1269       in
       
  1270         refute_term ctxt params assms (subst_bounds (frees, t))
       
  1271       end
       
  1272   end
       
  1273 
       
  1274 
       
  1275 (* ------------------------------------------------------------------------- *)
       
  1276 (* INTERPRETERS: Auxiliary Functions                                         *)
       
  1277 (* ------------------------------------------------------------------------- *)
       
  1278 
       
  1279 (* ------------------------------------------------------------------------- *)
       
  1280 (* make_constants: returns all interpretations for type 'T' that consist of  *)
       
  1281 (*                 unit vectors with 'True'/'False' only (no Boolean         *)
       
  1282 (*                 variables)                                                *)
       
  1283 (* ------------------------------------------------------------------------- *)
       
  1284 
       
  1285 fun make_constants ctxt model T =
       
  1286   let
       
  1287     (* returns a list with all unit vectors of length n *)
       
  1288     (* int -> interpretation list *)
       
  1289     fun unit_vectors n =
       
  1290       let
       
  1291         (* returns the k-th unit vector of length n *)
       
  1292         (* int * int -> interpretation *)
       
  1293         fun unit_vector (k, n) =
       
  1294           Leaf ((replicate (k-1) False) @ (True :: (replicate (n-k) False)))
       
  1295         (* int -> interpretation list *)
       
  1296         fun unit_vectors_loop k =
       
  1297           if k>n then [] else unit_vector (k,n) :: unit_vectors_loop (k+1)
       
  1298       in
       
  1299         unit_vectors_loop 1
       
  1300       end
       
  1301     (* returns a list of lists, each one consisting of n (possibly *)
       
  1302     (* identical) elements from 'xs'                               *)
       
  1303     (* int -> 'a list -> 'a list list *)
       
  1304     fun pick_all 1 xs = map single xs
       
  1305       | pick_all n xs =
       
  1306           let val rec_pick = pick_all (n - 1) xs in
       
  1307             maps (fn x => map (cons x) rec_pick) xs
       
  1308           end
       
  1309     (* returns all constant interpretations that have the same tree *)
       
  1310     (* structure as the interpretation argument                     *)
       
  1311     (* interpretation -> interpretation list *)
       
  1312     fun make_constants_intr (Leaf xs) = unit_vectors (length xs)
       
  1313       | make_constants_intr (Node xs) = map Node (pick_all (length xs)
       
  1314           (make_constants_intr (hd xs)))
       
  1315     (* obtain the interpretation for a variable of type 'T' *)
       
  1316     val (i, _, _) = interpret ctxt model {maxvars=0, def_eq=false, next_idx=1,
       
  1317       bounds=[], wellformed=True} (Free ("dummy", T))
       
  1318   in
       
  1319     make_constants_intr i
       
  1320   end;
       
  1321 
       
  1322 (* ------------------------------------------------------------------------- *)
       
  1323 (* size_of_type: returns the number of elements in a type 'T' (i.e. 'length  *)
       
  1324 (*               (make_constants T)', but implemented more efficiently)      *)
       
  1325 (* ------------------------------------------------------------------------- *)
       
  1326 
       
  1327 (* returns 0 for an empty ground type or a function type with empty      *)
       
  1328 (* codomain, but fails for a function type with empty domain --          *)
       
  1329 (* admissibility of datatype constructor argument types (see "Inductive  *)
       
  1330 (* datatypes in HOL - lessons learned ...", S. Berghofer, M. Wenzel,     *)
       
  1331 (* TPHOLs 99) ensures that recursive, possibly empty, datatype fragments *)
       
  1332 (* never occur as the domain of a function type that is the type of a    *)
       
  1333 (* constructor argument                                                  *)
       
  1334 
       
  1335 fun size_of_type ctxt model T =
       
  1336   let
       
  1337     (* returns the number of elements that have the same tree structure as a *)
       
  1338     (* given interpretation                                                  *)
       
  1339     fun size_of_intr (Leaf xs) = length xs
       
  1340       | size_of_intr (Node xs) = Integer.pow (length xs) (size_of_intr (hd xs))
       
  1341     (* obtain the interpretation for a variable of type 'T' *)
       
  1342     val (i, _, _) = interpret ctxt model {maxvars=0, def_eq=false, next_idx=1,
       
  1343       bounds=[], wellformed=True} (Free ("dummy", T))
       
  1344   in
       
  1345     size_of_intr i
       
  1346   end;
       
  1347 
       
  1348 (* ------------------------------------------------------------------------- *)
       
  1349 (* TT/FF: interpretations that denote "true" or "false", respectively        *)
       
  1350 (* ------------------------------------------------------------------------- *)
       
  1351 
       
  1352 (* interpretation *)
       
  1353 
       
  1354 val TT = Leaf [True, False];
       
  1355 
       
  1356 val FF = Leaf [False, True];
       
  1357 
       
  1358 (* ------------------------------------------------------------------------- *)
       
  1359 (* make_equality: returns an interpretation that denotes (extensional)       *)
       
  1360 (*                equality of two interpretations                            *)
       
  1361 (* - two interpretations are 'equal' iff they are both defined and denote    *)
       
  1362 (*   the same value                                                          *)
       
  1363 (* - two interpretations are 'not_equal' iff they are both defined at least  *)
       
  1364 (*   partially, and a defined part denotes different values                  *)
       
  1365 (* - a completely undefined interpretation is neither 'equal' nor            *)
       
  1366 (*   'not_equal' to another interpretation                                   *)
       
  1367 (* ------------------------------------------------------------------------- *)
       
  1368 
       
  1369 (* We could in principle represent '=' on a type T by a particular        *)
       
  1370 (* interpretation.  However, the size of that interpretation is quadratic *)
       
  1371 (* in the size of T.  Therefore comparing the interpretations 'i1' and    *)
       
  1372 (* 'i2' directly is more efficient than constructing the interpretation   *)
       
  1373 (* for equality on T first, and "applying" this interpretation to 'i1'    *)
       
  1374 (* and 'i2' in the usual way (cf. 'interpretation_apply') then.           *)
       
  1375 
       
  1376 (* interpretation * interpretation -> interpretation *)
       
  1377 
       
  1378 fun make_equality (i1, i2) =
       
  1379   let
       
  1380     (* interpretation * interpretation -> prop_formula *)
       
  1381     fun equal (i1, i2) =
       
  1382       (case i1 of
       
  1383         Leaf xs =>
       
  1384           (case i2 of
       
  1385             Leaf ys => Prop_Logic.dot_product (xs, ys)  (* defined and equal *)
       
  1386           | Node _  => raise REFUTE ("make_equality",
       
  1387             "second interpretation is higher"))
       
  1388       | Node xs =>
       
  1389           (case i2 of
       
  1390             Leaf _  => raise REFUTE ("make_equality",
       
  1391             "first interpretation is higher")
       
  1392           | Node ys => Prop_Logic.all (map equal (xs ~~ ys))))
       
  1393     (* interpretation * interpretation -> prop_formula *)
       
  1394     fun not_equal (i1, i2) =
       
  1395       (case i1 of
       
  1396         Leaf xs =>
       
  1397           (case i2 of
       
  1398             (* defined and not equal *)
       
  1399             Leaf ys => Prop_Logic.all ((Prop_Logic.exists xs)
       
  1400             :: (Prop_Logic.exists ys)
       
  1401             :: (map (fn (x,y) => SOr (SNot x, SNot y)) (xs ~~ ys)))
       
  1402           | Node _  => raise REFUTE ("make_equality",
       
  1403             "second interpretation is higher"))
       
  1404       | Node xs =>
       
  1405           (case i2 of
       
  1406             Leaf _  => raise REFUTE ("make_equality",
       
  1407             "first interpretation is higher")
       
  1408           | Node ys => Prop_Logic.exists (map not_equal (xs ~~ ys))))
       
  1409   in
       
  1410     (* a value may be undefined; therefore 'not_equal' is not just the *)
       
  1411     (* negation of 'equal'                                             *)
       
  1412     Leaf [equal (i1, i2), not_equal (i1, i2)]
       
  1413   end;
       
  1414 
       
  1415 (* ------------------------------------------------------------------------- *)
       
  1416 (* make_def_equality: returns an interpretation that denotes (extensional)   *)
       
  1417 (*                    equality of two interpretations                        *)
       
  1418 (* This function treats undefined/partially defined interpretations          *)
       
  1419 (* different from 'make_equality': two undefined interpretations are         *)
       
  1420 (* considered equal, while a defined interpretation is considered not equal  *)
       
  1421 (* to an undefined interpretation.                                           *)
       
  1422 (* ------------------------------------------------------------------------- *)
       
  1423 
       
  1424 (* interpretation * interpretation -> interpretation *)
       
  1425 
       
  1426 fun make_def_equality (i1, i2) =
       
  1427   let
       
  1428     (* interpretation * interpretation -> prop_formula *)
       
  1429     fun equal (i1, i2) =
       
  1430       (case i1 of
       
  1431         Leaf xs =>
       
  1432           (case i2 of
       
  1433             (* defined and equal, or both undefined *)
       
  1434             Leaf ys => SOr (Prop_Logic.dot_product (xs, ys),
       
  1435             SAnd (Prop_Logic.all (map SNot xs), Prop_Logic.all (map SNot ys)))
       
  1436           | Node _  => raise REFUTE ("make_def_equality",
       
  1437             "second interpretation is higher"))
       
  1438       | Node xs =>
       
  1439           (case i2 of
       
  1440             Leaf _  => raise REFUTE ("make_def_equality",
       
  1441             "first interpretation is higher")
       
  1442           | Node ys => Prop_Logic.all (map equal (xs ~~ ys))))
       
  1443     (* interpretation *)
       
  1444     val eq = equal (i1, i2)
       
  1445   in
       
  1446     Leaf [eq, SNot eq]
       
  1447   end;
       
  1448 
       
  1449 (* ------------------------------------------------------------------------- *)
       
  1450 (* interpretation_apply: returns an interpretation that denotes the result   *)
       
  1451 (*                       of applying the function denoted by 'i1' to the     *)
       
  1452 (*                       argument denoted by 'i2'                            *)
       
  1453 (* ------------------------------------------------------------------------- *)
       
  1454 
       
  1455 (* interpretation * interpretation -> interpretation *)
       
  1456 
       
  1457 fun interpretation_apply (i1, i2) =
       
  1458   let
       
  1459     (* interpretation * interpretation -> interpretation *)
       
  1460     fun interpretation_disjunction (tr1,tr2) =
       
  1461       tree_map (fn (xs,ys) => map (fn (x,y) => SOr(x,y)) (xs ~~ ys))
       
  1462         (tree_pair (tr1,tr2))
       
  1463     (* prop_formula * interpretation -> interpretation *)
       
  1464     fun prop_formula_times_interpretation (fm,tr) =
       
  1465       tree_map (map (fn x => SAnd (fm,x))) tr
       
  1466     (* prop_formula list * interpretation list -> interpretation *)
       
  1467     fun prop_formula_list_dot_product_interpretation_list ([fm],[tr]) =
       
  1468           prop_formula_times_interpretation (fm,tr)
       
  1469       | prop_formula_list_dot_product_interpretation_list (fm::fms,tr::trees) =
       
  1470           interpretation_disjunction (prop_formula_times_interpretation (fm,tr),
       
  1471             prop_formula_list_dot_product_interpretation_list (fms,trees))
       
  1472       | prop_formula_list_dot_product_interpretation_list (_,_) =
       
  1473           raise REFUTE ("interpretation_apply", "empty list (in dot product)")
       
  1474     (* returns a list of lists, each one consisting of one element from each *)
       
  1475     (* element of 'xss'                                                      *)
       
  1476     (* 'a list list -> 'a list list *)
       
  1477     fun pick_all [xs] = map single xs
       
  1478       | pick_all (xs::xss) =
       
  1479           let val rec_pick = pick_all xss in
       
  1480             maps (fn x => map (cons x) rec_pick) xs
       
  1481           end
       
  1482       | pick_all _ = raise REFUTE ("interpretation_apply", "empty list (in pick_all)")
       
  1483     (* interpretation -> prop_formula list *)
       
  1484     fun interpretation_to_prop_formula_list (Leaf xs) = xs
       
  1485       | interpretation_to_prop_formula_list (Node trees) =
       
  1486           map Prop_Logic.all (pick_all
       
  1487             (map interpretation_to_prop_formula_list trees))
       
  1488   in
       
  1489     case i1 of
       
  1490       Leaf _ =>
       
  1491         raise REFUTE ("interpretation_apply", "first interpretation is a leaf")
       
  1492     | Node xs =>
       
  1493         prop_formula_list_dot_product_interpretation_list
       
  1494           (interpretation_to_prop_formula_list i2, xs)
       
  1495   end;
       
  1496 
       
  1497 (* ------------------------------------------------------------------------- *)
       
  1498 (* eta_expand: eta-expands a term 't' by adding 'i' lambda abstractions      *)
       
  1499 (* ------------------------------------------------------------------------- *)
       
  1500 
       
  1501 (* Term.term -> int -> Term.term *)
       
  1502 
       
  1503 fun eta_expand t i =
       
  1504   let
       
  1505     val Ts = Term.binder_types (Term.fastype_of t)
       
  1506     val t' = Term.incr_boundvars i t
       
  1507   in
       
  1508     fold_rev (fn T => fn term => Abs ("<eta_expand>", T, term))
       
  1509       (List.take (Ts, i))
       
  1510       (Term.list_comb (t', map Bound (i-1 downto 0)))
       
  1511   end;
       
  1512 
       
  1513 (* ------------------------------------------------------------------------- *)
       
  1514 (* size_of_dtyp: the size of (an initial fragment of) an inductive data type *)
       
  1515 (*               is the sum (over its constructors) of the product (over     *)
       
  1516 (*               their arguments) of the size of the argument types          *)
       
  1517 (* ------------------------------------------------------------------------- *)
       
  1518 
       
  1519 fun size_of_dtyp ctxt typ_sizes descr typ_assoc constructors =
       
  1520   Integer.sum (map (fn (_, dtyps) =>
       
  1521     Integer.prod (map (size_of_type ctxt (typ_sizes, []) o
       
  1522       (typ_of_dtyp descr typ_assoc)) dtyps))
       
  1523         constructors);
       
  1524 
       
  1525 
       
  1526 (* ------------------------------------------------------------------------- *)
       
  1527 (* INTERPRETERS: Actual Interpreters                                         *)
       
  1528 (* ------------------------------------------------------------------------- *)
       
  1529 
       
  1530 (* simply typed lambda calculus: Isabelle's basic term syntax, with type *)
       
  1531 (* variables, function types, and propT                                  *)
       
  1532 
       
  1533 fun stlc_interpreter ctxt model args t =
       
  1534   let
       
  1535     val (typs, terms) = model
       
  1536     val {maxvars, def_eq, next_idx, bounds, wellformed} = args
       
  1537     (* Term.typ -> (interpretation * model * arguments) option *)
       
  1538     fun interpret_groundterm T =
       
  1539       let
       
  1540         (* unit -> (interpretation * model * arguments) option *)
       
  1541         fun interpret_groundtype () =
       
  1542           let
       
  1543             (* the model must specify a size for ground types *)
       
  1544             val size =
       
  1545               if T = Term.propT then 2
       
  1546               else the (AList.lookup (op =) typs T)
       
  1547             val next = next_idx + size
       
  1548             (* check if 'maxvars' is large enough *)
       
  1549             val _ = (if next - 1 > maxvars andalso maxvars > 0 then
       
  1550               raise MAXVARS_EXCEEDED else ())
       
  1551             (* prop_formula list *)
       
  1552             val fms  = map BoolVar (next_idx upto (next_idx + size - 1))
       
  1553             (* interpretation *)
       
  1554             val intr = Leaf fms
       
  1555             (* prop_formula list -> prop_formula *)
       
  1556             fun one_of_two_false [] = True
       
  1557               | one_of_two_false (x::xs) = SAnd (Prop_Logic.all (map (fn x' =>
       
  1558                   SOr (SNot x, SNot x')) xs), one_of_two_false xs)
       
  1559             (* prop_formula *)
       
  1560             val wf = one_of_two_false fms
       
  1561           in
       
  1562             (* extend the model, increase 'next_idx', add well-formedness *)
       
  1563             (* condition                                                  *)
       
  1564             SOME (intr, (typs, (t, intr)::terms), {maxvars = maxvars,
       
  1565               def_eq = def_eq, next_idx = next, bounds = bounds,
       
  1566               wellformed = SAnd (wellformed, wf)})
       
  1567           end
       
  1568       in
       
  1569         case T of
       
  1570           Type ("fun", [T1, T2]) =>
       
  1571             let
       
  1572               (* we create 'size_of_type ... T1' different copies of the        *)
       
  1573               (* interpretation for 'T2', which are then combined into a single *)
       
  1574               (* new interpretation                                             *)
       
  1575               (* make fresh copies, with different variable indices *)
       
  1576               (* 'idx': next variable index                         *)
       
  1577               (* 'n'  : number of copies                            *)
       
  1578               (* int -> int -> (int * interpretation list * prop_formula *)
       
  1579               fun make_copies idx 0 = (idx, [], True)
       
  1580                 | make_copies idx n =
       
  1581                     let
       
  1582                       val (copy, _, new_args) = interpret ctxt (typs, [])
       
  1583                         {maxvars = maxvars, def_eq = false, next_idx = idx,
       
  1584                         bounds = [], wellformed = True} (Free ("dummy", T2))
       
  1585                       val (idx', copies, wf') = make_copies (#next_idx new_args) (n-1)
       
  1586                     in
       
  1587                       (idx', copy :: copies, SAnd (#wellformed new_args, wf'))
       
  1588                     end
       
  1589               val (next, copies, wf) = make_copies next_idx
       
  1590                 (size_of_type ctxt model T1)
       
  1591               (* combine copies into a single interpretation *)
       
  1592               val intr = Node copies
       
  1593             in
       
  1594               (* extend the model, increase 'next_idx', add well-formedness *)
       
  1595               (* condition                                                  *)
       
  1596               SOME (intr, (typs, (t, intr)::terms), {maxvars = maxvars,
       
  1597                 def_eq = def_eq, next_idx = next, bounds = bounds,
       
  1598                 wellformed = SAnd (wellformed, wf)})
       
  1599             end
       
  1600         | Type _  => interpret_groundtype ()
       
  1601         | TFree _ => interpret_groundtype ()
       
  1602         | TVar  _ => interpret_groundtype ()
       
  1603       end
       
  1604   in
       
  1605     case AList.lookup (op =) terms t of
       
  1606       SOME intr =>
       
  1607         (* return an existing interpretation *)
       
  1608         SOME (intr, model, args)
       
  1609     | NONE =>
       
  1610         (case t of
       
  1611           Const (_, T) => interpret_groundterm T
       
  1612         | Free (_, T) => interpret_groundterm T
       
  1613         | Var (_, T) => interpret_groundterm T
       
  1614         | Bound i => SOME (nth (#bounds args) i, model, args)
       
  1615         | Abs (_, T, body) =>
       
  1616             let
       
  1617               (* create all constants of type 'T' *)
       
  1618               val constants = make_constants ctxt model T
       
  1619               (* interpret the 'body' separately for each constant *)
       
  1620               val (bodies, (model', args')) = fold_map
       
  1621                 (fn c => fn (m, a) =>
       
  1622                   let
       
  1623                     (* add 'c' to 'bounds' *)
       
  1624                     val (i', m', a') = interpret ctxt m {maxvars = #maxvars a,
       
  1625                       def_eq = #def_eq a, next_idx = #next_idx a,
       
  1626                       bounds = (c :: #bounds a), wellformed = #wellformed a} body
       
  1627                   in
       
  1628                     (* keep the new model m' and 'next_idx' and 'wellformed', *)
       
  1629                     (* but use old 'bounds'                                   *)
       
  1630                     (i', (m', {maxvars = maxvars, def_eq = def_eq,
       
  1631                       next_idx = #next_idx a', bounds = bounds,
       
  1632                       wellformed = #wellformed a'}))
       
  1633                   end)
       
  1634                 constants (model, args)
       
  1635             in
       
  1636               SOME (Node bodies, model', args')
       
  1637             end
       
  1638         | t1 $ t2 =>
       
  1639             let
       
  1640               (* interpret 't1' and 't2' separately *)
       
  1641               val (intr1, model1, args1) = interpret ctxt model args t1
       
  1642               val (intr2, model2, args2) = interpret ctxt model1 args1 t2
       
  1643             in
       
  1644               SOME (interpretation_apply (intr1, intr2), model2, args2)
       
  1645             end)
       
  1646   end;
       
  1647 
       
  1648 fun Pure_interpreter ctxt model args t =
       
  1649   case t of
       
  1650     Const (@{const_name all}, _) $ t1 =>
       
  1651       let
       
  1652         val (i, m, a) = interpret ctxt model args t1
       
  1653       in
       
  1654         case i of
       
  1655           Node xs =>
       
  1656             (* 3-valued logic *)
       
  1657             let
       
  1658               val fmTrue  = Prop_Logic.all (map toTrue xs)
       
  1659               val fmFalse = Prop_Logic.exists (map toFalse xs)
       
  1660             in
       
  1661               SOME (Leaf [fmTrue, fmFalse], m, a)
       
  1662             end
       
  1663         | _ =>
       
  1664           raise REFUTE ("Pure_interpreter",
       
  1665             "\"all\" is followed by a non-function")
       
  1666       end
       
  1667   | Const (@{const_name all}, _) =>
       
  1668       SOME (interpret ctxt model args (eta_expand t 1))
       
  1669   | Const (@{const_name "=="}, _) $ t1 $ t2 =>
       
  1670       let
       
  1671         val (i1, m1, a1) = interpret ctxt model args t1
       
  1672         val (i2, m2, a2) = interpret ctxt m1 a1 t2
       
  1673       in
       
  1674         (* we use either 'make_def_equality' or 'make_equality' *)
       
  1675         SOME ((if #def_eq args then make_def_equality else make_equality)
       
  1676           (i1, i2), m2, a2)
       
  1677       end
       
  1678   | Const (@{const_name "=="}, _) $ _ =>
       
  1679       SOME (interpret ctxt model args (eta_expand t 1))
       
  1680   | Const (@{const_name "=="}, _) =>
       
  1681       SOME (interpret ctxt model args (eta_expand t 2))
       
  1682   | Const (@{const_name "==>"}, _) $ t1 $ t2 =>
       
  1683       (* 3-valued logic *)
       
  1684       let
       
  1685         val (i1, m1, a1) = interpret ctxt model args t1
       
  1686         val (i2, m2, a2) = interpret ctxt m1 a1 t2
       
  1687         val fmTrue = Prop_Logic.SOr (toFalse i1, toTrue i2)
       
  1688         val fmFalse = Prop_Logic.SAnd (toTrue i1, toFalse i2)
       
  1689       in
       
  1690         SOME (Leaf [fmTrue, fmFalse], m2, a2)
       
  1691       end
       
  1692   | Const (@{const_name "==>"}, _) $ _ =>
       
  1693       SOME (interpret ctxt model args (eta_expand t 1))
       
  1694   | Const (@{const_name "==>"}, _) =>
       
  1695       SOME (interpret ctxt model args (eta_expand t 2))
       
  1696   | _ => NONE;
       
  1697 
       
  1698 fun HOLogic_interpreter ctxt model args t =
       
  1699 (* Providing interpretations directly is more efficient than unfolding the *)
       
  1700 (* logical constants.  In HOL however, logical constants can themselves be *)
       
  1701 (* arguments.  They are then translated using eta-expansion.               *)
       
  1702   case t of
       
  1703     Const (@{const_name Trueprop}, _) =>
       
  1704       SOME (Node [TT, FF], model, args)
       
  1705   | Const (@{const_name Not}, _) =>
       
  1706       SOME (Node [FF, TT], model, args)
       
  1707   (* redundant, since 'True' is also an IDT constructor *)
       
  1708   | Const (@{const_name True}, _) =>
       
  1709       SOME (TT, model, args)
       
  1710   (* redundant, since 'False' is also an IDT constructor *)
       
  1711   | Const (@{const_name False}, _) =>
       
  1712       SOME (FF, model, args)
       
  1713   | Const (@{const_name All}, _) $ t1 =>  (* similar to "all" (Pure) *)
       
  1714       let
       
  1715         val (i, m, a) = interpret ctxt model args t1
       
  1716       in
       
  1717         case i of
       
  1718           Node xs =>
       
  1719             (* 3-valued logic *)
       
  1720             let
       
  1721               val fmTrue = Prop_Logic.all (map toTrue xs)
       
  1722               val fmFalse = Prop_Logic.exists (map toFalse xs)
       
  1723             in
       
  1724               SOME (Leaf [fmTrue, fmFalse], m, a)
       
  1725             end
       
  1726         | _ =>
       
  1727           raise REFUTE ("HOLogic_interpreter",
       
  1728             "\"All\" is followed by a non-function")
       
  1729       end
       
  1730   | Const (@{const_name All}, _) =>
       
  1731       SOME (interpret ctxt model args (eta_expand t 1))
       
  1732   | Const (@{const_name Ex}, _) $ t1 =>
       
  1733       let
       
  1734         val (i, m, a) = interpret ctxt model args t1
       
  1735       in
       
  1736         case i of
       
  1737           Node xs =>
       
  1738             (* 3-valued logic *)
       
  1739             let
       
  1740               val fmTrue = Prop_Logic.exists (map toTrue xs)
       
  1741               val fmFalse = Prop_Logic.all (map toFalse xs)
       
  1742             in
       
  1743               SOME (Leaf [fmTrue, fmFalse], m, a)
       
  1744             end
       
  1745         | _ =>
       
  1746           raise REFUTE ("HOLogic_interpreter",
       
  1747             "\"Ex\" is followed by a non-function")
       
  1748       end
       
  1749   | Const (@{const_name Ex}, _) =>
       
  1750       SOME (interpret ctxt model args (eta_expand t 1))
       
  1751   | Const (@{const_name HOL.eq}, _) $ t1 $ t2 =>  (* similar to "==" (Pure) *)
       
  1752       let
       
  1753         val (i1, m1, a1) = interpret ctxt model args t1
       
  1754         val (i2, m2, a2) = interpret ctxt m1 a1 t2
       
  1755       in
       
  1756         SOME (make_equality (i1, i2), m2, a2)
       
  1757       end
       
  1758   | Const (@{const_name HOL.eq}, _) $ _ =>
       
  1759       SOME (interpret ctxt model args (eta_expand t 1))
       
  1760   | Const (@{const_name HOL.eq}, _) =>
       
  1761       SOME (interpret ctxt model args (eta_expand t 2))
       
  1762   | Const (@{const_name HOL.conj}, _) $ t1 $ t2 =>
       
  1763       (* 3-valued logic *)
       
  1764       let
       
  1765         val (i1, m1, a1) = interpret ctxt model args t1
       
  1766         val (i2, m2, a2) = interpret ctxt m1 a1 t2
       
  1767         val fmTrue = Prop_Logic.SAnd (toTrue i1, toTrue i2)
       
  1768         val fmFalse = Prop_Logic.SOr (toFalse i1, toFalse i2)
       
  1769       in
       
  1770         SOME (Leaf [fmTrue, fmFalse], m2, a2)
       
  1771       end
       
  1772   | Const (@{const_name HOL.conj}, _) $ _ =>
       
  1773       SOME (interpret ctxt model args (eta_expand t 1))
       
  1774   | Const (@{const_name HOL.conj}, _) =>
       
  1775       SOME (interpret ctxt model args (eta_expand t 2))
       
  1776       (* this would make "undef" propagate, even for formulae like *)
       
  1777       (* "False & undef":                                          *)
       
  1778       (* SOME (Node [Node [TT, FF], Node [FF, FF]], model, args) *)
       
  1779   | Const (@{const_name HOL.disj}, _) $ t1 $ t2 =>
       
  1780       (* 3-valued logic *)
       
  1781       let
       
  1782         val (i1, m1, a1) = interpret ctxt model args t1
       
  1783         val (i2, m2, a2) = interpret ctxt m1 a1 t2
       
  1784         val fmTrue = Prop_Logic.SOr (toTrue i1, toTrue i2)
       
  1785         val fmFalse = Prop_Logic.SAnd (toFalse i1, toFalse i2)
       
  1786       in
       
  1787         SOME (Leaf [fmTrue, fmFalse], m2, a2)
       
  1788       end
       
  1789   | Const (@{const_name HOL.disj}, _) $ _ =>
       
  1790       SOME (interpret ctxt model args (eta_expand t 1))
       
  1791   | Const (@{const_name HOL.disj}, _) =>
       
  1792       SOME (interpret ctxt model args (eta_expand t 2))
       
  1793       (* this would make "undef" propagate, even for formulae like *)
       
  1794       (* "True | undef":                                           *)
       
  1795       (* SOME (Node [Node [TT, TT], Node [TT, FF]], model, args) *)
       
  1796   | Const (@{const_name HOL.implies}, _) $ t1 $ t2 =>  (* similar to "==>" (Pure) *)
       
  1797       (* 3-valued logic *)
       
  1798       let
       
  1799         val (i1, m1, a1) = interpret ctxt model args t1
       
  1800         val (i2, m2, a2) = interpret ctxt m1 a1 t2
       
  1801         val fmTrue = Prop_Logic.SOr (toFalse i1, toTrue i2)
       
  1802         val fmFalse = Prop_Logic.SAnd (toTrue i1, toFalse i2)
       
  1803       in
       
  1804         SOME (Leaf [fmTrue, fmFalse], m2, a2)
       
  1805       end
       
  1806   | Const (@{const_name HOL.implies}, _) $ _ =>
       
  1807       SOME (interpret ctxt model args (eta_expand t 1))
       
  1808   | Const (@{const_name HOL.implies}, _) =>
       
  1809       SOME (interpret ctxt model args (eta_expand t 2))
       
  1810       (* this would make "undef" propagate, even for formulae like *)
       
  1811       (* "False --> undef":                                        *)
       
  1812       (* SOME (Node [Node [TT, FF], Node [TT, TT]], model, args) *)
       
  1813   | _ => NONE;
       
  1814 
       
  1815 (* interprets variables and constants whose type is an IDT (this is        *)
       
  1816 (* relatively easy and merely requires us to compute the size of the IDT); *)
       
  1817 (* constructors of IDTs however are properly interpreted by                *)
       
  1818 (* 'IDT_constructor_interpreter'                                           *)
       
  1819 
       
  1820 fun IDT_interpreter ctxt model args t =
       
  1821   let
       
  1822     val thy = Proof_Context.theory_of ctxt
       
  1823     val (typs, terms) = model
       
  1824     (* Term.typ -> (interpretation * model * arguments) option *)
       
  1825     fun interpret_term (Type (s, Ts)) =
       
  1826           (case Datatype.get_info thy s of
       
  1827             SOME info =>  (* inductive datatype *)
       
  1828               let
       
  1829                 (* int option -- only recursive IDTs have an associated depth *)
       
  1830                 val depth = AList.lookup (op =) typs (Type (s, Ts))
       
  1831                 (* sanity check: depth must be at least 0 *)
       
  1832                 val _ =
       
  1833                   (case depth of SOME n =>
       
  1834                     if n < 0 then
       
  1835                       raise REFUTE ("IDT_interpreter", "negative depth")
       
  1836                     else ()
       
  1837                   | _ => ())
       
  1838               in
       
  1839                 (* termination condition to avoid infinite recursion *)
       
  1840                 if depth = (SOME 0) then
       
  1841                   (* return a leaf of size 0 *)
       
  1842                   SOME (Leaf [], model, args)
       
  1843                 else
       
  1844                   let
       
  1845                     val index               = #index info
       
  1846                     val descr               = #descr info
       
  1847                     val (_, dtyps, constrs) = the (AList.lookup (op =) descr index)
       
  1848                     val typ_assoc           = dtyps ~~ Ts
       
  1849                     (* sanity check: every element in 'dtyps' must be a 'DtTFree' *)
       
  1850                     val _ =
       
  1851                       if Library.exists (fn d =>
       
  1852                         case d of Datatype.DtTFree _ => false | _ => true) dtyps
       
  1853                       then
       
  1854                         raise REFUTE ("IDT_interpreter",
       
  1855                           "datatype argument (for type "
       
  1856                           ^ Syntax.string_of_typ ctxt (Type (s, Ts))
       
  1857                           ^ ") is not a variable")
       
  1858                       else ()
       
  1859                     (* if the model specifies a depth for the current type, *)
       
  1860                     (* decrement it to avoid infinite recursion             *)
       
  1861                     val typs' = case depth of NONE => typs | SOME n =>
       
  1862                       AList.update (op =) (Type (s, Ts), n-1) typs
       
  1863                     (* recursively compute the size of the datatype *)
       
  1864                     val size     = size_of_dtyp ctxt typs' descr typ_assoc constrs
       
  1865                     val next_idx = #next_idx args
       
  1866                     val next     = next_idx+size
       
  1867                     (* check if 'maxvars' is large enough *)
       
  1868                     val _        = (if next-1 > #maxvars args andalso
       
  1869                       #maxvars args > 0 then raise MAXVARS_EXCEEDED else ())
       
  1870                     (* prop_formula list *)
       
  1871                     val fms      = map BoolVar (next_idx upto (next_idx+size-1))
       
  1872                     (* interpretation *)
       
  1873                     val intr     = Leaf fms
       
  1874                     (* prop_formula list -> prop_formula *)
       
  1875                     fun one_of_two_false [] = True
       
  1876                       | one_of_two_false (x::xs) = SAnd (Prop_Logic.all (map (fn x' =>
       
  1877                           SOr (SNot x, SNot x')) xs), one_of_two_false xs)
       
  1878                     (* prop_formula *)
       
  1879                     val wf = one_of_two_false fms
       
  1880                   in
       
  1881                     (* extend the model, increase 'next_idx', add well-formedness *)
       
  1882                     (* condition                                                  *)
       
  1883                     SOME (intr, (typs, (t, intr)::terms), {maxvars = #maxvars args,
       
  1884                       def_eq = #def_eq args, next_idx = next, bounds = #bounds args,
       
  1885                       wellformed = SAnd (#wellformed args, wf)})
       
  1886                   end
       
  1887               end
       
  1888           | NONE =>  (* not an inductive datatype *)
       
  1889               NONE)
       
  1890       | interpret_term _ =  (* a (free or schematic) type variable *)
       
  1891           NONE
       
  1892   in
       
  1893     case AList.lookup (op =) terms t of
       
  1894       SOME intr =>
       
  1895         (* return an existing interpretation *)
       
  1896         SOME (intr, model, args)
       
  1897     | NONE =>
       
  1898         (case t of
       
  1899           Free (_, T) => interpret_term T
       
  1900         | Var (_, T) => interpret_term T
       
  1901         | Const (_, T) => interpret_term T
       
  1902         | _ => NONE)
       
  1903   end;
       
  1904 
       
  1905 (* This function imposes an order on the elements of a datatype fragment  *)
       
  1906 (* as follows: C_i x_1 ... x_n < C_j y_1 ... y_m iff i < j or             *)
       
  1907 (* (x_1, ..., x_n) < (y_1, ..., y_m).  With this order, a constructor is  *)
       
  1908 (* a function C_i that maps some argument indices x_1, ..., x_n to the    *)
       
  1909 (* datatype element given by index C_i x_1 ... x_n.  The idea remains the *)
       
  1910 (* same for recursive datatypes, although the computation of indices gets *)
       
  1911 (* a little tricky.                                                       *)
       
  1912 
       
  1913 fun IDT_constructor_interpreter ctxt model args t =
       
  1914   let
       
  1915     val thy = Proof_Context.theory_of ctxt
       
  1916     (* returns a list of canonical representations for terms of the type 'T' *)
       
  1917     (* It would be nice if we could just use 'print' for this, but 'print'   *)
       
  1918     (* for IDTs calls 'IDT_constructor_interpreter' again, and this could    *)
       
  1919     (* lead to infinite recursion when we have (mutually) recursive IDTs.    *)
       
  1920     (* (Term.typ * int) list -> Term.typ -> Term.term list *)
       
  1921     fun canonical_terms typs T =
       
  1922           (case T of
       
  1923             Type ("fun", [T1, T2]) =>
       
  1924             (* 'T2' might contain a recursive IDT, so we cannot use 'print' (at *)
       
  1925             (* least not for 'T2'                                               *)
       
  1926             let
       
  1927               (* returns a list of lists, each one consisting of n (possibly *)
       
  1928               (* identical) elements from 'xs'                               *)
       
  1929               (* int -> 'a list -> 'a list list *)
       
  1930               fun pick_all 1 xs = map single xs
       
  1931                 | pick_all n xs =
       
  1932                     let val rec_pick = pick_all (n-1) xs in
       
  1933                       maps (fn x => map (cons x) rec_pick) xs
       
  1934                     end
       
  1935               (* ["x1", ..., "xn"] *)
       
  1936               val terms1 = canonical_terms typs T1
       
  1937               (* ["y1", ..., "ym"] *)
       
  1938               val terms2 = canonical_terms typs T2
       
  1939               (* [[("x1", "y1"), ..., ("xn", "y1")], ..., *)
       
  1940               (*   [("x1", "ym"), ..., ("xn", "ym")]]     *)
       
  1941               val functions = map (curry (op ~~) terms1)
       
  1942                 (pick_all (length terms1) terms2)
       
  1943               (* [["(x1, y1)", ..., "(xn, y1)"], ..., *)
       
  1944               (*   ["(x1, ym)", ..., "(xn, ym)"]]     *)
       
  1945               val pairss = map (map HOLogic.mk_prod) functions
       
  1946               (* Term.typ *)
       
  1947               val HOLogic_prodT = HOLogic.mk_prodT (T1, T2)
       
  1948               val HOLogic_setT  = HOLogic.mk_setT HOLogic_prodT
       
  1949               (* Term.term *)
       
  1950               val HOLogic_empty_set = Const (@{const_abbrev Set.empty}, HOLogic_setT)
       
  1951               val HOLogic_insert    =
       
  1952                 Const (@{const_name insert}, HOLogic_prodT --> HOLogic_setT --> HOLogic_setT)
       
  1953             in
       
  1954               (* functions as graphs, i.e. as a (HOL) set of pairs "(x, y)" *)
       
  1955               map (fn ps => fold_rev (fn pair => fn acc => HOLogic_insert $ pair $ acc) ps
       
  1956                 HOLogic_empty_set) pairss
       
  1957             end
       
  1958       | Type (s, Ts) =>
       
  1959           (case Datatype.get_info thy s of
       
  1960             SOME info =>
       
  1961               (case AList.lookup (op =) typs T of
       
  1962                 SOME 0 =>
       
  1963                   (* termination condition to avoid infinite recursion *)
       
  1964                   []  (* at depth 0, every IDT is empty *)
       
  1965               | _ =>
       
  1966                 let
       
  1967                   val index = #index info
       
  1968                   val descr = #descr info
       
  1969                   val (_, dtyps, constrs) = the (AList.lookup (op =) descr index)
       
  1970                   val typ_assoc = dtyps ~~ Ts
       
  1971                   (* sanity check: every element in 'dtyps' must be a 'DtTFree' *)
       
  1972                   val _ =
       
  1973                     if Library.exists (fn d =>
       
  1974                       case d of Datatype.DtTFree _ => false | _ => true) dtyps
       
  1975                     then
       
  1976                       raise REFUTE ("IDT_constructor_interpreter",
       
  1977                         "datatype argument (for type "
       
  1978                         ^ Syntax.string_of_typ ctxt T
       
  1979                         ^ ") is not a variable")
       
  1980                     else ()
       
  1981                   (* decrement depth for the IDT 'T' *)
       
  1982                   val typs' =
       
  1983                     (case AList.lookup (op =) typs T of NONE => typs
       
  1984                     | SOME n => AList.update (op =) (T, n-1) typs)
       
  1985                   fun constructor_terms terms [] = terms
       
  1986                     | constructor_terms terms (d::ds) =
       
  1987                         let
       
  1988                           val dT = typ_of_dtyp descr typ_assoc d
       
  1989                           val d_terms = canonical_terms typs' dT
       
  1990                         in
       
  1991                           (* C_i x_1 ... x_n < C_i y_1 ... y_n if *)
       
  1992                           (* (x_1, ..., x_n) < (y_1, ..., y_n)    *)
       
  1993                           constructor_terms
       
  1994                             (map_product (curry op $) terms d_terms) ds
       
  1995                         end
       
  1996                 in
       
  1997                   (* C_i ... < C_j ... if i < j *)
       
  1998                   maps (fn (cname, ctyps) =>
       
  1999                     let
       
  2000                       val cTerm = Const (cname,
       
  2001                         map (typ_of_dtyp descr typ_assoc) ctyps ---> T)
       
  2002                     in
       
  2003                       constructor_terms [cTerm] ctyps
       
  2004                     end) constrs
       
  2005                 end)
       
  2006           | NONE =>
       
  2007               (* not an inductive datatype; in this case the argument types in *)
       
  2008               (* 'Ts' may not be IDTs either, so 'print' should be safe        *)
       
  2009               map (fn intr => print ctxt (typs, []) T intr (K false))
       
  2010                 (make_constants ctxt (typs, []) T))
       
  2011       | _ =>  (* TFree ..., TVar ... *)
       
  2012           map (fn intr => print ctxt (typs, []) T intr (K false))
       
  2013             (make_constants ctxt (typs, []) T))
       
  2014     val (typs, terms) = model
       
  2015   in
       
  2016     case AList.lookup (op =) terms t of
       
  2017       SOME intr =>
       
  2018         (* return an existing interpretation *)
       
  2019         SOME (intr, model, args)
       
  2020     | NONE =>
       
  2021         (case t of
       
  2022           Const (s, T) =>
       
  2023             (case body_type T of
       
  2024               Type (s', Ts') =>
       
  2025                 (case Datatype.get_info thy s' of
       
  2026                   SOME info =>  (* body type is an inductive datatype *)
       
  2027                     let
       
  2028                       val index               = #index info
       
  2029                       val descr               = #descr info
       
  2030                       val (_, dtyps, constrs) = the (AList.lookup (op =) descr index)
       
  2031                       val typ_assoc           = dtyps ~~ Ts'
       
  2032                       (* sanity check: every element in 'dtyps' must be a 'DtTFree' *)
       
  2033                       val _ = if Library.exists (fn d =>
       
  2034                           case d of Datatype.DtTFree _ => false | _ => true) dtyps
       
  2035                         then
       
  2036                           raise REFUTE ("IDT_constructor_interpreter",
       
  2037                             "datatype argument (for type "
       
  2038                             ^ Syntax.string_of_typ ctxt (Type (s', Ts'))
       
  2039                             ^ ") is not a variable")
       
  2040                         else ()
       
  2041                       (* split the constructors into those occuring before/after *)
       
  2042                       (* 'Const (s, T)'                                          *)
       
  2043                       val (constrs1, constrs2) = take_prefix (fn (cname, ctypes) =>
       
  2044                         not (cname = s andalso Sign.typ_instance thy (T,
       
  2045                           map (typ_of_dtyp descr typ_assoc) ctypes
       
  2046                             ---> Type (s', Ts')))) constrs
       
  2047                     in
       
  2048                       case constrs2 of
       
  2049                         [] =>
       
  2050                           (* 'Const (s, T)' is not a constructor of this datatype *)
       
  2051                           NONE
       
  2052                       | (_, ctypes)::_ =>
       
  2053                           let
       
  2054                             (* int option -- only /recursive/ IDTs have an associated *)
       
  2055                             (*               depth                                    *)
       
  2056                             val depth = AList.lookup (op =) typs (Type (s', Ts'))
       
  2057                             (* this should never happen: at depth 0, this IDT fragment *)
       
  2058                             (* is definitely empty, and in this case we don't need to  *)
       
  2059                             (* interpret its constructors                              *)
       
  2060                             val _ = (case depth of SOME 0 =>
       
  2061                                 raise REFUTE ("IDT_constructor_interpreter",
       
  2062                                   "depth is 0")
       
  2063                               | _ => ())
       
  2064                             val typs' = (case depth of NONE => typs | SOME n =>
       
  2065                               AList.update (op =) (Type (s', Ts'), n-1) typs)
       
  2066                             (* elements of the datatype come before elements generated *)
       
  2067                             (* by 'Const (s, T)' iff they are generated by a           *)
       
  2068                             (* constructor in constrs1                                 *)
       
  2069                             val offset = size_of_dtyp ctxt typs' descr typ_assoc constrs1
       
  2070                             (* compute the total (current) size of the datatype *)
       
  2071                             val total = offset +
       
  2072                               size_of_dtyp ctxt typs' descr typ_assoc constrs2
       
  2073                             (* sanity check *)
       
  2074                             val _ = if total <> size_of_type ctxt (typs, [])
       
  2075                               (Type (s', Ts')) then
       
  2076                                 raise REFUTE ("IDT_constructor_interpreter",
       
  2077                                   "total is not equal to current size")
       
  2078                               else ()
       
  2079                             (* returns an interpretation where everything is mapped to *)
       
  2080                             (* an "undefined" element of the datatype                  *)
       
  2081                             fun make_undef [] = Leaf (replicate total False)
       
  2082                               | make_undef (d::ds) =
       
  2083                                   let
       
  2084                                     (* compute the current size of the type 'd' *)
       
  2085                                     val dT   = typ_of_dtyp descr typ_assoc d
       
  2086                                     val size = size_of_type ctxt (typs, []) dT
       
  2087                                   in
       
  2088                                     Node (replicate size (make_undef ds))
       
  2089                                   end
       
  2090                             (* returns the interpretation for a constructor *)
       
  2091                             fun make_constr [] offset =
       
  2092                                   if offset < total then
       
  2093                                     (Leaf (replicate offset False @ True ::
       
  2094                                       (replicate (total - offset - 1) False)), offset + 1)
       
  2095                                   else
       
  2096                                     raise REFUTE ("IDT_constructor_interpreter",
       
  2097                                       "offset >= total")
       
  2098                               | make_constr (d::ds) offset =
       
  2099                                   let
       
  2100                                     (* Term.typ *)
       
  2101                                     val dT = typ_of_dtyp descr typ_assoc d
       
  2102                                     (* compute canonical term representations for all   *)
       
  2103                                     (* elements of the type 'd' (with the reduced depth *)
       
  2104                                     (* for the IDT)                                     *)
       
  2105                                     val terms' = canonical_terms typs' dT
       
  2106                                     (* sanity check *)
       
  2107                                     val _ =
       
  2108                                       if length terms' <> size_of_type ctxt (typs', []) dT
       
  2109                                       then
       
  2110                                         raise REFUTE ("IDT_constructor_interpreter",
       
  2111                                           "length of terms' is not equal to old size")
       
  2112                                       else ()
       
  2113                                     (* compute canonical term representations for all   *)
       
  2114                                     (* elements of the type 'd' (with the current depth *)
       
  2115                                     (* for the IDT)                                     *)
       
  2116                                     val terms = canonical_terms typs dT
       
  2117                                     (* sanity check *)
       
  2118                                     val _ =
       
  2119                                       if length terms <> size_of_type ctxt (typs, []) dT
       
  2120                                       then
       
  2121                                         raise REFUTE ("IDT_constructor_interpreter",
       
  2122                                           "length of terms is not equal to current size")
       
  2123                                       else ()
       
  2124                                     (* sanity check *)
       
  2125                                     val _ =
       
  2126                                       if length terms < length terms' then
       
  2127                                         raise REFUTE ("IDT_constructor_interpreter",
       
  2128                                           "current size is less than old size")
       
  2129                                       else ()
       
  2130                                     (* sanity check: every element of terms' must also be *)
       
  2131                                     (*               present in terms                     *)
       
  2132                                     val _ =
       
  2133                                       if forall (member (op =) terms) terms' then ()
       
  2134                                       else
       
  2135                                         raise REFUTE ("IDT_constructor_interpreter",
       
  2136                                           "element has disappeared")
       
  2137                                     (* sanity check: the order on elements of terms' is    *)
       
  2138                                     (*               the same in terms, for those elements *)
       
  2139                                     val _ =
       
  2140                                       let
       
  2141                                         fun search (x::xs) (y::ys) =
       
  2142                                               if x = y then search xs ys else search (x::xs) ys
       
  2143                                           | search (_::_) [] =
       
  2144                                               raise REFUTE ("IDT_constructor_interpreter",
       
  2145                                                 "element order not preserved")
       
  2146                                           | search [] _ = ()
       
  2147                                       in search terms' terms end
       
  2148                                     (* int * interpretation list *)
       
  2149                                     val (intrs, new_offset) =
       
  2150                                       fold_map (fn t_elem => fn off =>
       
  2151                                         (* if 't_elem' existed at the previous depth,    *)
       
  2152                                         (* proceed recursively, otherwise map the entire *)
       
  2153                                         (* subtree to "undefined"                        *)
       
  2154                                         if member (op =) terms' t_elem then
       
  2155                                           make_constr ds off
       
  2156                                         else
       
  2157                                           (make_undef ds, off))
       
  2158                                       terms offset
       
  2159                                   in
       
  2160                                     (Node intrs, new_offset)
       
  2161                                   end
       
  2162                           in
       
  2163                             SOME (fst (make_constr ctypes offset), model, args)
       
  2164                           end
       
  2165                     end
       
  2166                 | NONE =>  (* body type is not an inductive datatype *)
       
  2167                     NONE)
       
  2168             | _ =>  (* body type is a (free or schematic) type variable *)
       
  2169               NONE)
       
  2170         | _ =>  (* term is not a constant *)
       
  2171           NONE)
       
  2172   end;
       
  2173 
       
  2174 (* Difficult code ahead.  Make sure you understand the                *)
       
  2175 (* 'IDT_constructor_interpreter' and the order in which it enumerates *)
       
  2176 (* elements of an IDT before you try to understand this function.     *)
       
  2177 
       
  2178 fun IDT_recursion_interpreter ctxt model args t =
       
  2179   let
       
  2180     val thy = Proof_Context.theory_of ctxt
       
  2181   in
       
  2182     (* careful: here we descend arbitrarily deep into 't', possibly before *)
       
  2183     (* any other interpreter for atomic terms has had a chance to look at  *)
       
  2184     (* 't'                                                                 *)
       
  2185     case strip_comb t of
       
  2186       (Const (s, T), params) =>
       
  2187         (* iterate over all datatypes in 'thy' *)
       
  2188         Symtab.fold (fn (_, info) => fn result =>
       
  2189           case result of
       
  2190             SOME _ =>
       
  2191               result  (* just keep 'result' *)
       
  2192           | NONE =>
       
  2193               if member (op =) (#rec_names info) s then
       
  2194                 (* we do have a recursion operator of one of the (mutually *)
       
  2195                 (* recursive) datatypes given by 'info'                    *)
       
  2196                 let
       
  2197                   (* number of all constructors, including those of different  *)
       
  2198                   (* (mutually recursive) datatypes within the same descriptor *)
       
  2199                   val mconstrs_count =
       
  2200                     Integer.sum (map (fn (_, (_, _, cs)) => length cs) (#descr info))
       
  2201                 in
       
  2202                   if mconstrs_count < length params then
       
  2203                     (* too many actual parameters; for now we'll use the *)
       
  2204                     (* 'stlc_interpreter' to strip off one application   *)
       
  2205                     NONE
       
  2206                   else if mconstrs_count > length params then
       
  2207                     (* too few actual parameters; we use eta expansion          *)
       
  2208                     (* Note that the resulting expansion of lambda abstractions *)
       
  2209                     (* by the 'stlc_interpreter' may be rather slow (depending  *)
       
  2210                     (* on the argument types and the size of the IDT, of        *)
       
  2211                     (* course).                                                 *)
       
  2212                     SOME (interpret ctxt model args (eta_expand t
       
  2213                       (mconstrs_count - length params)))
       
  2214                   else  (* mconstrs_count = length params *)
       
  2215                     let
       
  2216                       (* interpret each parameter separately *)
       
  2217                       val (p_intrs, (model', args')) = fold_map (fn p => fn (m, a) =>
       
  2218                         let
       
  2219                           val (i, m', a') = interpret ctxt m a p
       
  2220                         in
       
  2221                           (i, (m', a'))
       
  2222                         end) params (model, args)
       
  2223                       val (typs, _) = model'
       
  2224                       (* 'index' is /not/ necessarily the index of the IDT that *)
       
  2225                       (* the recursion operator is associated with, but merely  *)
       
  2226                       (* the index of some mutually recursive IDT               *)
       
  2227                       val index         = #index info
       
  2228                       val descr         = #descr info
       
  2229                       val (_, dtyps, _) = the (AList.lookup (op =) descr index)
       
  2230                       (* sanity check: we assume that the order of constructors *)
       
  2231                       (*               in 'descr' is the same as the order of   *)
       
  2232                       (*               corresponding parameters, otherwise the  *)
       
  2233                       (*               association code below won't match the   *)
       
  2234                       (*               right constructors/parameters; we also   *)
       
  2235                       (*               assume that the order of recursion       *)
       
  2236                       (*               operators in '#rec_names info' is the    *)
       
  2237                       (*               same as the order of corresponding       *)
       
  2238                       (*               datatypes in 'descr'                     *)
       
  2239                       val _ = if map fst descr <> (0 upto (length descr - 1)) then
       
  2240                           raise REFUTE ("IDT_recursion_interpreter",
       
  2241                             "order of constructors and corresponding parameters/" ^
       
  2242                               "recursion operators and corresponding datatypes " ^
       
  2243                               "different?")
       
  2244                         else ()
       
  2245                       (* sanity check: every element in 'dtyps' must be a *)
       
  2246                       (*               'DtTFree'                          *)
       
  2247                       val _ =
       
  2248                         if Library.exists (fn d =>
       
  2249                           case d of Datatype.DtTFree _ => false
       
  2250                                   | _ => true) dtyps
       
  2251                         then
       
  2252                           raise REFUTE ("IDT_recursion_interpreter",
       
  2253                             "datatype argument is not a variable")
       
  2254                         else ()
       
  2255                       (* the type of a recursion operator is *)
       
  2256                       (* [T1, ..., Tn, IDT] ---> Tresult     *)
       
  2257                       val IDT = nth (binder_types T) mconstrs_count
       
  2258                       (* by our assumption on the order of recursion operators *)
       
  2259                       (* and datatypes, this is the index of the datatype      *)
       
  2260                       (* corresponding to the given recursion operator         *)
       
  2261                       val idt_index = find_index (fn s' => s' = s) (#rec_names info)
       
  2262                       (* mutually recursive types must have the same type   *)
       
  2263                       (* parameters, unless the mutual recursion comes from *)
       
  2264                       (* indirect recursion                                 *)
       
  2265                       fun rec_typ_assoc acc [] = acc
       
  2266                         | rec_typ_assoc acc ((d, T)::xs) =
       
  2267                             (case AList.lookup op= acc d of
       
  2268                               NONE =>
       
  2269                                 (case d of
       
  2270                                   Datatype.DtTFree _ =>
       
  2271                                   (* add the association, proceed *)
       
  2272                                   rec_typ_assoc ((d, T)::acc) xs
       
  2273                                 | Datatype.DtType (s, ds) =>
       
  2274                                     let
       
  2275                                       val (s', Ts) = dest_Type T
       
  2276                                     in
       
  2277                                       if s=s' then
       
  2278                                         rec_typ_assoc ((d, T)::acc) ((ds ~~ Ts) @ xs)
       
  2279                                       else
       
  2280                                         raise REFUTE ("IDT_recursion_interpreter",
       
  2281                                           "DtType/Type mismatch")
       
  2282                                     end
       
  2283                                 | Datatype.DtRec i =>
       
  2284                                     let
       
  2285                                       val (_, ds, _) = the (AList.lookup (op =) descr i)
       
  2286                                       val (_, Ts)    = dest_Type T
       
  2287                                     in
       
  2288                                       rec_typ_assoc ((d, T)::acc) ((ds ~~ Ts) @ xs)
       
  2289                                     end)
       
  2290                             | SOME T' =>
       
  2291                                 if T=T' then
       
  2292                                   (* ignore the association since it's already *)
       
  2293                                   (* present, proceed                          *)
       
  2294                                   rec_typ_assoc acc xs
       
  2295                                 else
       
  2296                                   raise REFUTE ("IDT_recursion_interpreter",
       
  2297                                     "different type associations for the same dtyp"))
       
  2298                       val typ_assoc = filter
       
  2299                         (fn (Datatype.DtTFree _, _) => true | (_, _) => false)
       
  2300                         (rec_typ_assoc []
       
  2301                           (#2 (the (AList.lookup (op =) descr idt_index)) ~~ (snd o dest_Type) IDT))
       
  2302                       (* sanity check: typ_assoc must associate types to the   *)
       
  2303                       (*               elements of 'dtyps' (and only to those) *)
       
  2304                       val _ =
       
  2305                         if not (eq_set (op =) (dtyps, map fst typ_assoc))
       
  2306                         then
       
  2307                           raise REFUTE ("IDT_recursion_interpreter",
       
  2308                             "type association has extra/missing elements")
       
  2309                         else ()
       
  2310                       (* interpret each constructor in the descriptor (including *)
       
  2311                       (* those of mutually recursive datatypes)                  *)
       
  2312                       (* (int * interpretation list) list *)
       
  2313                       val mc_intrs = map (fn (idx, (_, _, cs)) =>
       
  2314                         let
       
  2315                           val c_return_typ = typ_of_dtyp descr typ_assoc
       
  2316                             (Datatype.DtRec idx)
       
  2317                         in
       
  2318                           (idx, map (fn (cname, cargs) =>
       
  2319                             (#1 o interpret ctxt (typs, []) {maxvars=0,
       
  2320                               def_eq=false, next_idx=1, bounds=[],
       
  2321                               wellformed=True}) (Const (cname, map (typ_of_dtyp
       
  2322                               descr typ_assoc) cargs ---> c_return_typ))) cs)
       
  2323                         end) descr
       
  2324                       (* associate constructors with corresponding parameters *)
       
  2325                       (* (int * (interpretation * interpretation) list) list *)
       
  2326                       val (mc_p_intrs, p_intrs') = fold_map
       
  2327                         (fn (idx, c_intrs) => fn p_intrs' =>
       
  2328                           let
       
  2329                             val len = length c_intrs
       
  2330                           in
       
  2331                             ((idx, c_intrs ~~ List.take (p_intrs', len)),
       
  2332                               List.drop (p_intrs', len))
       
  2333                           end) mc_intrs p_intrs
       
  2334                       (* sanity check: no 'p_intr' may be left afterwards *)
       
  2335                       val _ =
       
  2336                         if p_intrs' <> [] then
       
  2337                           raise REFUTE ("IDT_recursion_interpreter",
       
  2338                             "more parameter than constructor interpretations")
       
  2339                         else ()
       
  2340                       (* The recursion operator, applied to 'mconstrs_count'     *)
       
  2341                       (* arguments, is a function that maps every element of the *)
       
  2342                       (* inductive datatype to an element of some result type.   *)
       
  2343                       (* Recursion operators for mutually recursive IDTs are     *)
       
  2344                       (* translated simultaneously.                              *)
       
  2345                       (* Since the order on datatype elements is given by an     *)
       
  2346                       (* order on constructors (and then by the order on         *)
       
  2347                       (* argument tuples), we can simply copy corresponding      *)
       
  2348                       (* subtrees from 'p_intrs', in the order in which they are *)
       
  2349                       (* given.                                                  *)
       
  2350                       (* interpretation * interpretation -> interpretation list *)
       
  2351                       fun ci_pi (Leaf xs, pi) =
       
  2352                             (* if the constructor does not match the arguments to a *)
       
  2353                             (* defined element of the IDT, the corresponding value  *)
       
  2354                             (* of the parameter must be ignored                     *)
       
  2355                             if List.exists (equal True) xs then [pi] else []
       
  2356                         | ci_pi (Node xs, Node ys) = maps ci_pi (xs ~~ ys)
       
  2357                         | ci_pi (Node _, Leaf _) =
       
  2358                             raise REFUTE ("IDT_recursion_interpreter",
       
  2359                               "constructor takes more arguments than the " ^
       
  2360                                 "associated parameter")
       
  2361                       (* (int * interpretation list) list *)
       
  2362                       val rec_operators = map (fn (idx, c_p_intrs) =>
       
  2363                         (idx, maps ci_pi c_p_intrs)) mc_p_intrs
       
  2364                       (* sanity check: every recursion operator must provide as  *)
       
  2365                       (*               many values as the corresponding datatype *)
       
  2366                       (*               has elements                              *)
       
  2367                       val _ = map (fn (idx, intrs) =>
       
  2368                         let
       
  2369                           val T = typ_of_dtyp descr typ_assoc
       
  2370                             (Datatype.DtRec idx)
       
  2371                         in
       
  2372                           if length intrs <> size_of_type ctxt (typs, []) T then
       
  2373                             raise REFUTE ("IDT_recursion_interpreter",
       
  2374                               "wrong number of interpretations for rec. operator")
       
  2375                           else ()
       
  2376                         end) rec_operators
       
  2377                       (* For non-recursive datatypes, we are pretty much done at *)
       
  2378                       (* this point.  For recursive datatypes however, we still  *)
       
  2379                       (* need to apply the interpretations in 'rec_operators' to *)
       
  2380                       (* (recursively obtained) interpretations for recursive    *)
       
  2381                       (* constructor arguments.  To do so more efficiently, we   *)
       
  2382                       (* copy 'rec_operators' into arrays first.  Each Boolean   *)
       
  2383                       (* indicates whether the recursive arguments have been     *)
       
  2384                       (* considered already.                                     *)
       
  2385                       (* (int * (bool * interpretation) Array.array) list *)
       
  2386                       val REC_OPERATORS = map (fn (idx, intrs) =>
       
  2387                         (idx, Array.fromList (map (pair false) intrs)))
       
  2388                         rec_operators
       
  2389                       (* takes an interpretation, and if some leaf of this     *)
       
  2390                       (* interpretation is the 'elem'-th element of the type,  *)
       
  2391                       (* the indices of the arguments leading to this leaf are *)
       
  2392                       (* returned                                              *)
       
  2393                       (* interpretation -> int -> int list option *)
       
  2394                       fun get_args (Leaf xs) elem =
       
  2395                             if find_index (fn x => x = True) xs = elem then
       
  2396                               SOME []
       
  2397                             else
       
  2398                               NONE
       
  2399                         | get_args (Node xs) elem =
       
  2400                             let
       
  2401                               (* interpretation list * int -> int list option *)
       
  2402                               fun search ([], _) =
       
  2403                                 NONE
       
  2404                                 | search (x::xs, n) =
       
  2405                                 (case get_args x elem of
       
  2406                                   SOME result => SOME (n::result)
       
  2407                                 | NONE        => search (xs, n+1))
       
  2408                             in
       
  2409                               search (xs, 0)
       
  2410                             end
       
  2411                       (* returns the index of the constructor and indices for *)
       
  2412                       (* its arguments that generate the 'elem'-th element of *)
       
  2413                       (* the datatype given by 'idx'                          *)
       
  2414                       (* int -> int -> int * int list *)
       
  2415                       fun get_cargs idx elem =
       
  2416                         let
       
  2417                           (* int * interpretation list -> int * int list *)
       
  2418                           fun get_cargs_rec (_, []) =
       
  2419                                 raise REFUTE ("IDT_recursion_interpreter",
       
  2420                                   "no matching constructor found for datatype element")
       
  2421                             | get_cargs_rec (n, x::xs) =
       
  2422                                 (case get_args x elem of
       
  2423                                   SOME args => (n, args)
       
  2424                                 | NONE => get_cargs_rec (n+1, xs))
       
  2425                         in
       
  2426                           get_cargs_rec (0, the (AList.lookup (op =) mc_intrs idx))
       
  2427                         end
       
  2428                       (* computes one entry in 'REC_OPERATORS', and recursively *)
       
  2429                       (* all entries needed for it, where 'idx' gives the       *)
       
  2430                       (* datatype and 'elem' the element of it                  *)
       
  2431                       (* int -> int -> interpretation *)
       
  2432                       fun compute_array_entry idx elem =
       
  2433                         let
       
  2434                           val arr = the (AList.lookup (op =) REC_OPERATORS idx)
       
  2435                           val (flag, intr) = Array.sub (arr, elem)
       
  2436                         in
       
  2437                           if flag then
       
  2438                             (* simply return the previously computed result *)
       
  2439                             intr
       
  2440                           else
       
  2441                             (* we have to apply 'intr' to interpretations for all *)
       
  2442                             (* recursive arguments                                *)
       
  2443                             let
       
  2444                               (* int * int list *)
       
  2445                               val (c, args) = get_cargs idx elem
       
  2446                               (* find the indices of the constructor's /recursive/ *)
       
  2447                               (* arguments                                         *)
       
  2448                               val (_, _, constrs) = the (AList.lookup (op =) descr idx)
       
  2449                               val (_, dtyps) = nth constrs c
       
  2450                               val rec_dtyps_args = filter
       
  2451                                 (Datatype_Aux.is_rec_type o fst) (dtyps ~~ args)
       
  2452                               (* map those indices to interpretations *)
       
  2453                               val rec_dtyps_intrs = map (fn (dtyp, arg) =>
       
  2454                                 let
       
  2455                                   val dT = typ_of_dtyp descr typ_assoc dtyp
       
  2456                                   val consts = make_constants ctxt (typs, []) dT
       
  2457                                   val arg_i = nth consts arg
       
  2458                                 in
       
  2459                                   (dtyp, arg_i)
       
  2460                                 end) rec_dtyps_args
       
  2461                               (* takes the dtyp and interpretation of an element, *)
       
  2462                               (* and computes the interpretation for the          *)
       
  2463                               (* corresponding recursive argument                 *)
       
  2464                               fun rec_intr (Datatype.DtRec i) (Leaf xs) =
       
  2465                                     (* recursive argument is "rec_i params elem" *)
       
  2466                                     compute_array_entry i (find_index (fn x => x = True) xs)
       
  2467                                 | rec_intr (Datatype.DtRec _) (Node _) =
       
  2468                                     raise REFUTE ("IDT_recursion_interpreter",
       
  2469                                       "interpretation for IDT is a node")
       
  2470                                 | rec_intr (Datatype.DtType ("fun", [_, dt2])) (Node xs) =
       
  2471                                     (* recursive argument is something like     *)
       
  2472                                     (* "\<lambda>x::dt1. rec_? params (elem x)" *)
       
  2473                                     Node (map (rec_intr dt2) xs)
       
  2474                                 | rec_intr (Datatype.DtType ("fun", [_, _])) (Leaf _) =
       
  2475                                     raise REFUTE ("IDT_recursion_interpreter",
       
  2476                                       "interpretation for function dtyp is a leaf")
       
  2477                                 | rec_intr _ _ =
       
  2478                                     (* admissibility ensures that every recursive type *)
       
  2479                                     (* is of the form 'Dt_1 -> ... -> Dt_k ->          *)
       
  2480                                     (* (DtRec i)'                                      *)
       
  2481                                     raise REFUTE ("IDT_recursion_interpreter",
       
  2482                                       "non-recursive codomain in recursive dtyp")
       
  2483                               (* obtain interpretations for recursive arguments *)
       
  2484                               (* interpretation list *)
       
  2485                               val arg_intrs = map (uncurry rec_intr) rec_dtyps_intrs
       
  2486                               (* apply 'intr' to all recursive arguments *)
       
  2487                               val result = fold (fn arg_i => fn i =>
       
  2488                                 interpretation_apply (i, arg_i)) arg_intrs intr
       
  2489                               (* update 'REC_OPERATORS' *)
       
  2490                               val _ = Array.update (arr, elem, (true, result))
       
  2491                             in
       
  2492                               result
       
  2493                             end
       
  2494                         end
       
  2495                       val idt_size = Array.length (the (AList.lookup (op =) REC_OPERATORS idt_index))
       
  2496                       (* sanity check: the size of 'IDT' should be 'idt_size' *)
       
  2497                       val _ =
       
  2498                           if idt_size <> size_of_type ctxt (typs, []) IDT then
       
  2499                             raise REFUTE ("IDT_recursion_interpreter",
       
  2500                               "unexpected size of IDT (wrong type associated?)")
       
  2501                           else ()
       
  2502                       (* interpretation *)
       
  2503                       val rec_op = Node (map_range (compute_array_entry idt_index) idt_size)
       
  2504                     in
       
  2505                       SOME (rec_op, model', args')
       
  2506                     end
       
  2507                 end
       
  2508               else
       
  2509                 NONE  (* not a recursion operator of this datatype *)
       
  2510           ) (Datatype.get_all thy) NONE
       
  2511     | _ =>  (* head of term is not a constant *)
       
  2512       NONE
       
  2513   end;
       
  2514 
       
  2515 fun set_interpreter ctxt model args t =
       
  2516   let
       
  2517     val (typs, terms) = model
       
  2518   in
       
  2519     case AList.lookup (op =) terms t of
       
  2520       SOME intr =>
       
  2521         (* return an existing interpretation *)
       
  2522         SOME (intr, model, args)
       
  2523     | NONE =>
       
  2524         (case t of
       
  2525           Free (x, Type (@{type_name set}, [T])) =>
       
  2526           let
       
  2527             val (intr, _, args') =
       
  2528               interpret ctxt (typs, []) args (Free (x, T --> HOLogic.boolT))
       
  2529           in
       
  2530             SOME (intr, (typs, (t, intr)::terms), args')
       
  2531           end
       
  2532         | Var ((x, i), Type (@{type_name set}, [T])) =>
       
  2533           let
       
  2534             val (intr, _, args') =
       
  2535               interpret ctxt (typs, []) args (Var ((x,i), T --> HOLogic.boolT))
       
  2536           in
       
  2537             SOME (intr, (typs, (t, intr)::terms), args')
       
  2538           end
       
  2539         | Const (s, Type (@{type_name set}, [T])) =>
       
  2540           let
       
  2541             val (intr, _, args') =
       
  2542               interpret ctxt (typs, []) args (Const (s, T --> HOLogic.boolT))
       
  2543           in
       
  2544             SOME (intr, (typs, (t, intr)::terms), args')
       
  2545           end
       
  2546         (* 'Collect' == identity *)
       
  2547         | Const (@{const_name Collect}, _) $ t1 =>
       
  2548             SOME (interpret ctxt model args t1)
       
  2549         | Const (@{const_name Collect}, _) =>
       
  2550             SOME (interpret ctxt model args (eta_expand t 1))
       
  2551         (* 'op :' == application *)
       
  2552         | Const (@{const_name Set.member}, _) $ t1 $ t2 =>
       
  2553             SOME (interpret ctxt model args (t2 $ t1))
       
  2554         | Const (@{const_name Set.member}, _) $ _ =>
       
  2555             SOME (interpret ctxt model args (eta_expand t 1))
       
  2556         | Const (@{const_name Set.member}, _) =>
       
  2557             SOME (interpret ctxt model args (eta_expand t 2))
       
  2558         | _ => NONE)
       
  2559   end;
       
  2560 
       
  2561 (* only an optimization: 'card' could in principle be interpreted with *)
       
  2562 (* interpreters available already (using its definition), but the code *)
       
  2563 (* below is more efficient                                             *)
       
  2564 
       
  2565 fun Finite_Set_card_interpreter ctxt model args t =
       
  2566   case t of
       
  2567     Const (@{const_name Finite_Set.card},
       
  2568         Type ("fun", [Type (@{type_name set}, [T]), @{typ nat}])) =>
       
  2569       let
       
  2570         (* interpretation -> int *)
       
  2571         fun number_of_elements (Node xs) =
       
  2572             fold (fn x => fn n =>
       
  2573               if x = TT then
       
  2574                 n + 1
       
  2575               else if x = FF then
       
  2576                 n
       
  2577               else
       
  2578                 raise REFUTE ("Finite_Set_card_interpreter",
       
  2579                   "interpretation for set type does not yield a Boolean"))
       
  2580               xs 0
       
  2581           | number_of_elements (Leaf _) =
       
  2582               raise REFUTE ("Finite_Set_card_interpreter",
       
  2583                 "interpretation for set type is a leaf")
       
  2584         val size_of_nat = size_of_type ctxt model (@{typ nat})
       
  2585         (* takes an interpretation for a set and returns an interpretation *)
       
  2586         (* for a 'nat' denoting the set's cardinality                      *)
       
  2587         (* interpretation -> interpretation *)
       
  2588         fun card i =
       
  2589           let
       
  2590             val n = number_of_elements i
       
  2591           in
       
  2592             if n < size_of_nat then
       
  2593               Leaf ((replicate n False) @ True ::
       
  2594                 (replicate (size_of_nat-n-1) False))
       
  2595             else
       
  2596               Leaf (replicate size_of_nat False)
       
  2597           end
       
  2598         val set_constants = make_constants ctxt model (HOLogic.mk_setT T)
       
  2599       in
       
  2600         SOME (Node (map card set_constants), model, args)
       
  2601       end
       
  2602   | _ => NONE;
       
  2603 
       
  2604 (* only an optimization: 'finite' could in principle be interpreted with  *)
       
  2605 (* interpreters available already (using its definition), but the code    *)
       
  2606 (* below is more efficient                                                *)
       
  2607 
       
  2608 fun Finite_Set_finite_interpreter ctxt model args t =
       
  2609   case t of
       
  2610     Const (@{const_name Finite_Set.finite},
       
  2611            Type ("fun", [_, @{typ bool}])) $ _ =>
       
  2612         (* we only consider finite models anyway, hence EVERY set is *)
       
  2613         (* "finite"                                                  *)
       
  2614         SOME (TT, model, args)
       
  2615   | Const (@{const_name Finite_Set.finite},
       
  2616            Type ("fun", [set_T, @{typ bool}])) =>
       
  2617       let
       
  2618         val size_of_set = size_of_type ctxt model set_T
       
  2619       in
       
  2620         (* we only consider finite models anyway, hence EVERY set is *)
       
  2621         (* "finite"                                                  *)
       
  2622         SOME (Node (replicate size_of_set TT), model, args)
       
  2623       end
       
  2624   | _ => NONE;
       
  2625 
       
  2626 (* only an optimization: 'less' could in principle be interpreted with *)
       
  2627 (* interpreters available already (using its definition), but the code     *)
       
  2628 (* below is more efficient                                                 *)
       
  2629 
       
  2630 fun Nat_less_interpreter ctxt model args t =
       
  2631   case t of
       
  2632     Const (@{const_name Orderings.less}, Type ("fun", [@{typ nat},
       
  2633         Type ("fun", [@{typ nat}, @{typ bool}])])) =>
       
  2634       let
       
  2635         val size_of_nat = size_of_type ctxt model (@{typ nat})
       
  2636         (* the 'n'-th nat is not less than the first 'n' nats, while it *)
       
  2637         (* is less than the remaining 'size_of_nat - n' nats            *)
       
  2638         (* int -> interpretation *)
       
  2639         fun less n = Node ((replicate n FF) @ (replicate (size_of_nat - n) TT))
       
  2640       in
       
  2641         SOME (Node (map less (1 upto size_of_nat)), model, args)
       
  2642       end
       
  2643   | _ => NONE;
       
  2644 
       
  2645 (* only an optimization: 'plus' could in principle be interpreted with *)
       
  2646 (* interpreters available already (using its definition), but the code     *)
       
  2647 (* below is more efficient                                                 *)
       
  2648 
       
  2649 fun Nat_plus_interpreter ctxt model args t =
       
  2650   case t of
       
  2651     Const (@{const_name Groups.plus}, Type ("fun", [@{typ nat},
       
  2652         Type ("fun", [@{typ nat}, @{typ nat}])])) =>
       
  2653       let
       
  2654         val size_of_nat = size_of_type ctxt model (@{typ nat})
       
  2655         (* int -> int -> interpretation *)
       
  2656         fun plus m n =
       
  2657           let
       
  2658             val element = m + n
       
  2659           in
       
  2660             if element > size_of_nat - 1 then
       
  2661               Leaf (replicate size_of_nat False)
       
  2662             else
       
  2663               Leaf ((replicate element False) @ True ::
       
  2664                 (replicate (size_of_nat - element - 1) False))
       
  2665           end
       
  2666       in
       
  2667         SOME (Node (map_range (fn m => Node (map_range (plus m) size_of_nat)) size_of_nat),
       
  2668           model, args)
       
  2669       end
       
  2670   | _ => NONE;
       
  2671 
       
  2672 (* only an optimization: 'minus' could in principle be interpreted *)
       
  2673 (* with interpreters available already (using its definition), but the *)
       
  2674 (* code below is more efficient                                        *)
       
  2675 
       
  2676 fun Nat_minus_interpreter ctxt model args t =
       
  2677   case t of
       
  2678     Const (@{const_name Groups.minus}, Type ("fun", [@{typ nat},
       
  2679         Type ("fun", [@{typ nat}, @{typ nat}])])) =>
       
  2680       let
       
  2681         val size_of_nat = size_of_type ctxt model (@{typ nat})
       
  2682         (* int -> int -> interpretation *)
       
  2683         fun minus m n =
       
  2684           let
       
  2685             val element = Int.max (m-n, 0)
       
  2686           in
       
  2687             Leaf ((replicate element False) @ True ::
       
  2688               (replicate (size_of_nat - element - 1) False))
       
  2689           end
       
  2690       in
       
  2691         SOME (Node (map_range (fn m => Node (map_range (minus m) size_of_nat)) size_of_nat),
       
  2692           model, args)
       
  2693       end
       
  2694   | _ => NONE;
       
  2695 
       
  2696 (* only an optimization: 'times' could in principle be interpreted *)
       
  2697 (* with interpreters available already (using its definition), but the *)
       
  2698 (* code below is more efficient                                        *)
       
  2699 
       
  2700 fun Nat_times_interpreter ctxt model args t =
       
  2701   case t of
       
  2702     Const (@{const_name Groups.times}, Type ("fun", [@{typ nat},
       
  2703         Type ("fun", [@{typ nat}, @{typ nat}])])) =>
       
  2704       let
       
  2705         val size_of_nat = size_of_type ctxt model (@{typ nat})
       
  2706         (* nat -> nat -> interpretation *)
       
  2707         fun mult m n =
       
  2708           let
       
  2709             val element = m * n
       
  2710           in
       
  2711             if element > size_of_nat - 1 then
       
  2712               Leaf (replicate size_of_nat False)
       
  2713             else
       
  2714               Leaf ((replicate element False) @ True ::
       
  2715                 (replicate (size_of_nat - element - 1) False))
       
  2716           end
       
  2717       in
       
  2718         SOME (Node (map_range (fn m => Node (map_range (mult m) size_of_nat)) size_of_nat),
       
  2719           model, args)
       
  2720       end
       
  2721   | _ => NONE;
       
  2722 
       
  2723 (* only an optimization: 'append' could in principle be interpreted with *)
       
  2724 (* interpreters available already (using its definition), but the code   *)
       
  2725 (* below is more efficient                                               *)
       
  2726 
       
  2727 fun List_append_interpreter ctxt model args t =
       
  2728   case t of
       
  2729     Const (@{const_name List.append}, Type ("fun", [Type ("List.list", [T]), Type ("fun",
       
  2730         [Type ("List.list", [_]), Type ("List.list", [_])])])) =>
       
  2731       let
       
  2732         val size_elem = size_of_type ctxt model T
       
  2733         val size_list = size_of_type ctxt model (Type ("List.list", [T]))
       
  2734         (* maximal length of lists; 0 if we only consider the empty list *)
       
  2735         val list_length =
       
  2736           let
       
  2737             (* int -> int -> int -> int *)
       
  2738             fun list_length_acc len lists total =
       
  2739               if lists = total then
       
  2740                 len
       
  2741               else if lists < total then
       
  2742                 list_length_acc (len+1) (lists*size_elem) (total-lists)
       
  2743               else
       
  2744                 raise REFUTE ("List_append_interpreter",
       
  2745                   "size_list not equal to 1 + size_elem + ... + " ^
       
  2746                     "size_elem^len, for some len")
       
  2747           in
       
  2748             list_length_acc 0 1 size_list
       
  2749           end
       
  2750         val elements = 0 upto (size_list-1)
       
  2751         (* FIXME: there should be a nice formula, which computes the same as *)
       
  2752         (*        the following, but without all this intermediate tree      *)
       
  2753         (*        length/offset stuff                                        *)
       
  2754         (* associate each list with its length and offset in a complete tree *)
       
  2755         (* of width 'size_elem' and depth 'length_list' (with 'size_list'    *)
       
  2756         (* nodes total)                                                      *)
       
  2757         (* (int * (int * int)) list *)
       
  2758         val (lenoff_lists, _) = fold_map (fn elem => fn (offsets, off) =>
       
  2759           (* corresponds to a pre-order traversal of the tree *)
       
  2760           let
       
  2761             val len = length offsets
       
  2762             (* associate the given element with len/off *)
       
  2763             val assoc = (elem, (len, off))
       
  2764           in
       
  2765             if len < list_length then
       
  2766               (* go to first child node *)
       
  2767               (assoc, (off :: offsets, off * size_elem))
       
  2768             else if off mod size_elem < size_elem - 1 then
       
  2769               (* go to next sibling node *)
       
  2770               (assoc, (offsets, off + 1))
       
  2771             else
       
  2772               (* go back up the stack until we find a level where we can go *)
       
  2773               (* to the next sibling node                                   *)
       
  2774               let
       
  2775                 val offsets' = snd (take_prefix
       
  2776                   (fn off' => off' mod size_elem = size_elem - 1) offsets)
       
  2777               in
       
  2778                 case offsets' of
       
  2779                   [] =>
       
  2780                     (* we're at the last node in the tree; the next value *)
       
  2781                     (* won't be used anyway                               *)
       
  2782                     (assoc, ([], 0))
       
  2783                 | off'::offs' =>
       
  2784                     (* go to next sibling node *)
       
  2785                     (assoc, (offs', off' + 1))
       
  2786               end
       
  2787           end) elements ([], 0)
       
  2788         (* we also need the reverse association (from length/offset to *)
       
  2789         (* index)                                                      *)
       
  2790         val lenoff'_lists = map Library.swap lenoff_lists
       
  2791         (* returns the interpretation for "(list no. m) @ (list no. n)" *)
       
  2792         (* nat -> nat -> interpretation *)
       
  2793         fun append m n =
       
  2794           let
       
  2795             val (len_m, off_m) = the (AList.lookup (op =) lenoff_lists m)
       
  2796             val (len_n, off_n) = the (AList.lookup (op =) lenoff_lists n)
       
  2797             val len_elem = len_m + len_n
       
  2798             val off_elem = off_m * Integer.pow len_n size_elem + off_n
       
  2799           in
       
  2800             case AList.lookup op= lenoff'_lists (len_elem, off_elem) of
       
  2801               NONE =>
       
  2802                 (* undefined *)
       
  2803                 Leaf (replicate size_list False)
       
  2804             | SOME element =>
       
  2805                 Leaf ((replicate element False) @ True ::
       
  2806                   (replicate (size_list - element - 1) False))
       
  2807           end
       
  2808       in
       
  2809         SOME (Node (map (fn m => Node (map (append m) elements)) elements),
       
  2810           model, args)
       
  2811       end
       
  2812   | _ => NONE;
       
  2813 
       
  2814 (* only an optimization: 'lfp' could in principle be interpreted with  *)
       
  2815 (* interpreters available already (using its definition), but the code *)
       
  2816 (* below is more efficient                                             *)
       
  2817 
       
  2818 fun lfp_interpreter ctxt model args t =
       
  2819   case t of
       
  2820     Const (@{const_name lfp}, Type ("fun", [Type ("fun",
       
  2821       [Type (@{type_name set}, [T]),
       
  2822        Type (@{type_name set}, [_])]),
       
  2823        Type (@{type_name set}, [_])])) =>
       
  2824       let
       
  2825         val size_elem = size_of_type ctxt model T
       
  2826         (* the universe (i.e. the set that contains every element) *)
       
  2827         val i_univ = Node (replicate size_elem TT)
       
  2828         (* all sets with elements from type 'T' *)
       
  2829         val i_sets = make_constants ctxt model (HOLogic.mk_setT T)
       
  2830         (* all functions that map sets to sets *)
       
  2831         val i_funs = make_constants ctxt model (Type ("fun",
       
  2832           [HOLogic.mk_setT T, HOLogic.mk_setT T]))
       
  2833         (* "lfp(f) == Inter({u. f(u) <= u})" *)
       
  2834         (* interpretation * interpretation -> bool *)
       
  2835         fun is_subset (Node subs, Node sups) =
       
  2836               forall (fn (sub, sup) => (sub = FF) orelse (sup = TT)) (subs ~~ sups)
       
  2837           | is_subset (_, _) =
       
  2838               raise REFUTE ("lfp_interpreter",
       
  2839                 "is_subset: interpretation for set is not a node")
       
  2840         (* interpretation * interpretation -> interpretation *)
       
  2841         fun intersection (Node xs, Node ys) =
       
  2842               Node (map (fn (x, y) => if x=TT andalso y=TT then TT else FF)
       
  2843                 (xs ~~ ys))
       
  2844           | intersection (_, _) =
       
  2845               raise REFUTE ("lfp_interpreter",
       
  2846                 "intersection: interpretation for set is not a node")
       
  2847         (* interpretation -> interpretaion *)
       
  2848         fun lfp (Node resultsets) =
       
  2849               fold (fn (set, resultset) => fn acc =>
       
  2850                 if is_subset (resultset, set) then
       
  2851                   intersection (acc, set)
       
  2852                 else
       
  2853                   acc) (i_sets ~~ resultsets) i_univ
       
  2854           | lfp _ =
       
  2855               raise REFUTE ("lfp_interpreter",
       
  2856                 "lfp: interpretation for function is not a node")
       
  2857       in
       
  2858         SOME (Node (map lfp i_funs), model, args)
       
  2859       end
       
  2860   | _ => NONE;
       
  2861 
       
  2862 (* only an optimization: 'gfp' could in principle be interpreted with  *)
       
  2863 (* interpreters available already (using its definition), but the code *)
       
  2864 (* below is more efficient                                             *)
       
  2865 
       
  2866 fun gfp_interpreter ctxt model args t =
       
  2867   case t of
       
  2868     Const (@{const_name gfp}, Type ("fun", [Type ("fun",
       
  2869       [Type (@{type_name set}, [T]),
       
  2870        Type (@{type_name set}, [_])]),
       
  2871        Type (@{type_name set}, [_])])) =>
       
  2872       let
       
  2873         val size_elem = size_of_type ctxt model T
       
  2874         (* the universe (i.e. the set that contains every element) *)
       
  2875         val i_univ = Node (replicate size_elem TT)
       
  2876         (* all sets with elements from type 'T' *)
       
  2877         val i_sets = make_constants ctxt model (HOLogic.mk_setT T)
       
  2878         (* all functions that map sets to sets *)
       
  2879         val i_funs = make_constants ctxt model (Type ("fun",
       
  2880           [HOLogic.mk_setT T, HOLogic.mk_setT T]))
       
  2881         (* "gfp(f) == Union({u. u <= f(u)})" *)
       
  2882         (* interpretation * interpretation -> bool *)
       
  2883         fun is_subset (Node subs, Node sups) =
       
  2884               forall (fn (sub, sup) => (sub = FF) orelse (sup = TT))
       
  2885                 (subs ~~ sups)
       
  2886           | is_subset (_, _) =
       
  2887               raise REFUTE ("gfp_interpreter",
       
  2888                 "is_subset: interpretation for set is not a node")
       
  2889         (* interpretation * interpretation -> interpretation *)
       
  2890         fun union (Node xs, Node ys) =
       
  2891               Node (map (fn (x,y) => if x=TT orelse y=TT then TT else FF)
       
  2892                    (xs ~~ ys))
       
  2893           | union (_, _) =
       
  2894               raise REFUTE ("gfp_interpreter",
       
  2895                 "union: interpretation for set is not a node")
       
  2896         (* interpretation -> interpretaion *)
       
  2897         fun gfp (Node resultsets) =
       
  2898               fold (fn (set, resultset) => fn acc =>
       
  2899                 if is_subset (set, resultset) then
       
  2900                   union (acc, set)
       
  2901                 else
       
  2902                   acc) (i_sets ~~ resultsets) i_univ
       
  2903           | gfp _ =
       
  2904               raise REFUTE ("gfp_interpreter",
       
  2905                 "gfp: interpretation for function is not a node")
       
  2906       in
       
  2907         SOME (Node (map gfp i_funs), model, args)
       
  2908       end
       
  2909   | _ => NONE;
       
  2910 
       
  2911 (* only an optimization: 'fst' could in principle be interpreted with  *)
       
  2912 (* interpreters available already (using its definition), but the code *)
       
  2913 (* below is more efficient                                             *)
       
  2914 
       
  2915 fun Product_Type_fst_interpreter ctxt model args t =
       
  2916   case t of
       
  2917     Const (@{const_name fst}, Type ("fun", [Type (@{type_name Product_Type.prod}, [T, U]), _])) =>
       
  2918       let
       
  2919         val constants_T = make_constants ctxt model T
       
  2920         val size_U = size_of_type ctxt model U
       
  2921       in
       
  2922         SOME (Node (maps (replicate size_U) constants_T), model, args)
       
  2923       end
       
  2924   | _ => NONE;
       
  2925 
       
  2926 (* only an optimization: 'snd' could in principle be interpreted with  *)
       
  2927 (* interpreters available already (using its definition), but the code *)
       
  2928 (* below is more efficient                                             *)
       
  2929 
       
  2930 fun Product_Type_snd_interpreter ctxt model args t =
       
  2931   case t of
       
  2932     Const (@{const_name snd}, Type ("fun", [Type (@{type_name Product_Type.prod}, [T, U]), _])) =>
       
  2933       let
       
  2934         val size_T = size_of_type ctxt model T
       
  2935         val constants_U = make_constants ctxt model U
       
  2936       in
       
  2937         SOME (Node (flat (replicate size_T constants_U)), model, args)
       
  2938       end
       
  2939   | _ => NONE;
       
  2940 
       
  2941 
       
  2942 (* ------------------------------------------------------------------------- *)
       
  2943 (* PRINTERS                                                                  *)
       
  2944 (* ------------------------------------------------------------------------- *)
       
  2945 
       
  2946 fun stlc_printer ctxt model T intr assignment =
       
  2947   let
       
  2948     (* string -> string *)
       
  2949     val strip_leading_quote = perhaps (try (unprefix "'"))
       
  2950     (* Term.typ -> string *)
       
  2951     fun string_of_typ (Type (s, _)) = s
       
  2952       | string_of_typ (TFree (x, _)) = strip_leading_quote x
       
  2953       | string_of_typ (TVar ((x,i), _)) =
       
  2954           strip_leading_quote x ^ string_of_int i
       
  2955     (* interpretation -> int *)
       
  2956     fun index_from_interpretation (Leaf xs) =
       
  2957           find_index (Prop_Logic.eval assignment) xs
       
  2958       | index_from_interpretation _ =
       
  2959           raise REFUTE ("stlc_printer",
       
  2960             "interpretation for ground type is not a leaf")
       
  2961   in
       
  2962     case T of
       
  2963       Type ("fun", [T1, T2]) =>
       
  2964         let
       
  2965           (* create all constants of type 'T1' *)
       
  2966           val constants = make_constants ctxt model T1
       
  2967           (* interpretation list *)
       
  2968           val results =
       
  2969             (case intr of
       
  2970               Node xs => xs
       
  2971             | _ => raise REFUTE ("stlc_printer",
       
  2972               "interpretation for function type is a leaf"))
       
  2973           (* Term.term list *)
       
  2974           val pairs = map (fn (arg, result) =>
       
  2975             HOLogic.mk_prod
       
  2976               (print ctxt model T1 arg assignment,
       
  2977                print ctxt model T2 result assignment))
       
  2978             (constants ~~ results)
       
  2979           (* Term.typ *)
       
  2980           val HOLogic_prodT = HOLogic.mk_prodT (T1, T2)
       
  2981           val HOLogic_setT  = HOLogic.mk_setT HOLogic_prodT
       
  2982           (* Term.term *)
       
  2983           val HOLogic_empty_set = Const (@{const_abbrev Set.empty}, HOLogic_setT)
       
  2984           val HOLogic_insert    =
       
  2985             Const (@{const_name insert}, HOLogic_prodT --> HOLogic_setT --> HOLogic_setT)
       
  2986         in
       
  2987           SOME (fold_rev (fn pair => fn acc => HOLogic_insert $ pair $ acc) pairs HOLogic_empty_set)
       
  2988         end
       
  2989     | Type ("prop", []) =>
       
  2990         (case index_from_interpretation intr of
       
  2991           ~1 => SOME (HOLogic.mk_Trueprop (Const (@{const_name undefined}, HOLogic.boolT)))
       
  2992         | 0  => SOME (HOLogic.mk_Trueprop @{term True})
       
  2993         | 1  => SOME (HOLogic.mk_Trueprop @{term False})
       
  2994         | _  => raise REFUTE ("stlc_interpreter",
       
  2995           "illegal interpretation for a propositional value"))
       
  2996     | Type _  =>
       
  2997         if index_from_interpretation intr = (~1) then
       
  2998           SOME (Const (@{const_name undefined}, T))
       
  2999         else
       
  3000           SOME (Const (string_of_typ T ^
       
  3001             string_of_int (index_from_interpretation intr), T))
       
  3002     | TFree _ =>
       
  3003         if index_from_interpretation intr = (~1) then
       
  3004           SOME (Const (@{const_name undefined}, T))
       
  3005         else
       
  3006           SOME (Const (string_of_typ T ^
       
  3007             string_of_int (index_from_interpretation intr), T))
       
  3008     | TVar _  =>
       
  3009         if index_from_interpretation intr = (~1) then
       
  3010           SOME (Const (@{const_name undefined}, T))
       
  3011         else
       
  3012           SOME (Const (string_of_typ T ^
       
  3013             string_of_int (index_from_interpretation intr), T))
       
  3014   end;
       
  3015 
       
  3016 fun set_printer ctxt model T intr assignment =
       
  3017   (case T of
       
  3018     Type (@{type_name set}, [T1]) =>
       
  3019     let
       
  3020       (* create all constants of type 'T1' *)
       
  3021       val constants = make_constants ctxt model T1
       
  3022       (* interpretation list *)
       
  3023       val results = (case intr of
       
  3024           Node xs => xs
       
  3025         | _       => raise REFUTE ("set_printer",
       
  3026           "interpretation for set type is a leaf"))
       
  3027       (* Term.term list *)
       
  3028       val elements = List.mapPartial (fn (arg, result) =>
       
  3029         case result of
       
  3030           Leaf [fmTrue, (* fmFalse *) _] =>
       
  3031           if Prop_Logic.eval assignment fmTrue then
       
  3032             SOME (print ctxt model T1 arg assignment)
       
  3033           else (* if Prop_Logic.eval assignment fmFalse then *)
       
  3034             NONE
       
  3035         | _ =>
       
  3036           raise REFUTE ("set_printer",
       
  3037             "illegal interpretation for a Boolean value"))
       
  3038         (constants ~~ results)
       
  3039       (* Term.typ *)
       
  3040       val HOLogic_setT1     = HOLogic.mk_setT T1
       
  3041       (* Term.term *)
       
  3042       val HOLogic_empty_set = Const (@{const_abbrev Set.empty}, HOLogic_setT1)
       
  3043       val HOLogic_insert    =
       
  3044         Const (@{const_name insert}, T1 --> HOLogic_setT1 --> HOLogic_setT1)
       
  3045     in
       
  3046       SOME (Library.foldl (fn (acc, elem) => HOLogic_insert $ elem $ acc)
       
  3047         (HOLogic_empty_set, elements))
       
  3048     end
       
  3049   | _ =>
       
  3050     NONE);
       
  3051 
       
  3052 fun IDT_printer ctxt model T intr assignment =
       
  3053   let
       
  3054     val thy = Proof_Context.theory_of ctxt
       
  3055   in
       
  3056     (case T of
       
  3057       Type (s, Ts) =>
       
  3058         (case Datatype.get_info thy s of
       
  3059           SOME info =>  (* inductive datatype *)
       
  3060             let
       
  3061               val (typs, _)           = model
       
  3062               val index               = #index info
       
  3063               val descr               = #descr info
       
  3064               val (_, dtyps, constrs) = the (AList.lookup (op =) descr index)
       
  3065               val typ_assoc           = dtyps ~~ Ts
       
  3066               (* sanity check: every element in 'dtyps' must be a 'DtTFree' *)
       
  3067               val _ =
       
  3068                 if Library.exists (fn d =>
       
  3069                   case d of Datatype.DtTFree _ => false | _ => true) dtyps
       
  3070                 then
       
  3071                   raise REFUTE ("IDT_printer", "datatype argument (for type " ^
       
  3072                     Syntax.string_of_typ ctxt (Type (s, Ts)) ^ ") is not a variable")
       
  3073                 else ()
       
  3074               (* the index of the element in the datatype *)
       
  3075               val element =
       
  3076                 (case intr of
       
  3077                   Leaf xs => find_index (Prop_Logic.eval assignment) xs
       
  3078                 | Node _  => raise REFUTE ("IDT_printer",
       
  3079                   "interpretation is not a leaf"))
       
  3080             in
       
  3081               if element < 0 then
       
  3082                 SOME (Const (@{const_name undefined}, Type (s, Ts)))
       
  3083               else
       
  3084                 let
       
  3085                   (* takes a datatype constructor, and if for some arguments this  *)
       
  3086                   (* constructor generates the datatype's element that is given by *)
       
  3087                   (* 'element', returns the constructor (as a term) as well as the *)
       
  3088                   (* indices of the arguments                                      *)
       
  3089                   fun get_constr_args (cname, cargs) =
       
  3090                     let
       
  3091                       val cTerm      = Const (cname,
       
  3092                         map (typ_of_dtyp descr typ_assoc) cargs ---> Type (s, Ts))
       
  3093                       val (iC, _, _) = interpret ctxt (typs, []) {maxvars=0,
       
  3094                         def_eq=false, next_idx=1, bounds=[], wellformed=True} cTerm
       
  3095                       (* interpretation -> int list option *)
       
  3096                       fun get_args (Leaf xs) =
       
  3097                             if find_index (fn x => x = True) xs = element then
       
  3098                               SOME []
       
  3099                             else
       
  3100                               NONE
       
  3101                         | get_args (Node xs) =
       
  3102                             let
       
  3103                               (* interpretation * int -> int list option *)
       
  3104                               fun search ([], _) =
       
  3105                                 NONE
       
  3106                                 | search (x::xs, n) =
       
  3107                                 (case get_args x of
       
  3108                                   SOME result => SOME (n::result)
       
  3109                                 | NONE        => search (xs, n+1))
       
  3110                             in
       
  3111                               search (xs, 0)
       
  3112                             end
       
  3113                     in
       
  3114                       Option.map (fn args => (cTerm, cargs, args)) (get_args iC)
       
  3115                     end
       
  3116                   val (cTerm, cargs, args) =
       
  3117                     (* we could speed things up by computing the correct          *)
       
  3118                     (* constructor directly (rather than testing all              *)
       
  3119                     (* constructors), based on the order in which constructors    *)
       
  3120                     (* generate elements of datatypes; the current implementation *)
       
  3121                     (* of 'IDT_printer' however is independent of the internals   *)
       
  3122                     (* of 'IDT_constructor_interpreter'                           *)
       
  3123                     (case get_first get_constr_args constrs of
       
  3124                       SOME x => x
       
  3125                     | NONE   => raise REFUTE ("IDT_printer",
       
  3126                       "no matching constructor found for element " ^
       
  3127                       string_of_int element))
       
  3128                   val argsTerms = map (fn (d, n) =>
       
  3129                     let
       
  3130                       val dT = typ_of_dtyp descr typ_assoc d
       
  3131                       (* we only need the n-th element of this list, so there   *)
       
  3132                       (* might be a more efficient implementation that does not *)
       
  3133                       (* generate all constants                                 *)
       
  3134                       val consts = make_constants ctxt (typs, []) dT
       
  3135                     in
       
  3136                       print ctxt (typs, []) dT (nth consts n) assignment
       
  3137                     end) (cargs ~~ args)
       
  3138                 in
       
  3139                   SOME (list_comb (cTerm, argsTerms))
       
  3140                 end
       
  3141             end
       
  3142         | NONE =>  (* not an inductive datatype *)
       
  3143             NONE)
       
  3144     | _ =>  (* a (free or schematic) type variable *)
       
  3145         NONE)
       
  3146   end;
       
  3147 
       
  3148 
       
  3149 (* ------------------------------------------------------------------------- *)
       
  3150 (* use 'setup Refute.setup' in an Isabelle theory to initialize the 'Refute' *)
       
  3151 (* structure                                                                 *)
       
  3152 (* ------------------------------------------------------------------------- *)
       
  3153 
       
  3154 (* ------------------------------------------------------------------------- *)
       
  3155 (* Note: the interpreters and printers are used in reverse order; however,   *)
       
  3156 (*       an interpreter that can handle non-atomic terms ends up being       *)
       
  3157 (*       applied before the 'stlc_interpreter' breaks the term apart into    *)
       
  3158 (*       subterms that are then passed to other interpreters!                *)
       
  3159 (* ------------------------------------------------------------------------- *)
       
  3160 
       
  3161 val setup =
       
  3162    add_interpreter "stlc"    stlc_interpreter #>
       
  3163    add_interpreter "Pure"    Pure_interpreter #>
       
  3164    add_interpreter "HOLogic" HOLogic_interpreter #>
       
  3165    add_interpreter "set"     set_interpreter #>
       
  3166    add_interpreter "IDT"             IDT_interpreter #>
       
  3167    add_interpreter "IDT_constructor" IDT_constructor_interpreter #>
       
  3168    add_interpreter "IDT_recursion"   IDT_recursion_interpreter #>
       
  3169    add_interpreter "Finite_Set.card"    Finite_Set_card_interpreter #>
       
  3170    add_interpreter "Finite_Set.finite"  Finite_Set_finite_interpreter #>
       
  3171    add_interpreter "Nat_Orderings.less" Nat_less_interpreter #>
       
  3172    add_interpreter "Nat_HOL.plus"       Nat_plus_interpreter #>
       
  3173    add_interpreter "Nat_HOL.minus"      Nat_minus_interpreter #>
       
  3174    add_interpreter "Nat_HOL.times"      Nat_times_interpreter #>
       
  3175    add_interpreter "List.append" List_append_interpreter #>
       
  3176 (* UNSOUND
       
  3177    add_interpreter "lfp" lfp_interpreter #>
       
  3178    add_interpreter "gfp" gfp_interpreter #>
       
  3179 *)
       
  3180    add_interpreter "Product_Type.fst" Product_Type_fst_interpreter #>
       
  3181    add_interpreter "Product_Type.snd" Product_Type_snd_interpreter #>
       
  3182    add_printer "stlc" stlc_printer #>
       
  3183    add_printer "set" set_printer #>
       
  3184    add_printer "IDT"  IDT_printer;
       
  3185 
       
  3186 
       
  3187 
       
  3188 (** outer syntax commands 'refute' and 'refute_params' **)
       
  3189 
       
  3190 (* argument parsing *)
       
  3191 
       
  3192 (*optional list of arguments of the form [name1=value1, name2=value2, ...]*)
       
  3193 
       
  3194 val scan_parm = Parse.name -- (Scan.optional (@{keyword "="} |-- Parse.name) "true")
       
  3195 val scan_parms = Scan.optional (@{keyword "["} |-- Parse.list scan_parm --| @{keyword "]"}) [];
       
  3196 
       
  3197 
       
  3198 (* 'refute' command *)
       
  3199 
       
  3200 val _ =
       
  3201   Outer_Syntax.improper_command @{command_spec "refute"}
       
  3202     "try to find a model that refutes a given subgoal"
       
  3203     (scan_parms -- Scan.optional Parse.nat 1 >>
       
  3204       (fn (parms, i) =>
       
  3205         Toplevel.keep (fn state =>
       
  3206           let
       
  3207             val ctxt = Toplevel.context_of state;
       
  3208             val {goal = st, ...} = Proof.raw_goal (Toplevel.proof_of state);
       
  3209           in refute_goal ctxt parms st i; () end)));
       
  3210 
       
  3211 
       
  3212 (* 'refute_params' command *)
       
  3213 
       
  3214 val _ =
       
  3215   Outer_Syntax.command @{command_spec "refute_params"}
       
  3216     "show/store default parameters for the 'refute' command"
       
  3217     (scan_parms >> (fn parms =>
       
  3218       Toplevel.theory (fn thy =>
       
  3219         let
       
  3220           val thy' = fold set_default_param parms thy;
       
  3221           val output =
       
  3222             (case get_default_params (Proof_Context.init_global thy') of
       
  3223               [] => "none"
       
  3224             | new_defaults => cat_lines (map (fn (x, y) => x ^ "=" ^ y) new_defaults));
       
  3225           val _ = writeln ("Default parameters for 'refute':\n" ^ output);
       
  3226         in thy' end)));
       
  3227 
       
  3228 end;
       
  3229