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