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