1 (* Title: HOL/Tools/refute.ML
6 Finite model generation for HOL formulas, using a SAT solver.
9 (* ------------------------------------------------------------------------- *)
10 (* Declares the 'REFUTE' signature as well as a structure 'Refute'. *)
11 (* Documentation is available in the Isabelle/Isar theory 'HOL/Refute.thy'. *)
12 (* ------------------------------------------------------------------------- *)
17 exception REFUTE of string * string
19 (* ------------------------------------------------------------------------- *)
20 (* Model/interpretation related code (translation HOL -> propositional logic *)
21 (* ------------------------------------------------------------------------- *)
28 exception MAXVARS_EXCEEDED
30 val add_interpreter : string -> (theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option) -> theory -> theory
31 val add_printer : string -> (theory -> model -> Term.term -> interpretation -> (int -> bool) -> Term.term option) -> theory -> theory
33 val interpret : theory -> model -> arguments -> Term.term -> (interpretation * model * arguments)
35 val print : theory -> model -> Term.term -> interpretation -> (int -> bool) -> Term.term
36 val print_model : theory -> model -> (int -> bool) -> string
38 (* ------------------------------------------------------------------------- *)
40 (* ------------------------------------------------------------------------- *)
42 val set_default_param : (string * string) -> theory -> theory
43 val get_default_param : theory -> string -> string option
44 val get_default_params : theory -> (string * string) list
45 val actual_params : theory -> (string * string) list -> params
47 val find_model : theory -> params -> Term.term -> bool -> unit
49 val satisfy_term : theory -> (string * string) list -> Term.term -> unit (* tries to find a model for a formula *)
50 val refute_term : theory -> (string * string) list -> Term.term -> unit (* tries to find a model that refutes a formula *)
51 val refute_subgoal : theory -> (string * string) list -> Thm.thm -> int -> unit
53 val setup : theory -> theory
56 structure Refute : REFUTE =
61 (* We use 'REFUTE' only for internal error conditions that should *)
62 (* never occur in the first place (i.e. errors caused by bugs in our *)
63 (* code). Otherwise (e.g. to indicate invalid input data) we use *)
65 exception REFUTE of string * string; (* ("in function", "cause") *)
67 (* should be raised by an interpreter when more variables would be *)
68 (* required than allowed by 'maxvars' *)
69 exception MAXVARS_EXCEEDED;
71 (* ------------------------------------------------------------------------- *)
73 (* ------------------------------------------------------------------------- *)
75 (* ------------------------------------------------------------------------- *)
76 (* tree: implements an arbitrarily (but finitely) branching tree as a list *)
77 (* of (lists of ...) elements *)
78 (* ------------------------------------------------------------------------- *)
82 | Node of ('a tree) list;
84 (* ('a -> 'b) -> 'a tree -> 'b tree *)
89 | Node xs => Node (map (tree_map f) xs);
91 (* ('a * 'b -> 'a) -> 'a * ('b tree) -> 'a *)
95 fun itl (e, Leaf x) = f(e,x)
96 | itl (e, Node xs) = Library.foldl (tree_foldl f) (e,xs)
101 (* 'a tree * 'b tree -> ('a * 'b) tree *)
103 fun tree_pair (t1, t2) =
108 | Node _ => raise REFUTE ("tree_pair", "trees are of different height (second tree is higher)"))
111 (* '~~' will raise an exception if the number of branches in *)
112 (* both trees is different at the current node *)
113 Node ys => Node (map tree_pair (xs ~~ ys))
114 | Leaf _ => raise REFUTE ("tree_pair", "trees are of different height (first tree is higher)"));
116 (* ------------------------------------------------------------------------- *)
117 (* params: parameters that control the translation into a propositional *)
118 (* formula/model generation *)
120 (* The following parameters are supported (and required (!), except for *)
123 (* Name Type Description *)
125 (* "sizes" (string * int) list *)
126 (* Size of ground types (e.g. 'a=2), or depth of IDTs. *)
127 (* "minsize" int If >0, minimal size of each ground type/IDT depth. *)
128 (* "maxsize" int If >0, maximal size of each ground type/IDT depth. *)
129 (* "maxvars" int If >0, use at most 'maxvars' Boolean variables *)
130 (* when transforming the term into a propositional *)
132 (* "maxtime" int If >0, terminate after at most 'maxtime' seconds. *)
133 (* "satsolver" string SAT solver to be used. *)
134 (* ------------------------------------------------------------------------- *)
138 sizes : (string * int) list,
146 (* ------------------------------------------------------------------------- *)
147 (* interpretation: a term's interpretation is given by a variable of type *)
148 (* 'interpretation' *)
149 (* ------------------------------------------------------------------------- *)
151 type interpretation =
152 prop_formula list tree;
154 (* ------------------------------------------------------------------------- *)
155 (* model: a model specifies the size of types and the interpretation of *)
157 (* ------------------------------------------------------------------------- *)
160 (Term.typ * int) list * (Term.term * interpretation) list;
162 (* ------------------------------------------------------------------------- *)
163 (* arguments: additional arguments required during interpretation of terms *)
164 (* ------------------------------------------------------------------------- *)
168 maxvars : int, (* just passed unchanged from 'params' *)
169 def_eq : bool, (* whether to use 'make_equality' or 'make_def_equality' *)
170 (* the following may change during the translation *)
172 bounds : interpretation list,
173 wellformed: prop_formula
177 structure RefuteDataArgs =
179 val name = "HOL/refute";
181 {interpreters: (string * (theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option)) list,
182 printers: (string * (theory -> model -> Term.term -> interpretation -> (int -> bool) -> Term.term option)) list,
183 parameters: string Symtab.table};
184 val empty = {interpreters = [], printers = [], parameters = Symtab.empty};
188 ({interpreters = in1, printers = pr1, parameters = pa1},
189 {interpreters = in2, printers = pr2, parameters = pa2}) =
190 {interpreters = rev (merge_alists (rev in1) (rev in2)),
191 printers = rev (merge_alists (rev pr1) (rev pr2)),
192 parameters = Symtab.merge (op=) (pa1, pa2)};
193 fun print sg {interpreters, printers, parameters} =
194 Pretty.writeln (Pretty.chunks
195 [Pretty.strs ("default parameters:" :: List.concat (map (fn (name, value) => [name, "=", value]) (Symtab.dest parameters))),
196 Pretty.strs ("interpreters:" :: map fst interpreters),
197 Pretty.strs ("printers:" :: map fst printers)]);
200 structure RefuteData = TheoryDataFun(RefuteDataArgs);
203 (* ------------------------------------------------------------------------- *)
204 (* interpret: interprets the term 't' using a suitable interpreter; returns *)
205 (* the interpretation and a (possibly extended) model that keeps *)
206 (* track of the interpretation of subterms *)
207 (* ------------------------------------------------------------------------- *)
209 (* theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) *)
211 fun interpret thy model args t =
212 (case get_first (fn (_, f) => f thy model args t) (#interpreters (RefuteData.get thy)) of
213 NONE => raise REFUTE ("interpret", "no interpreter for term " ^ quote (Sign.string_of_term (sign_of thy) t))
216 (* ------------------------------------------------------------------------- *)
217 (* print: converts the constant denoted by the term 't' into a term using a *)
218 (* suitable printer *)
219 (* ------------------------------------------------------------------------- *)
221 (* theory -> model -> Term.term -> interpretation -> (int -> bool) -> Term.term *)
223 fun print thy model t intr assignment =
224 (case get_first (fn (_, f) => f thy model t intr assignment) (#printers (RefuteData.get thy)) of
225 NONE => raise REFUTE ("print", "no printer for term " ^ quote (Sign.string_of_term (sign_of thy) t))
228 (* ------------------------------------------------------------------------- *)
229 (* print_model: turns the model into a string, using a fixed interpretation *)
230 (* (given by an assignment for Boolean variables) and suitable *)
232 (* ------------------------------------------------------------------------- *)
234 (* theory -> model -> (int -> bool) -> string *)
236 fun print_model thy model assignment =
238 val (typs, terms) = model
241 "empty universe (no type variables in term)\n"
243 "Size of types: " ^ commas (map (fn (T, i) => Sign.string_of_typ (sign_of thy) T ^ ": " ^ string_of_int i) typs) ^ "\n"
244 val show_consts_msg =
245 if not (!show_consts) andalso Library.exists (is_Const o fst) terms then
246 "set \"show_consts\" to show the interpretation of constants\n"
251 "empty interpretation (no free variables in term)\n"
253 space_implode "\n" (List.mapPartial (fn (t, intr) =>
254 (* print constants only if 'show_consts' is true *)
255 if (!show_consts) orelse not (is_Const t) then
256 SOME (Sign.string_of_term (sign_of thy) t ^ ": " ^ Sign.string_of_term (sign_of thy) (print thy model t intr assignment))
260 typs_msg ^ show_consts_msg ^ terms_msg
264 (* ------------------------------------------------------------------------- *)
265 (* PARAMETER MANAGEMENT *)
266 (* ------------------------------------------------------------------------- *)
268 (* string -> (theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option) -> theory -> theory *)
270 fun add_interpreter name f thy =
272 val {interpreters, printers, parameters} = RefuteData.get thy
274 case AList.lookup (op =) interpreters name of
275 NONE => RefuteData.put {interpreters = (name, f) :: interpreters, printers = printers, parameters = parameters} thy
276 | SOME _ => error ("Interpreter " ^ name ^ " already declared")
279 (* string -> (theory -> model -> Term.term -> interpretation -> (int -> bool) -> Term.term option) -> theory -> theory *)
281 fun add_printer name f thy =
283 val {interpreters, printers, parameters} = RefuteData.get thy
285 case AList.lookup (op =) printers name of
286 NONE => RefuteData.put {interpreters = interpreters, printers = (name, f) :: printers, parameters = parameters} thy
287 | SOME _ => error ("Printer " ^ name ^ " already declared")
290 (* ------------------------------------------------------------------------- *)
291 (* set_default_param: stores the '(name, value)' pair in RefuteData's *)
292 (* parameter table *)
293 (* ------------------------------------------------------------------------- *)
295 (* (string * string) -> theory -> theory *)
297 fun set_default_param (name, value) thy =
299 val {interpreters, printers, parameters} = RefuteData.get thy
301 case Symtab.lookup parameters name of
302 NONE => RefuteData.put
303 {interpreters = interpreters, printers = printers, parameters = Symtab.extend (parameters, [(name, value)])} thy
304 | SOME _ => RefuteData.put
305 {interpreters = interpreters, printers = printers, parameters = Symtab.update (name, value) parameters} thy
308 (* ------------------------------------------------------------------------- *)
309 (* get_default_param: retrieves the value associated with 'name' from *)
310 (* RefuteData's parameter table *)
311 (* ------------------------------------------------------------------------- *)
313 (* theory -> string -> string option *)
315 val get_default_param = Symtab.lookup o #parameters o RefuteData.get;
317 (* ------------------------------------------------------------------------- *)
318 (* get_default_params: returns a list of all '(name, value)' pairs that are *)
319 (* stored in RefuteData's parameter table *)
320 (* ------------------------------------------------------------------------- *)
322 (* theory -> (string * string) list *)
324 val get_default_params = Symtab.dest o #parameters o RefuteData.get;
326 (* ------------------------------------------------------------------------- *)
327 (* actual_params: takes a (possibly empty) list 'params' of parameters that *)
328 (* override the default parameters currently specified in 'thy', and *)
329 (* returns a record that can be passed to 'find_model'. *)
330 (* ------------------------------------------------------------------------- *)
332 (* theory -> (string * string) list -> params *)
334 fun actual_params thy override =
336 (* (string * string) list * string -> int *)
337 fun read_int (parms, name) =
338 case AList.lookup (op =) parms name of
339 SOME s => (case Int.fromString s of
341 | NONE => error ("parameter " ^ quote name ^ " (value is " ^ quote s ^ ") must be an integer value"))
342 | NONE => error ("parameter " ^ quote name ^ " must be assigned a value")
343 (* (string * string) list * string -> string *)
344 fun read_string (parms, name) =
345 case AList.lookup (op =) parms name of
347 | NONE => error ("parameter " ^ quote name ^ " must be assigned a value")
348 (* (string * string) list *)
349 val allparams = override @ (get_default_params thy) (* 'override' first, defaults last *)
351 val minsize = read_int (allparams, "minsize")
352 val maxsize = read_int (allparams, "maxsize")
353 val maxvars = read_int (allparams, "maxvars")
354 val maxtime = read_int (allparams, "maxtime")
356 val satsolver = read_string (allparams, "satsolver")
357 (* all remaining parameters of the form "string=int" are collected in *)
359 (* TODO: it is currently not possible to specify a size for a type *)
360 (* whose name is one of the other parameters (e.g. 'maxvars') *)
361 (* (string * int) list *)
362 val sizes = List.mapPartial
363 (fn (name,value) => (case Int.fromString value of SOME i => SOME (name, i) | NONE => NONE))
364 (List.filter (fn (name,_) => name<>"minsize" andalso name<>"maxsize" andalso name<>"maxvars" andalso name<>"maxtime" andalso name<>"satsolver")
367 {sizes=sizes, minsize=minsize, maxsize=maxsize, maxvars=maxvars, maxtime=maxtime, satsolver=satsolver}
371 (* ------------------------------------------------------------------------- *)
372 (* TRANSLATION HOL -> PROPOSITIONAL LOGIC, BOOLEAN ASSIGNMENT -> MODEL *)
373 (* ------------------------------------------------------------------------- *)
375 (* ------------------------------------------------------------------------- *)
376 (* typ_of_dtyp: converts a data type ('DatatypeAux.dtyp') into a type *)
377 (* ('Term.typ'), given type parameters for the data type's type *)
379 (* ------------------------------------------------------------------------- *)
381 (* DatatypeAux.descr -> (DatatypeAux.dtyp * Term.typ) list -> DatatypeAux.dtyp -> Term.typ *)
383 fun typ_of_dtyp descr typ_assoc (DatatypeAux.DtTFree a) =
384 (* replace a 'DtTFree' variable by the associated type *)
385 (the o AList.lookup (op =) typ_assoc) (DatatypeAux.DtTFree a)
386 | typ_of_dtyp descr typ_assoc (DatatypeAux.DtType (s, ds)) =
387 Type (s, map (typ_of_dtyp descr typ_assoc) ds)
388 | typ_of_dtyp descr typ_assoc (DatatypeAux.DtRec i) =
390 val (s, ds, _) = (the o AList.lookup (op =) descr) i
392 Type (s, map (typ_of_dtyp descr typ_assoc) ds)
395 (* ------------------------------------------------------------------------- *)
396 (* collect_axioms: collects (monomorphic, universally quantified versions *)
397 (* of) all HOL axioms that are relevant w.r.t 't' *)
398 (* ------------------------------------------------------------------------- *)
400 (* Note: to make the collection of axioms more easily extensible, this *)
401 (* function could be based on user-supplied "axiom collectors", *)
402 (* similar to 'interpret'/interpreters or 'print'/printers *)
404 (* theory -> Term.term -> Term.term list *)
406 (* Which axioms are "relevant" for a particular term/type goes hand in *)
407 (* hand with the interpretation of that term/type by its interpreter (see *)
408 (* way below): if the interpretation respects an axiom anyway, the axiom *)
409 (* does not need to be added as a constraint here. *)
411 (* When an axiom is added as relevant, further axioms may need to be *)
412 (* added as well (e.g. when a constant is defined in terms of other *)
413 (* constants). To avoid infinite recursion (which should not happen for *)
414 (* constants anyway, but it could happen for "typedef"-related axioms, *)
415 (* since they contain the type again), we use an accumulator 'axs' and *)
416 (* add a relevant axiom only if it is not in 'axs' yet. *)
418 fun collect_axioms thy t =
420 val _ = immediate_output "Adding axioms..."
421 (* (string * Term.term) list *)
422 val axioms = Theory.all_axioms_of thy;
424 val rec_names = Symtab.fold (append o #rec_names o snd)
425 (DatatypePackage.get_datatypes thy) [];
427 val const_of_class_names = map Logic.const_of_class (Sign.classes thy)
428 (* given a constant 's' of type 'T', which is a subterm of 't', where *)
429 (* 't' has a (possibly) more general type, the schematic type *)
430 (* variables in 't' are instantiated to match the type 'T' (may raise *)
431 (* Type.TYPE_MATCH) *)
432 (* (string * Term.typ) * Term.term -> Term.term *)
433 fun specialize_type ((s, T), t) =
435 fun find_typeSubs (Const (s', T')) =
437 SOME (Sign.typ_match thy (T', T) Vartab.empty) handle Type.TYPE_MATCH => NONE
440 | find_typeSubs (Free _) = NONE
441 | find_typeSubs (Var _) = NONE
442 | find_typeSubs (Bound _) = NONE
443 | find_typeSubs (Abs (_, _, body)) = find_typeSubs body
444 | find_typeSubs (t1 $ t2) = (case find_typeSubs t1 of SOME x => SOME x | NONE => find_typeSubs t2)
445 val typeSubs = (case find_typeSubs t of
447 | NONE => raise Type.TYPE_MATCH (* no match found - perhaps due to sort constraints *))
452 case Type.lookup (typeSubs, v) of
454 (* schematic type variable not instantiated *)
455 raise REFUTE ("collect_axioms", "term " ^ Sign.string_of_term (sign_of thy) t ^ " still has a polymorphic type (after instantiating type of " ^ quote s ^ ")")
460 (* applies a type substitution 'typeSubs' for all type variables in a *)
462 (* (Term.sort * Term.typ) Term.Vartab.table -> Term.term -> Term.term *)
463 fun monomorphic_term typeSubs t =
464 map_types (map_type_tvar
466 case Type.lookup (typeSubs, v) of
468 (* schematic type variable not instantiated *)
472 (* Term.term list * Term.typ -> Term.term list *)
473 fun collect_sort_axioms (axs, T) =
475 (* collect the axioms for a single 'class' (but not for its superclasses) *)
476 (* Term.term list * string -> Term.term list *)
477 fun collect_class_axioms (axs, class) =
479 (* obtain the axioms generated by the "axclass" command *)
480 (* (string * Term.term) list *)
481 val class_axioms = List.filter (fn (s, _) => String.isPrefix (Logic.const_of_class class ^ ".axioms_") s) axioms
482 (* replace the one schematic type variable in each axiom by the actual type 'T' *)
483 (* (string * Term.term) list *)
484 val monomorphic_class_axioms = map (fn (axname, ax) =>
486 val (idx, S) = (case term_tvars ax of
488 | _ => raise REFUTE ("collect_axioms", "class axiom " ^ axname ^ " (" ^ Sign.string_of_term (sign_of thy) ax ^ ") does not contain exactly one type variable"))
490 (axname, monomorphic_term (Vartab.make [(idx, (S, T))]) ax)
492 (* Term.term list * (string * Term.term) list -> Term.term list *)
493 fun collect_axiom (axs, (axname, ax)) =
494 if member (op aconv) axs ax then
497 immediate_output (" " ^ axname);
498 collect_term_axioms (ax :: axs, ax)
501 Library.foldl collect_axiom (axs, monomorphic_class_axioms)
504 val sort = (case T of
505 TFree (_, sort) => sort
506 | TVar (_, sort) => sort
507 | _ => raise REFUTE ("collect_axioms", "type " ^ Sign.string_of_typ (sign_of thy) T ^ " is not a variable"))
508 (* obtain all superclasses of classes in 'sort' *)
510 val superclasses = distinct (op =) (sort @ maps (Sign.super_classes thy) sort)
512 Library.foldl collect_class_axioms (axs, superclasses)
514 (* Term.term list * Term.typ -> Term.term list *)
515 and collect_type_axioms (axs, T) =
518 Type ("prop", []) => axs
519 | Type ("fun", [T1, T2]) => collect_type_axioms (collect_type_axioms (axs, T1), T2)
520 | Type ("set", [T1]) => collect_type_axioms (axs, T1)
521 | Type ("itself", [T1]) => collect_type_axioms (axs, T1) (* axiomatic type classes *)
524 (* look up the definition of a type, as created by "typedef" *)
525 (* (string * Term.term) list -> (string * Term.term) option *)
526 fun get_typedefn [] =
528 | get_typedefn ((axname,ax)::axms) =
530 (* Term.term -> Term.typ option *)
531 fun type_of_type_definition (Const (s', T')) =
532 if s'="Typedef.type_definition" then
536 | type_of_type_definition (Free _) = NONE
537 | type_of_type_definition (Var _) = NONE
538 | type_of_type_definition (Bound _) = NONE
539 | type_of_type_definition (Abs (_, _, body)) = type_of_type_definition body
540 | type_of_type_definition (t1 $ t2) = (case type_of_type_definition t1 of SOME x => SOME x | NONE => type_of_type_definition t2)
542 case type_of_type_definition ax of
545 val T'' = (domain_type o domain_type) T'
546 val typeSubs = Sign.typ_match thy (T'', T) Vartab.empty
548 SOME (axname, monomorphic_term typeSubs ax)
553 handle ERROR _ => get_typedefn axms
554 | MATCH => get_typedefn axms
555 | Type.TYPE_MATCH => get_typedefn axms)
557 case DatatypePackage.get_datatype thy s of
558 SOME info => (* inductive datatype *)
559 (* only collect relevant type axioms for the argument types *)
560 Library.foldl collect_type_axioms (axs, Ts)
562 (case get_typedefn axioms of
564 if member (op aconv) axs ax then
565 (* only collect relevant type axioms for the argument types *)
566 Library.foldl collect_type_axioms (axs, Ts)
568 (immediate_output (" " ^ axname);
569 collect_term_axioms (ax :: axs, ax))
571 (* unspecified type, perhaps introduced with 'typedecl' *)
572 (* at least collect relevant type axioms for the argument types *)
573 Library.foldl collect_type_axioms (axs, Ts))
575 | TFree _ => collect_sort_axioms (axs, T) (* axiomatic type classes *)
576 | TVar _ => collect_sort_axioms (axs, T) (* axiomatic type classes *)
577 (* Term.term list * Term.term -> Term.term list *)
578 and collect_term_axioms (axs, t) =
581 Const ("all", _) => axs
582 | Const ("==", _) => axs
583 | Const ("==>", _) => axs
584 | Const ("TYPE", T) => collect_type_axioms (axs, T) (* axiomatic type classes *)
586 | Const ("Trueprop", _) => axs
587 | Const ("Not", _) => axs
588 | Const ("True", _) => axs (* redundant, since 'True' is also an IDT constructor *)
589 | Const ("False", _) => axs (* redundant, since 'False' is also an IDT constructor *)
590 | Const ("arbitrary", T) => collect_type_axioms (axs, T)
591 | Const ("The", T) =>
593 val ax = specialize_type (("The", T), (the o AList.lookup (op =) axioms) "HOL.the_eq_trivial")
595 if member (op aconv) axs ax then
596 collect_type_axioms (axs, T)
598 (immediate_output " HOL.the_eq_trivial";
599 collect_term_axioms (ax :: axs, ax))
601 | Const ("Hilbert_Choice.Eps", T) =>
603 val ax = specialize_type (("Hilbert_Choice.Eps", T),
604 (the o AList.lookup (op =) axioms) "Hilbert_Choice.someI")
606 if member (op aconv) axs ax then
607 collect_type_axioms (axs, T)
609 (immediate_output " Hilbert_Choice.someI";
610 collect_term_axioms (ax :: axs, ax))
612 | Const ("All", _) $ t1 => collect_term_axioms (axs, t1)
613 | Const ("Ex", _) $ t1 => collect_term_axioms (axs, t1)
614 | Const ("op =", T) => collect_type_axioms (axs, T)
615 | Const ("op &", _) => axs
616 | Const ("op |", _) => axs
617 | Const ("op -->", _) => axs
619 | Const ("Collect", T) => collect_type_axioms (axs, T)
620 | Const ("op :", T) => collect_type_axioms (axs, T)
621 (* other optimizations *)
622 | Const ("Finite_Set.card", T) => collect_type_axioms (axs, T)
623 | Const ("Finite_Set.Finites", T) => collect_type_axioms (axs, T)
624 | Const ("Orderings.less", T as Type ("fun", [Type ("nat", []), Type ("fun", [Type ("nat", []), Type ("bool", [])])])) => collect_type_axioms (axs, T)
625 | Const ("HOL.plus", T as Type ("fun", [Type ("nat", []), Type ("fun", [Type ("nat", []), Type ("nat", [])])])) => collect_type_axioms (axs, T)
626 | Const ("HOL.minus", T as Type ("fun", [Type ("nat", []), Type ("fun", [Type ("nat", []), Type ("nat", [])])])) => collect_type_axioms (axs, T)
627 | Const ("HOL.times", T as Type ("fun", [Type ("nat", []), Type ("fun", [Type ("nat", []), Type ("nat", [])])])) => collect_type_axioms (axs, T)
628 | Const ("List.op @", T) => collect_type_axioms (axs, T)
629 | Const ("Lfp.lfp", T) => collect_type_axioms (axs, T)
630 | Const ("Gfp.gfp", T) => collect_type_axioms (axs, T)
631 (* simply-typed lambda calculus *)
634 (* look up the definition of a constant, as created by "constdefs" *)
635 (* string -> Term.typ -> (string * Term.term) list -> (string * Term.term) option *)
638 | get_defn ((axname, ax)::axms) =
640 val (lhs, _) = Logic.dest_equals ax (* equations only *)
642 val (s', T') = dest_Const c
646 val typeSubs = Sign.typ_match thy (T', T) Vartab.empty
648 SOME (axname, monomorphic_term typeSubs ax)
653 handle ERROR _ => get_defn axms
654 | TERM _ => get_defn axms
655 | Type.TYPE_MATCH => get_defn axms)
656 (* axiomatic type classes *)
658 fun is_const_of_class () =
659 (* I'm not quite sure if checking the name 's' is sufficient, *)
660 (* or if we should also check the type 'T' *)
661 s mem const_of_class_names
662 (* inductive data types *)
664 fun is_IDT_constructor () =
667 (case DatatypePackage.get_datatype_constrs thy s' of
669 Library.exists (fn (cname, cty) =>
670 cname = s andalso Sign.typ_instance thy (T, cty))
677 fun is_IDT_recursor () =
678 (* I'm not quite sure if checking the name 's' is sufficient, *)
679 (* or if we should also check the type 'T' *)
680 member (op =) rec_names s
682 if is_const_of_class () then
683 (* axiomatic type classes: add "OFCLASS(?'a::c, c_class)" and *)
684 (* the introduction rule "class.intro" as axioms *)
686 val class = Logic.class_of_const s
687 val inclass = Logic.mk_inclass (TVar (("'a", 0), [class]), class)
688 (* Term.term option *)
689 val ofclass_ax = (SOME (specialize_type ((s, T), inclass)) handle Type.TYPE_MATCH => NONE)
690 val intro_ax = (Option.map (fn t => specialize_type ((s, T), t))
691 (AList.lookup (op =) axioms (class ^ ".intro")) handle Type.TYPE_MATCH => NONE)
692 val axs' = (case ofclass_ax of NONE => axs | SOME ax => if member (op aconv) axs ax then
693 (* collect relevant type axioms *)
694 collect_type_axioms (axs, T)
696 (immediate_output (" " ^ Sign.string_of_term (sign_of thy) ax);
697 collect_term_axioms (ax :: axs, ax)))
698 val axs'' = (case intro_ax of NONE => axs' | SOME ax => if member (op aconv) axs' ax then
699 (* collect relevant type axioms *)
700 collect_type_axioms (axs', T)
702 (immediate_output (" " ^ s ^ ".intro");
703 collect_term_axioms (ax :: axs', ax)))
707 else if is_IDT_constructor () then
708 (* only collect relevant type axioms *)
709 collect_type_axioms (axs, T)
710 else if is_IDT_recursor () then
711 (* only collect relevant type axioms *)
712 collect_type_axioms (axs, T)
714 case get_defn axioms of
716 if member (op aconv) axs ax then
717 (* collect relevant type axioms *)
718 collect_type_axioms (axs, T)
720 (immediate_output (" " ^ axname);
721 collect_term_axioms (ax :: axs, ax))
723 (* collect relevant type axioms *)
724 collect_type_axioms (axs, T)
727 | Free (_, T) => collect_type_axioms (axs, T)
728 | Var (_, T) => collect_type_axioms (axs, T)
730 | Abs (_, T, body) => collect_term_axioms (collect_type_axioms (axs, T), body)
731 | t1 $ t2 => collect_term_axioms (collect_term_axioms (axs, t1), t2)
732 (* universal closure over schematic variables *)
733 (* Term.term -> Term.term *)
736 (* (Term.indexname * Term.typ) list *)
737 val vars = sort_wrt (fst o fst) (map dest_Var (term_vars t))
740 (fn (t', ((x, i), T)) => (Term.all T) $ Abs (x, T, abstract_over (Var((x, i), T), t')))
744 val result = map close_form (collect_term_axioms ([], t))
745 val _ = writeln " ...done."
750 (* ------------------------------------------------------------------------- *)
751 (* ground_types: collects all ground types in a term (including argument *)
752 (* types of other types), suppressing duplicates. Does not *)
753 (* return function types, set types, non-recursive IDTs, or *)
754 (* 'propT'. For IDTs, also the argument types of constructors *)
755 (* are considered. *)
756 (* ------------------------------------------------------------------------- *)
758 (* theory -> Term.term -> Term.typ list *)
760 fun ground_types thy t =
762 (* Term.typ * Term.typ list -> Term.typ list *)
763 fun collect_types (T, acc) =
765 acc (* prevent infinite recursion (for IDTs) *)
768 Type ("fun", [T1, T2]) => collect_types (T1, collect_types (T2, acc))
769 | Type ("prop", []) => acc
770 | Type ("set", [T1]) => collect_types (T1, acc)
772 (case DatatypePackage.get_datatype thy s of
773 SOME info => (* inductive datatype *)
775 val index = #index info
776 val descr = #descr info
777 val (_, dtyps, constrs) = (the o AList.lookup (op =) descr) index
778 val typ_assoc = dtyps ~~ Ts
779 (* sanity check: every element in 'dtyps' must be a 'DtTFree' *)
780 val _ = (if Library.exists (fn d =>
781 case d of DatatypeAux.DtTFree _ => false | _ => true) dtyps
783 raise REFUTE ("ground_types", "datatype argument (for type " ^ Sign.string_of_typ (sign_of thy) (Type (s, Ts)) ^ ") is not a variable")
786 (* if the current type is a recursive IDT (i.e. a depth is required), add it to 'acc' *)
787 val acc' = (if Library.exists (fn (_, ds) => Library.exists DatatypeAux.is_rec_type ds) constrs then
791 (* collect argument types *)
792 val acc_args = foldr collect_types acc' Ts
793 (* collect constructor types *)
794 val acc_constrs = foldr collect_types acc_args (List.concat (map (fn (_, ds) => map (typ_of_dtyp descr typ_assoc) ds) constrs))
798 | NONE => (* not an inductive datatype, e.g. defined via "typedef" or "typedecl" *)
799 insert (op =) T (foldr collect_types acc Ts))
800 | TFree _ => insert (op =) T acc
801 | TVar _ => insert (op =) T acc)
803 it_term_types collect_types (t, [])
806 (* ------------------------------------------------------------------------- *)
807 (* string_of_typ: (rather naive) conversion from types to strings, used to *)
808 (* look up the size of a type in 'sizes'. Parameterized *)
809 (* types with different parameters (e.g. "'a list" vs. "bool *)
810 (* list") are identified. *)
811 (* ------------------------------------------------------------------------- *)
813 (* Term.typ -> string *)
815 fun string_of_typ (Type (s, _)) = s
816 | string_of_typ (TFree (s, _)) = s
817 | string_of_typ (TVar ((s,_), _)) = s;
819 (* ------------------------------------------------------------------------- *)
820 (* first_universe: returns the "first" (i.e. smallest) universe by assigning *)
821 (* 'minsize' to every type for which no size is specified in *)
823 (* ------------------------------------------------------------------------- *)
825 (* Term.typ list -> (string * int) list -> int -> (Term.typ * int) list *)
827 fun first_universe xs sizes minsize =
830 case AList.lookup (op =) sizes (string_of_typ T) of
834 map (fn T => (T, size_of_typ T)) xs
837 (* ------------------------------------------------------------------------- *)
838 (* next_universe: enumerates all universes (i.e. assignments of sizes to *)
839 (* types), where the minimal size of a type is given by *)
840 (* 'minsize', the maximal size is given by 'maxsize', and a *)
841 (* type may have a fixed size given in 'sizes' *)
842 (* ------------------------------------------------------------------------- *)
844 (* (Term.typ * int) list -> (string * int) list -> int -> int -> (Term.typ * int) list option *)
846 fun next_universe xs sizes minsize maxsize =
848 (* creates the "first" list of length 'len', where the sum of all list *)
849 (* elements is 'sum', and the length of the list is 'len' *)
850 (* int -> int -> int -> int list option *)
851 fun make_first _ 0 sum =
856 | make_first max len sum =
857 if sum<=max orelse max<0 then
858 Option.map (fn xs' => sum :: xs') (make_first max (len-1) 0)
860 Option.map (fn xs' => max :: xs') (make_first max (len-1) (sum-max))
861 (* enumerates all int lists with a fixed length, where 0<=x<='max' for *)
862 (* all list elements x (unless 'max'<0) *)
863 (* int -> int -> int -> int list -> int list option *)
864 fun next max len sum [] =
866 | next max len sum [x] =
867 (* we've reached the last list element, so there's no shift possible *)
868 make_first max (len+1) (sum+x+1) (* increment 'sum' by 1 *)
869 | next max len sum (x1::x2::xs) =
870 if x1>0 andalso (x2<max orelse max<0) then
872 SOME (valOf (make_first max (len+1) (sum+x1-1)) @ (x2+1) :: xs)
874 (* continue search *)
875 next max (len+1) (sum+x1) (x2::xs)
876 (* only consider those types for which the size is not fixed *)
877 val mutables = List.filter (not o (AList.defined (op =) sizes) o string_of_typ o fst) xs
878 (* subtract 'minsize' from every size (will be added again at the end) *)
879 val diffs = map (fn (_, n) => n-minsize) mutables
881 case next (maxsize-minsize) 0 0 diffs of
883 (* merge with those types for which the size is fixed *)
884 SOME (snd (foldl_map (fn (ds, (T, _)) =>
885 case AList.lookup (op =) sizes (string_of_typ T) of
886 SOME n => (ds, (T, n)) (* return the fixed size *)
887 | NONE => (tl ds, (T, minsize + hd ds))) (* consume the head of 'ds', add 'minsize' *)
893 (* ------------------------------------------------------------------------- *)
894 (* toTrue: converts the interpretation of a Boolean value to a propositional *)
895 (* formula that is true iff the interpretation denotes "true" *)
896 (* ------------------------------------------------------------------------- *)
898 (* interpretation -> prop_formula *)
900 fun toTrue (Leaf [fm, _]) = fm
901 | toTrue _ = raise REFUTE ("toTrue", "interpretation does not denote a Boolean value");
903 (* ------------------------------------------------------------------------- *)
904 (* toFalse: converts the interpretation of a Boolean value to a *)
905 (* propositional formula that is true iff the interpretation *)
906 (* denotes "false" *)
907 (* ------------------------------------------------------------------------- *)
909 (* interpretation -> prop_formula *)
911 fun toFalse (Leaf [_, fm]) = fm
912 | toFalse _ = raise REFUTE ("toFalse", "interpretation does not denote a Boolean value");
914 (* ------------------------------------------------------------------------- *)
915 (* find_model: repeatedly calls 'interpret' with appropriate parameters, *)
916 (* applies a SAT solver, and (in case a model is found) displays *)
917 (* the model to the user by calling 'print_model' *)
918 (* thy : the current theory *)
919 (* {...} : parameters that control the translation/model generation *)
920 (* t : term to be translated into a propositional formula *)
921 (* negate : if true, find a model that makes 't' false (rather than true) *)
922 (* ------------------------------------------------------------------------- *)
924 (* theory -> params -> Term.term -> bool -> unit *)
926 fun find_model thy {sizes, minsize, maxsize, maxvars, maxtime, satsolver} t negate =
932 val axioms = collect_axioms thy t
934 val types = Library.foldl (fn (acc, t') => acc union (ground_types thy t')) ([], t :: axioms)
935 val _ = writeln ("Ground types: "
936 ^ (if null types then "none."
937 else commas (map (Sign.string_of_typ (sign_of thy)) types)))
938 (* we can only consider fragments of recursive IDTs, so we issue a *)
939 (* warning if the formula contains a recursive IDT *)
940 (* TODO: no warning needed for /positive/ occurrences of IDTs *)
941 val _ = if Library.exists (fn
943 (case DatatypePackage.get_datatype thy s of
944 SOME info => (* inductive datatype *)
946 val index = #index info
947 val descr = #descr info
948 val (_, _, constrs) = (the o AList.lookup (op =) descr) index
950 (* recursive datatype? *)
951 Library.exists (fn (_, ds) => Library.exists DatatypeAux.is_rec_type ds) constrs
954 | _ => false) types then
955 warning "Term contains a recursive datatype; countermodel(s) may be spurious!"
958 (* (Term.typ * int) list -> unit *)
959 fun find_model_loop universe =
961 val init_model = (universe, [])
962 val init_args = {maxvars = maxvars, def_eq = false, next_idx = 1, bounds = [], wellformed = True}
963 val _ = immediate_output ("Translating term (sizes: " ^ commas (map (fn (_, n) => string_of_int n) universe) ^ ") ...")
964 (* translate 't' and all axioms *)
965 val ((model, args), intrs) = foldl_map (fn ((m, a), t') =>
967 val (i, m', a') = interpret thy m a t'
969 (* set 'def_eq' to 'true' *)
970 ((m', {maxvars = #maxvars a', def_eq = true, next_idx = #next_idx a', bounds = #bounds a', wellformed = #wellformed a'}), i)
971 end) ((init_model, init_args), t :: axioms)
972 (* make 't' either true or false, and make all axioms true, and *)
973 (* add the well-formedness side condition *)
974 val fm_t = (if negate then toFalse else toTrue) (hd intrs)
975 val fm_ax = PropLogic.all (map toTrue (tl intrs))
976 val fm = PropLogic.all [#wellformed args, fm_ax, fm_t]
978 immediate_output " invoking SAT solver...";
979 (case SatSolver.invoke_solver satsolver fm of
980 SatSolver.SATISFIABLE assignment =>
981 (writeln " model found!";
982 writeln ("*** Model found: ***\n" ^ print_model thy model (fn i => case assignment i of SOME b => b | NONE => true)))
983 | SatSolver.UNSATISFIABLE _ =>
984 (immediate_output " no model exists.\n";
985 case next_universe universe sizes minsize maxsize of
986 SOME universe' => find_model_loop universe'
987 | NONE => writeln "Search terminated, no larger universe within the given limits.")
988 | SatSolver.UNKNOWN =>
989 (immediate_output " no model found.\n";
990 case next_universe universe sizes minsize maxsize of
991 SOME universe' => find_model_loop universe'
992 | NONE => writeln "Search terminated, no larger universe within the given limits.")
993 ) handle SatSolver.NOT_CONFIGURED =>
994 error ("SAT solver " ^ quote satsolver ^ " is not configured.")
995 end handle MAXVARS_EXCEEDED =>
996 writeln ("\nSearch terminated, number of Boolean variables (" ^ string_of_int maxvars ^ " allowed) exceeded.")
998 find_model_loop (first_universe types sizes minsize)
1001 (* some parameter sanity checks *)
1002 assert (minsize>=1) ("\"minsize\" is " ^ string_of_int minsize ^ ", must be at least 1");
1003 assert (maxsize>=1) ("\"maxsize\" is " ^ string_of_int maxsize ^ ", must be at least 1");
1004 assert (maxsize>=minsize) ("\"maxsize\" (=" ^ string_of_int maxsize ^ ") is less than \"minsize\" (=" ^ string_of_int minsize ^ ").");
1005 assert (maxvars>=0) ("\"maxvars\" is " ^ string_of_int maxvars ^ ", must be at least 0");
1006 assert (maxtime>=0) ("\"maxtime\" is " ^ string_of_int maxtime ^ ", must be at least 0");
1007 (* enter loop with or without time limit *)
1008 writeln ("Trying to find a model that " ^ (if negate then "refutes" else "satisfies") ^ ": "
1009 ^ Sign.string_of_term (sign_of thy) t);
1011 interrupt_timeout (Time.fromSeconds (Int.toLarge maxtime))
1014 writeln ("\nSearch terminated, time limit ("
1015 ^ string_of_int maxtime ^ (if maxtime=1 then " second" else " seconds")
1022 (* ------------------------------------------------------------------------- *)
1023 (* INTERFACE, PART 2: FINDING A MODEL *)
1024 (* ------------------------------------------------------------------------- *)
1026 (* ------------------------------------------------------------------------- *)
1027 (* satisfy_term: calls 'find_model' to find a model that satisfies 't' *)
1028 (* params : list of '(name, value)' pairs used to override default *)
1030 (* ------------------------------------------------------------------------- *)
1032 (* theory -> (string * string) list -> Term.term -> unit *)
1034 fun satisfy_term thy params t =
1035 find_model thy (actual_params thy params) t false;
1037 (* ------------------------------------------------------------------------- *)
1038 (* refute_term: calls 'find_model' to find a model that refutes 't' *)
1039 (* params : list of '(name, value)' pairs used to override default *)
1041 (* ------------------------------------------------------------------------- *)
1043 (* theory -> (string * string) list -> Term.term -> unit *)
1045 fun refute_term thy params t =
1047 (* disallow schematic type variables, since we cannot properly negate *)
1048 (* terms containing them (their logical meaning is that there EXISTS a *)
1049 (* type s.t. ...; to refute such a formula, we would have to show that *)
1050 (* for ALL types, not ...) *)
1051 val _ = assert (null (term_tvars t)) "Term to be refuted contains schematic type variables"
1052 (* existential closure over schematic variables *)
1053 (* (Term.indexname * Term.typ) list *)
1054 val vars = sort_wrt (fst o fst) (map dest_Var (term_vars t))
1056 val ex_closure = Library.foldl
1057 (fn (t', ((x, i), T)) => (HOLogic.exists_const T) $ Abs (x, T, abstract_over (Var ((x, i), T), t')))
1059 (* If 't' is of type 'propT' (rather than 'boolT'), applying *)
1060 (* 'HOLogic.exists_const' is not type-correct. However, this *)
1061 (* is not really a problem as long as 'find_model' still *)
1062 (* interprets the resulting term correctly, without checking *)
1065 find_model thy (actual_params thy params) ex_closure true
1068 (* ------------------------------------------------------------------------- *)
1069 (* refute_subgoal: calls 'refute_term' on a specific subgoal *)
1070 (* params : list of '(name, value)' pairs used to override default *)
1072 (* subgoal : 0-based index specifying the subgoal number *)
1073 (* ------------------------------------------------------------------------- *)
1075 (* theory -> (string * string) list -> Thm.thm -> int -> unit *)
1077 fun refute_subgoal thy params thm subgoal =
1078 refute_term thy params (List.nth (prems_of thm, subgoal));
1081 (* ------------------------------------------------------------------------- *)
1082 (* INTERPRETERS: Auxiliary Functions *)
1083 (* ------------------------------------------------------------------------- *)
1085 (* ------------------------------------------------------------------------- *)
1086 (* make_constants: returns all interpretations that have the same tree *)
1087 (* structure as 'intr', but consist of unit vectors with *)
1088 (* 'True'/'False' only (no Boolean variables) *)
1089 (* ------------------------------------------------------------------------- *)
1091 (* interpretation -> interpretation list *)
1093 fun make_constants intr =
1095 (* returns a list with all unit vectors of length n *)
1096 (* int -> interpretation list *)
1097 fun unit_vectors n =
1099 (* returns the k-th unit vector of length n *)
1100 (* int * int -> interpretation *)
1101 fun unit_vector (k,n) =
1102 Leaf ((replicate (k-1) False) @ (True :: (replicate (n-k) False)))
1103 (* int -> interpretation list -> interpretation list *)
1104 fun unit_vectors_acc k vs =
1105 if k>n then [] else (unit_vector (k,n))::(unit_vectors_acc (k+1) vs)
1107 unit_vectors_acc 1 []
1109 (* concatenates 'x' with every list in 'xss', returning a new list of lists *)
1110 (* 'a -> 'a list list -> 'a list list *)
1111 fun cons_list x xss =
1112 map (fn xs => x::xs) xss
1113 (* returns a list of lists, each one consisting of n (possibly identical) elements from 'xs' *)
1114 (* int -> 'a list -> 'a list list *)
1116 map (fn x => [x]) xs
1118 let val rec_pick = pick_all (n-1) xs in
1119 Library.foldl (fn (acc, x) => (cons_list x rec_pick) @ acc) ([], xs)
1123 Leaf xs => unit_vectors (length xs)
1124 | Node xs => map (fn xs' => Node xs') (pick_all (length xs) (make_constants (hd xs)))
1127 (* ------------------------------------------------------------------------- *)
1128 (* size_of_type: returns the number of constants in a type (i.e. 'length *)
1129 (* (make_constants intr)', but implemented more efficiently) *)
1130 (* ------------------------------------------------------------------------- *)
1132 (* interpretation -> int *)
1134 fun size_of_type intr =
1136 (* power (a, b) computes a^b, for a>=0, b>=0 *)
1137 (* int * int -> int *)
1138 fun power (a, 0) = 1
1140 | power (a, b) = let val ab = power(a, b div 2) in ab * ab * power(a, b mod 2) end
1143 Leaf xs => length xs
1144 | Node xs => power (size_of_type (hd xs), length xs)
1147 (* ------------------------------------------------------------------------- *)
1148 (* TT/FF: interpretations that denote "true" or "false", respectively *)
1149 (* ------------------------------------------------------------------------- *)
1151 (* interpretation *)
1153 val TT = Leaf [True, False];
1155 val FF = Leaf [False, True];
1157 (* ------------------------------------------------------------------------- *)
1158 (* make_equality: returns an interpretation that denotes (extensional) *)
1159 (* equality of two interpretations *)
1160 (* - two interpretations are 'equal' iff they are both defined and denote *)
1161 (* the same value *)
1162 (* - two interpretations are 'not_equal' iff they are both defined at least *)
1163 (* partially, and a defined part denotes different values *)
1164 (* - a completely undefined interpretation is neither 'equal' nor *)
1165 (* 'not_equal' to another interpretation *)
1166 (* ------------------------------------------------------------------------- *)
1168 (* We could in principle represent '=' on a type T by a particular *)
1169 (* interpretation. However, the size of that interpretation is quadratic *)
1170 (* in the size of T. Therefore comparing the interpretations 'i1' and *)
1171 (* 'i2' directly is more efficient than constructing the interpretation *)
1172 (* for equality on T first, and "applying" this interpretation to 'i1' *)
1173 (* and 'i2' in the usual way (cf. 'interpretation_apply') then. *)
1175 (* interpretation * interpretation -> interpretation *)
1177 fun make_equality (i1, i2) =
1179 (* interpretation * interpretation -> prop_formula *)
1180 fun equal (i1, i2) =
1184 Leaf ys => PropLogic.dot_product (xs, ys) (* defined and equal *)
1185 | Node _ => raise REFUTE ("make_equality", "second interpretation is higher"))
1188 Leaf _ => raise REFUTE ("make_equality", "first interpretation is higher")
1189 | Node ys => PropLogic.all (map equal (xs ~~ ys))))
1190 (* interpretation * interpretation -> prop_formula *)
1191 fun not_equal (i1, i2) =
1195 Leaf ys => PropLogic.all ((PropLogic.exists xs) :: (PropLogic.exists ys) ::
1196 (map (fn (x,y) => SOr (SNot x, SNot y)) (xs ~~ ys))) (* defined and not equal *)
1197 | Node _ => raise REFUTE ("make_equality", "second interpretation is higher"))
1200 Leaf _ => raise REFUTE ("make_equality", "first interpretation is higher")
1201 | Node ys => PropLogic.exists (map not_equal (xs ~~ ys))))
1203 (* a value may be undefined; therefore 'not_equal' is not just the *)
1204 (* negation of 'equal' *)
1205 Leaf [equal (i1, i2), not_equal (i1, i2)]
1208 (* ------------------------------------------------------------------------- *)
1209 (* make_def_equality: returns an interpretation that denotes (extensional) *)
1210 (* equality of two interpretations *)
1211 (* This function treats undefined/partially defined interpretations *)
1212 (* different from 'make_equality': two undefined interpretations are *)
1213 (* considered equal, while a defined interpretation is considered not equal *)
1214 (* to an undefined interpretation. *)
1215 (* ------------------------------------------------------------------------- *)
1217 (* interpretation * interpretation -> interpretation *)
1219 fun make_def_equality (i1, i2) =
1221 (* interpretation * interpretation -> prop_formula *)
1222 fun equal (i1, i2) =
1226 Leaf ys => SOr (PropLogic.dot_product (xs, ys), (* defined and equal, or both undefined *)
1227 SAnd (PropLogic.all (map SNot xs), PropLogic.all (map SNot ys)))
1228 | Node _ => raise REFUTE ("make_def_equality", "second interpretation is higher"))
1231 Leaf _ => raise REFUTE ("make_def_equality", "first interpretation is higher")
1232 | Node ys => PropLogic.all (map equal (xs ~~ ys))))
1233 (* interpretation *)
1234 val eq = equal (i1, i2)
1239 (* ------------------------------------------------------------------------- *)
1240 (* interpretation_apply: returns an interpretation that denotes the result *)
1241 (* of applying the function denoted by 'i2' to the *)
1242 (* argument denoted by 'i2' *)
1243 (* ------------------------------------------------------------------------- *)
1245 (* interpretation * interpretation -> interpretation *)
1247 fun interpretation_apply (i1, i2) =
1249 (* interpretation * interpretation -> interpretation *)
1250 fun interpretation_disjunction (tr1,tr2) =
1251 tree_map (fn (xs,ys) => map (fn (x,y) => SOr(x,y)) (xs ~~ ys)) (tree_pair (tr1,tr2))
1252 (* prop_formula * interpretation -> interpretation *)
1253 fun prop_formula_times_interpretation (fm,tr) =
1254 tree_map (map (fn x => SAnd (fm,x))) tr
1255 (* prop_formula list * interpretation list -> interpretation *)
1256 fun prop_formula_list_dot_product_interpretation_list ([fm],[tr]) =
1257 prop_formula_times_interpretation (fm,tr)
1258 | prop_formula_list_dot_product_interpretation_list (fm::fms,tr::trees) =
1259 interpretation_disjunction (prop_formula_times_interpretation (fm,tr), prop_formula_list_dot_product_interpretation_list (fms,trees))
1260 | prop_formula_list_dot_product_interpretation_list (_,_) =
1261 raise REFUTE ("interpretation_apply", "empty list (in dot product)")
1262 (* concatenates 'x' with every list in 'xss', returning a new list of lists *)
1263 (* 'a -> 'a list list -> 'a list list *)
1264 fun cons_list x xss =
1265 map (fn xs => x::xs) xss
1266 (* returns a list of lists, each one consisting of one element from each element of 'xss' *)
1267 (* 'a list list -> 'a list list *)
1269 map (fn x => [x]) xs
1270 | pick_all (xs::xss) =
1271 let val rec_pick = pick_all xss in
1272 Library.foldl (fn (acc, x) => (cons_list x rec_pick) @ acc) ([], xs)
1275 raise REFUTE ("interpretation_apply", "empty list (in pick_all)")
1276 (* interpretation -> prop_formula list *)
1277 fun interpretation_to_prop_formula_list (Leaf xs) =
1279 | interpretation_to_prop_formula_list (Node trees) =
1280 map PropLogic.all (pick_all (map interpretation_to_prop_formula_list trees))
1284 raise REFUTE ("interpretation_apply", "first interpretation is a leaf")
1286 prop_formula_list_dot_product_interpretation_list (interpretation_to_prop_formula_list i2, xs)
1289 (* ------------------------------------------------------------------------- *)
1290 (* eta_expand: eta-expands a term 't' by adding 'i' lambda abstractions *)
1291 (* ------------------------------------------------------------------------- *)
1293 (* Term.term -> int -> Term.term *)
1295 fun eta_expand t i =
1297 val Ts = binder_types (fastype_of t)
1299 foldr (fn (T, t) => Abs ("<eta_expand>", T, t))
1300 (list_comb (t, map Bound (i-1 downto 0))) (Library.take (i, Ts))
1303 (* ------------------------------------------------------------------------- *)
1304 (* sum: returns the sum of a list 'xs' of integers *)
1305 (* ------------------------------------------------------------------------- *)
1307 (* int list -> int *)
1309 fun sum xs = foldl op+ 0 xs;
1311 (* ------------------------------------------------------------------------- *)
1312 (* product: returns the product of a list 'xs' of integers *)
1313 (* ------------------------------------------------------------------------- *)
1315 (* int list -> int *)
1317 fun product xs = foldl op* 1 xs;
1319 (* ------------------------------------------------------------------------- *)
1320 (* size_of_dtyp: the size of (an initial fragment of) an inductive data type *)
1321 (* is the sum (over its constructors) of the product (over *)
1322 (* their arguments) of the size of the argument types *)
1323 (* ------------------------------------------------------------------------- *)
1325 (* theory -> (Term.typ * int) list -> DatatypeAux.descr -> (DatatypeAux.dtyp * Term.typ) list -> (string * DatatypeAux.dtyp list) list -> int *)
1327 fun size_of_dtyp thy typ_sizes descr typ_assoc constructors =
1328 sum (map (fn (_, dtyps) =>
1329 product (map (fn dtyp =>
1331 val T = typ_of_dtyp descr typ_assoc dtyp
1332 val (i, _, _) = interpret thy (typ_sizes, []) {maxvars=0, def_eq = false, next_idx=1, bounds=[], wellformed=True} (Free ("dummy", T))
1335 end) dtyps)) constructors);
1338 (* ------------------------------------------------------------------------- *)
1339 (* INTERPRETERS: Actual Interpreters *)
1340 (* ------------------------------------------------------------------------- *)
1342 (* theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option *)
1344 (* simply typed lambda calculus: Isabelle's basic term syntax, with type *)
1345 (* variables, function types, and propT *)
1347 fun stlc_interpreter thy model args t =
1349 val (typs, terms) = model
1350 val {maxvars, def_eq, next_idx, bounds, wellformed} = args
1351 (* Term.typ -> (interpretation * model * arguments) option *)
1352 fun interpret_groundterm T =
1354 (* unit -> (interpretation * model * arguments) option *)
1355 fun interpret_groundtype () =
1357 val size = (if T = Term.propT then 2 else (the o AList.lookup (op =) typs) T) (* the model MUST specify a size for ground types *)
1358 val next = next_idx+size
1359 val _ = (if next-1>maxvars andalso maxvars>0 then raise MAXVARS_EXCEEDED else ()) (* check if 'maxvars' is large enough *)
1360 (* prop_formula list *)
1361 val fms = map BoolVar (next_idx upto (next_idx+size-1))
1362 (* interpretation *)
1364 (* prop_formula list -> prop_formula *)
1365 fun one_of_two_false [] = True
1366 | one_of_two_false (x::xs) = SAnd (PropLogic.all (map (fn x' => SOr (SNot x, SNot x')) xs), one_of_two_false xs)
1368 val wf = one_of_two_false fms
1370 (* extend the model, increase 'next_idx', add well-formedness condition *)
1371 SOME (intr, (typs, (t, intr)::terms), {maxvars = maxvars, def_eq = def_eq, next_idx = next, bounds = bounds, wellformed = SAnd (wellformed, wf)})
1375 Type ("fun", [T1, T2]) =>
1377 (* we create 'size_of_type (interpret (... T1))' different copies *)
1378 (* of the interpretation for 'T2', which are then combined into a *)
1379 (* single new interpretation *)
1380 val (i1, _, _) = interpret thy model {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True} (Free ("dummy", T1))
1381 (* make fresh copies, with different variable indices *)
1382 (* 'idx': next variable index *)
1383 (* 'n' : number of copies *)
1384 (* int -> int -> (int * interpretation list * prop_formula *)
1385 fun make_copies idx 0 =
1387 | make_copies idx n =
1389 val (copy, _, new_args) = interpret thy (typs, []) {maxvars = maxvars, def_eq = false, next_idx = idx, bounds = [], wellformed = True} (Free ("dummy", T2))
1390 val (idx', copies, wf') = make_copies (#next_idx new_args) (n-1)
1392 (idx', copy :: copies, SAnd (#wellformed new_args, wf'))
1394 val (next, copies, wf) = make_copies next_idx (size_of_type i1)
1395 (* combine copies into a single interpretation *)
1396 val intr = Node copies
1398 (* extend the model, increase 'next_idx', add well-formedness condition *)
1399 SOME (intr, (typs, (t, intr)::terms), {maxvars = maxvars, def_eq = def_eq, next_idx = next, bounds = bounds, wellformed = SAnd (wellformed, wf)})
1401 | Type _ => interpret_groundtype ()
1402 | TFree _ => interpret_groundtype ()
1403 | TVar _ => interpret_groundtype ()
1406 case AList.lookup (op =) terms t of
1408 (* return an existing interpretation *)
1409 SOME (intr, model, args)
1413 interpret_groundterm T
1415 interpret_groundterm T
1417 interpret_groundterm T
1419 SOME (List.nth (#bounds args, i), model, args)
1420 | Abs (x, T, body) =>
1422 (* create all constants of type 'T' *)
1423 val (i, _, _) = interpret thy model {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True} (Free ("dummy", T))
1424 val constants = make_constants i
1425 (* interpret the 'body' separately for each constant *)
1426 val ((model', args'), bodies) = foldl_map
1429 (* add 'c' to 'bounds' *)
1430 val (i', m', a') = interpret thy m {maxvars = #maxvars a, def_eq = #def_eq a, next_idx = #next_idx a, bounds = (c :: #bounds a), wellformed = #wellformed a} body
1432 (* keep the new model m' and 'next_idx' and 'wellformed', but use old 'bounds' *)
1433 ((m', {maxvars = maxvars, def_eq = def_eq, next_idx = #next_idx a', bounds = bounds, wellformed = #wellformed a'}), i')
1435 ((model, args), constants)
1437 SOME (Node bodies, model', args')
1441 (* interpret 't1' and 't2' separately *)
1442 val (intr1, model1, args1) = interpret thy model args t1
1443 val (intr2, model2, args2) = interpret thy model1 args1 t2
1445 SOME (interpretation_apply (intr1, intr2), model2, args2)
1449 (* theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option *)
1451 fun Pure_interpreter thy model args t =
1453 Const ("all", _) $ t1 => (* in the meta-logic, 'all' MUST be followed by an argument term *)
1455 val (i, m, a) = interpret thy model args t1
1460 val fmTrue = PropLogic.all (map toTrue xs)
1461 val fmFalse = PropLogic.exists (map toFalse xs)
1463 SOME (Leaf [fmTrue, fmFalse], m, a)
1466 raise REFUTE ("Pure_interpreter", "\"all\" is not followed by a function")
1468 | Const ("==", _) $ t1 $ t2 =>
1470 val (i1, m1, a1) = interpret thy model args t1
1471 val (i2, m2, a2) = interpret thy m1 a1 t2
1473 (* we use either 'make_def_equality' or 'make_equality' *)
1474 SOME ((if #def_eq args then make_def_equality else make_equality) (i1, i2), m2, a2)
1476 | Const ("==>", _) => (* simpler than translating 'Const ("==>", _) $ t1 $ t2' *)
1477 SOME (Node [Node [TT, FF], Node [TT, TT]], model, args)
1480 (* theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option *)
1482 fun HOLogic_interpreter thy model args t =
1483 (* ------------------------------------------------------------------------- *)
1484 (* Providing interpretations directly is more efficient than unfolding the *)
1485 (* logical constants. In HOL however, logical constants can themselves be *)
1486 (* arguments. "All" and "Ex" are then translated just like any other *)
1487 (* constant, with the relevant axiom being added by 'collect_axioms'. *)
1488 (* ------------------------------------------------------------------------- *)
1490 Const ("Trueprop", _) =>
1491 SOME (Node [TT, FF], model, args)
1492 | Const ("Not", _) =>
1493 SOME (Node [FF, TT], model, args)
1494 | Const ("True", _) => (* redundant, since 'True' is also an IDT constructor *)
1495 SOME (TT, model, args)
1496 | Const ("False", _) => (* redundant, since 'False' is also an IDT constructor *)
1497 SOME (FF, model, args)
1498 | Const ("All", _) $ t1 =>
1499 (* if "All" occurs without an argument (i.e. as argument to a higher-order *)
1500 (* function or predicate), it is handled by the 'stlc_interpreter' (i.e. *)
1501 (* by unfolding its definition) *)
1503 val (i, m, a) = interpret thy model args t1
1508 val fmTrue = PropLogic.all (map toTrue xs)
1509 val fmFalse = PropLogic.exists (map toFalse xs)
1511 SOME (Leaf [fmTrue, fmFalse], m, a)
1514 raise REFUTE ("HOLogic_interpreter", "\"All\" is followed by a non-function")
1516 | Const ("Ex", _) $ t1 =>
1517 (* if "Ex" occurs without an argument (i.e. as argument to a higher-order *)
1518 (* function or predicate), it is handled by the 'stlc_interpreter' (i.e. *)
1519 (* by unfolding its definition) *)
1521 val (i, m, a) = interpret thy model args t1
1526 val fmTrue = PropLogic.exists (map toTrue xs)
1527 val fmFalse = PropLogic.all (map toFalse xs)
1529 SOME (Leaf [fmTrue, fmFalse], m, a)
1532 raise REFUTE ("HOLogic_interpreter", "\"Ex\" is followed by a non-function")
1534 | Const ("op =", _) $ t1 $ t2 =>
1536 val (i1, m1, a1) = interpret thy model args t1
1537 val (i2, m2, a2) = interpret thy m1 a1 t2
1539 SOME (make_equality (i1, i2), m2, a2)
1541 | Const ("op =", _) $ t1 =>
1542 SOME (interpret thy model args (eta_expand t 1))
1543 | Const ("op =", _) =>
1544 SOME (interpret thy model args (eta_expand t 2))
1545 | Const ("op &", _) $ t1 $ t2 =>
1546 (* 3-valued logic *)
1548 val (i1, m1, a1) = interpret thy model args t1
1549 val (i2, m2, a2) = interpret thy m1 a1 t2
1550 val fmTrue = PropLogic.SAnd (toTrue i1, toTrue i2)
1551 val fmFalse = PropLogic.SOr (toFalse i1, toFalse i2)
1553 SOME (Leaf [fmTrue, fmFalse], m2, a2)
1555 | Const ("op &", _) $ t1 =>
1556 SOME (interpret thy model args (eta_expand t 1))
1557 | Const ("op &", _) =>
1558 SOME (interpret thy model args (eta_expand t 2))
1559 (* SOME (Node [Node [TT, FF], Node [FF, FF]], model, args) *)
1560 | Const ("op |", _) $ t1 $ t2 =>
1561 (* 3-valued logic *)
1563 val (i1, m1, a1) = interpret thy model args t1
1564 val (i2, m2, a2) = interpret thy m1 a1 t2
1565 val fmTrue = PropLogic.SOr (toTrue i1, toTrue i2)
1566 val fmFalse = PropLogic.SAnd (toFalse i1, toFalse i2)
1568 SOME (Leaf [fmTrue, fmFalse], m2, a2)
1570 | Const ("op |", _) $ t1 =>
1571 SOME (interpret thy model args (eta_expand t 1))
1572 | Const ("op |", _) =>
1573 SOME (interpret thy model args (eta_expand t 2))
1574 (* SOME (Node [Node [TT, TT], Node [TT, FF]], model, args) *)
1575 | Const ("op -->", _) $ t1 $ t2 =>
1576 (* 3-valued logic *)
1578 val (i1, m1, a1) = interpret thy model args t1
1579 val (i2, m2, a2) = interpret thy m1 a1 t2
1580 val fmTrue = PropLogic.SOr (toFalse i1, toTrue i2)
1581 val fmFalse = PropLogic.SAnd (toTrue i1, toFalse i2)
1583 SOME (Leaf [fmTrue, fmFalse], m2, a2)
1585 | Const ("op -->", _) =>
1586 (* SOME (Node [Node [TT, FF], Node [TT, TT]], model, args) *)
1587 SOME (interpret thy model args (eta_expand t 2))
1590 (* theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option *)
1592 fun set_interpreter thy model args t =
1593 (* "T set" is isomorphic to "T --> bool" *)
1595 val (typs, terms) = model
1597 case AList.lookup (op =) terms t of
1599 (* return an existing interpretation *)
1600 SOME (intr, model, args)
1603 Free (x, Type ("set", [T])) =>
1605 val (intr, _, args') = interpret thy (typs, []) args (Free (x, T --> HOLogic.boolT))
1607 SOME (intr, (typs, (t, intr)::terms), args')
1609 | Var ((x, i), Type ("set", [T])) =>
1611 val (intr, _, args') = interpret thy (typs, []) args (Var ((x,i), T --> HOLogic.boolT))
1613 SOME (intr, (typs, (t, intr)::terms), args')
1615 | Const (s, Type ("set", [T])) =>
1617 val (intr, _, args') = interpret thy (typs, []) args (Const (s, T --> HOLogic.boolT))
1619 SOME (intr, (typs, (t, intr)::terms), args')
1621 (* 'Collect' == identity *)
1622 | Const ("Collect", _) $ t1 =>
1623 SOME (interpret thy model args t1)
1624 | Const ("Collect", _) =>
1625 SOME (interpret thy model args (eta_expand t 1))
1626 (* 'op :' == application *)
1627 | Const ("op :", _) $ t1 $ t2 =>
1628 SOME (interpret thy model args (t2 $ t1))
1629 | Const ("op :", _) $ t1 =>
1630 SOME (interpret thy model args (eta_expand t 1))
1631 | Const ("op :", _) =>
1632 SOME (interpret thy model args (eta_expand t 2))
1636 (* theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option *)
1638 (* interprets variables and constants whose type is an IDT; constructors of *)
1639 (* IDTs are properly interpreted by 'IDT_constructor_interpreter' however *)
1641 fun IDT_interpreter thy model args t =
1643 val (typs, terms) = model
1644 (* Term.typ -> (interpretation * model * arguments) option *)
1645 fun interpret_term (Type (s, Ts)) =
1646 (case DatatypePackage.get_datatype thy s of
1647 SOME info => (* inductive datatype *)
1649 (* int option -- only recursive IDTs have an associated depth *)
1650 val depth = AList.lookup (op =) typs (Type (s, Ts))
1652 if depth = (SOME 0) then (* termination condition to avoid infinite recursion *)
1653 (* return a leaf of size 0 *)
1654 SOME (Leaf [], model, args)
1657 val index = #index info
1658 val descr = #descr info
1659 val (_, dtyps, constrs) = (the o AList.lookup (op =) descr) index
1660 val typ_assoc = dtyps ~~ Ts
1661 (* sanity check: every element in 'dtyps' must be a 'DtTFree' *)
1662 val _ = (if Library.exists (fn d =>
1663 case d of DatatypeAux.DtTFree _ => false | _ => true) dtyps
1665 raise REFUTE ("IDT_interpreter", "datatype argument (for type " ^ Sign.string_of_typ (sign_of thy) (Type (s, Ts)) ^ ") is not a variable")
1668 (* if the model specifies a depth for the current type, decrement it to avoid infinite recursion *)
1669 val typs' = case depth of NONE => typs | SOME n =>
1670 AList.update (op =) (Type (s, Ts), n-1) typs
1671 (* recursively compute the size of the datatype *)
1672 val size = size_of_dtyp thy typs' descr typ_assoc constrs
1673 val next_idx = #next_idx args
1674 val next = next_idx+size
1675 val _ = (if next-1>(#maxvars args) andalso (#maxvars args)>0 then raise MAXVARS_EXCEEDED else ()) (* check if 'maxvars' is large enough *)
1676 (* prop_formula list *)
1677 val fms = map BoolVar (next_idx upto (next_idx+size-1))
1678 (* interpretation *)
1680 (* prop_formula list -> prop_formula *)
1681 fun one_of_two_false [] = True
1682 | one_of_two_false (x::xs) = SAnd (PropLogic.all (map (fn x' => SOr (SNot x, SNot x')) xs), one_of_two_false xs)
1684 val wf = one_of_two_false fms
1686 (* extend the model, increase 'next_idx', add well-formedness condition *)
1687 SOME (intr, (typs, (t, intr)::terms), {maxvars = #maxvars args, def_eq = #def_eq args, next_idx = next, bounds = #bounds args, wellformed = SAnd (#wellformed args, wf)})
1690 | NONE => (* not an inductive datatype *)
1692 | interpret_term _ = (* a (free or schematic) type variable *)
1695 case AList.lookup (op =) terms t of
1697 (* return an existing interpretation *)
1698 SOME (intr, model, args)
1701 Free (_, T) => interpret_term T
1702 | Var (_, T) => interpret_term T
1703 | Const (_, T) => interpret_term T
1707 (* theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option *)
1709 fun IDT_constructor_interpreter thy model args t =
1711 val (typs, terms) = model
1713 case AList.lookup (op =) terms t of
1715 (* return an existing interpretation *)
1716 SOME (intr, model, args)
1720 (case body_type T of
1722 (case DatatypePackage.get_datatype thy s' of
1723 SOME info => (* body type is an inductive datatype *)
1725 val index = #index info
1726 val descr = #descr info
1727 val (_, dtyps, constrs) = (the o AList.lookup (op =) descr) index
1728 val typ_assoc = dtyps ~~ Ts'
1729 (* sanity check: every element in 'dtyps' must be a 'DtTFree' *)
1730 val _ = (if Library.exists (fn d =>
1731 case d of DatatypeAux.DtTFree _ => false | _ => true) dtyps
1733 raise REFUTE ("IDT_constructor_interpreter", "datatype argument (for type " ^ Sign.string_of_typ (sign_of thy) (Type (s', Ts')) ^ ") is not a variable")
1736 (* split the constructors into those occuring before/after 'Const (s, T)' *)
1737 val (constrs1, constrs2) = take_prefix (fn (cname, ctypes) =>
1738 not (cname = s andalso Sign.typ_instance thy (T,
1739 map (typ_of_dtyp descr typ_assoc) ctypes ---> Type (s', Ts')))) constrs
1743 (* 'Const (s, T)' is not a constructor of this datatype *)
1745 | (_, ctypes)::cs =>
1747 (* compute the total size of the datatype (with the current depth) *)
1748 val (i, _, _) = interpret thy (typs, []) {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True} (Free ("dummy", Type (s', Ts')))
1749 val total = size_of_type i
1750 (* int option -- only recursive IDTs have an associated depth *)
1751 val depth = AList.lookup (op =) typs (Type (s', Ts'))
1752 val typs' = (case depth of NONE => typs | SOME n =>
1753 AList.update (op =) (Type (s', Ts'), n-1) typs)
1754 (* returns an interpretation where everything is mapped to "undefined" *)
1755 (* DatatypeAux.dtyp list -> interpretation *)
1757 Leaf (replicate total False)
1758 | make_undef (d::ds) =
1760 (* compute the current size of the type 'd' *)
1761 val T = typ_of_dtyp descr typ_assoc d
1762 val (i, _, _) = interpret thy (typs, []) {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True} (Free ("dummy", T))
1763 val size = size_of_type i
1765 Node (replicate size (make_undef ds))
1767 (* returns the interpretation for a constructor at depth 1 *)
1768 (* int * DatatypeAux.dtyp list -> int * interpretation *)
1769 fun make_constr (offset, []) =
1770 if offset<total then
1771 (offset+1, Leaf ((replicate offset False) @ True :: (replicate (total-offset-1) False)))
1773 raise REFUTE ("IDT_constructor_interpreter", "offset >= total")
1774 | make_constr (offset, d::ds) =
1776 (* compute the current and the old size of the type 'd' *)
1777 val T = typ_of_dtyp descr typ_assoc d
1778 val (i, _, _) = interpret thy (typs, []) {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True} (Free ("dummy", T))
1779 val size = size_of_type i
1780 val (i', _, _) = interpret thy (typs', []) {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True} (Free ("dummy", T))
1781 val size' = size_of_type i'
1783 val _ = if size < size' then
1784 raise REFUTE ("IDT_constructor_interpreter", "current size is less than old size")
1787 (* int * interpretation list *)
1788 val (new_offset, intrs) = foldl_map make_constr (offset, replicate size' ds)
1789 (* interpretation list *)
1790 val undefs = replicate (size - size') (make_undef ds)
1792 (* elements that exist at the previous depth are mapped to a defined *)
1793 (* value, while new elements are mapped to "undefined" by the *)
1794 (* recursive constructor *)
1795 (new_offset, Node (intrs @ undefs))
1797 (* extends the interpretation for a constructor (both recursive *)
1798 (* and non-recursive) obtained at depth n (n>=1) to depth n+1 *)
1799 (* int * DatatypeAux.dtyp list * interpretation -> int * interpretation *)
1800 fun extend_constr (offset, [], Leaf xs) =
1802 (* returns the k-th unit vector of length n *)
1803 (* int * int -> interpretation *)
1804 fun unit_vector (k, n) =
1805 Leaf ((replicate (k-1) False) @ (True :: (replicate (n-k) False)))
1807 val k = find_index_eq True xs
1810 (* if the element was mapped to "undefined" before, map it to *)
1811 (* the value given by 'offset' now (and extend the length of *)
1813 (offset+1, unit_vector (offset+1, total))
1815 (* if the element was already mapped to a defined value, map it *)
1816 (* to the same value again, just extend the length of the leaf, *)
1817 (* do not increment the 'offset' *)
1818 (offset, unit_vector (k+1, total))
1820 | extend_constr (_, [], Node _) =
1821 raise REFUTE ("IDT_constructor_interpreter", "interpretation for constructor (with no arguments left) is a node")
1822 | extend_constr (offset, d::ds, Node xs) =
1824 (* compute the size of the type 'd' *)
1825 val T = typ_of_dtyp descr typ_assoc d
1826 val (i, _, _) = interpret thy (typs, []) {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True} (Free ("dummy", T))
1827 val size = size_of_type i
1829 val _ = if size < length xs then
1830 raise REFUTE ("IDT_constructor_interpreter", "new size of type is less than old size")
1833 (* extend the existing interpretations *)
1834 (* int * interpretation list *)
1835 val (new_offset, intrs) = foldl_map (fn (off, i) => extend_constr (off, ds, i)) (offset, xs)
1836 (* new elements of the type 'd' are mapped to "undefined" *)
1837 val undefs = replicate (size - length xs) (make_undef ds)
1839 (new_offset, Node (intrs @ undefs))
1841 | extend_constr (_, d::ds, Leaf _) =
1842 raise REFUTE ("IDT_constructor_interpreter", "interpretation for constructor (with arguments left) is a leaf")
1843 (* returns 'true' iff the constructor has a recursive argument *)
1844 (* DatatypeAux.dtyp list -> bool *)
1845 fun is_rec_constr ds =
1846 Library.exists DatatypeAux.is_rec_type ds
1847 (* constructors before 'Const (s, T)' generate elements of the datatype *)
1848 val offset = size_of_dtyp thy typs' descr typ_assoc constrs1
1851 NONE => (* equivalent to a depth of 1 *)
1852 SOME (snd (make_constr (offset, ctypes)), model, args)
1854 raise REFUTE ("IDT_constructor_interpreter", "depth is 0")
1856 SOME (snd (make_constr (offset, ctypes)), model, args)
1857 | SOME n => (* n > 1 *)
1859 (* interpret the constructor at depth-1 *)
1860 val (iC, _, _) = interpret thy (typs', []) {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True} (Const (s, T))
1861 (* elements generated by the constructor at depth-1 must be added to 'offset' *)
1862 (* interpretation -> int *)
1863 fun number_of_defined_elements (Leaf xs) =
1864 if find_index_eq True xs = (~1) then 0 else 1
1865 | number_of_defined_elements (Node xs) =
1866 sum (map number_of_defined_elements xs)
1868 val offset' = offset + number_of_defined_elements iC
1870 SOME (snd (extend_constr (offset', ctypes, iC)), model, args)
1874 | NONE => (* body type is not an inductive datatype *)
1876 | _ => (* body type is a (free or schematic) type variable *)
1878 | _ => (* term is not a constant *)
1882 (* theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option *)
1884 (* Difficult code ahead. Make sure you understand the 'IDT_constructor_interpreter' *)
1885 (* and the order in which it enumerates elements of an IDT before you try to *)
1886 (* understand this function. *)
1888 fun IDT_recursion_interpreter thy model args t =
1889 case strip_comb t of (* careful: here we descend arbitrarily deep into 't', *)
1890 (* possibly before any other interpreter for atomic *)
1891 (* terms has had a chance to look at 't' *)
1892 (Const (s, T), params) =>
1893 (* iterate over all datatypes in 'thy' *)
1894 Symtab.fold (fn (_, info) => fn result =>
1897 result (* just keep 'result' *)
1899 if member (op =) (#rec_names info) s then
1900 (* we do have a recursion operator of the datatype given by 'info', *)
1901 (* or of a mutually recursive datatype *)
1903 val index = #index info
1904 val descr = #descr info
1905 val (dtname, dtyps, _) = (the o AList.lookup (op =) descr) index
1906 (* number of all constructors, including those of different *)
1907 (* (mutually recursive) datatypes within the same descriptor 'descr' *)
1908 val mconstrs_count = sum (map (fn (_, (_, _, cs)) => length cs) descr)
1909 val params_count = length params
1910 (* the type of a recursion operator: [T1, ..., Tn, IDT] ---> Tresult *)
1911 val IDT = List.nth (binder_types T, mconstrs_count)
1913 if (fst o dest_Type) IDT <> dtname then
1914 (* recursion operator of a mutually recursive datatype *)
1916 else if mconstrs_count < params_count then
1917 (* too many actual parameters; for now we'll use the *)
1918 (* 'stlc_interpreter' to strip off one application *)
1920 else if mconstrs_count > params_count then
1921 (* too few actual parameters; we use eta expansion *)
1922 (* Note that the resulting expansion of lambda abstractions *)
1923 (* by the 'stlc_interpreter' may be rather slow (depending on *)
1924 (* the argument types and the size of the IDT, of course). *)
1925 SOME (interpret thy model args (eta_expand t (mconstrs_count - params_count)))
1926 else (* mconstrs_count = params_count *)
1928 (* interpret each parameter separately *)
1929 val ((model', args'), p_intrs) = foldl_map (fn ((m, a), p) =>
1931 val (i, m', a') = interpret thy m a p
1934 end) ((model, args), params)
1935 val (typs, _) = model'
1936 val typ_assoc = dtyps ~~ (snd o dest_Type) IDT
1937 (* interpret each constructor in the descriptor (including *)
1938 (* those of mutually recursive datatypes) *)
1939 (* (int * interpretation list) list *)
1940 val mc_intrs = map (fn (idx, (_, _, cs)) =>
1942 val c_return_typ = typ_of_dtyp descr typ_assoc (DatatypeAux.DtRec idx)
1944 (idx, map (fn (cname, cargs) =>
1945 (#1 o interpret thy (typs, []) {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True})
1946 (Const (cname, map (typ_of_dtyp descr typ_assoc) cargs ---> c_return_typ))) cs)
1948 (* the recursion operator is a function that maps every element of *)
1949 (* the inductive datatype (and of mutually recursive types) to an *)
1950 (* element of some result type; an array entry of NONE means that *)
1951 (* the actual result has not been computed yet *)
1952 (* (int * interpretation option Array.array) list *)
1953 val INTRS = map (fn (idx, _) =>
1955 val T = typ_of_dtyp descr typ_assoc (DatatypeAux.DtRec idx)
1956 val (i, _, _) = interpret thy (typs, []) {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True} (Free ("dummy", T))
1957 val size = size_of_type i
1959 (idx, Array.array (size, NONE))
1961 (* takes an interpretation, and if some leaf of this interpretation *)
1962 (* is the 'elem'-th element of the type, the indices of the arguments *)
1963 (* leading to this leaf are returned *)
1964 (* interpretation -> int -> int list option *)
1965 fun get_args (Leaf xs) elem =
1966 if find_index_eq True xs = elem then
1970 | get_args (Node xs) elem =
1972 (* interpretation * int -> int list option *)
1973 fun search ([], _) =
1975 | search (x::xs, n) =
1976 (case get_args x elem of
1977 SOME result => SOME (n::result)
1978 | NONE => search (xs, n+1))
1982 (* returns the index of the constructor and indices for its *)
1983 (* arguments that generate the 'elem'-th element of the datatype *)
1984 (* given by 'idx' *)
1985 (* int -> int -> int * int list *)
1986 fun get_cargs idx elem =
1988 (* int * interpretation list -> int * int list *)
1989 fun get_cargs_rec (_, []) =
1990 raise REFUTE ("IDT_recursion_interpreter", "no matching constructor found for element " ^ string_of_int elem ^ " in datatype " ^ Sign.string_of_typ (sign_of thy) IDT ^ " (datatype index " ^ string_of_int idx ^ ")")
1991 | get_cargs_rec (n, x::xs) =
1992 (case get_args x elem of
1993 SOME args => (n, args)
1994 | NONE => get_cargs_rec (n+1, xs))
1996 get_cargs_rec (0, (the o AList.lookup (op =) mc_intrs) idx)
1998 (* returns the number of constructors in datatypes that occur in *)
1999 (* the descriptor 'descr' before the datatype given by 'idx' *)
2000 fun get_coffset idx =
2002 fun get_coffset_acc _ [] =
2003 raise REFUTE ("IDT_recursion_interpreter", "index " ^ string_of_int idx ^ " not found in descriptor")
2004 | get_coffset_acc sum ((i, (_, _, cs))::descr') =
2008 get_coffset_acc (sum + length cs) descr'
2010 get_coffset_acc 0 descr
2012 (* computes one entry in INTRS, and recursively all entries needed for it, *)
2013 (* where 'idx' gives the datatype and 'elem' the element of it *)
2014 (* int -> int -> interpretation *)
2015 fun compute_array_entry idx elem =
2016 case Array.sub ((the o AList.lookup (op =) INTRS) idx, elem) of
2018 (* simply return the previously computed result *)
2022 (* int * int list *)
2023 val (c, args) = get_cargs idx elem
2024 (* interpretation * int list -> interpretation *)
2025 fun select_subtree (tr, []) =
2026 tr (* return the whole tree *)
2027 | select_subtree (Leaf _, _) =
2028 raise REFUTE ("IDT_recursion_interpreter", "interpretation for parameter is a leaf; cannot select a subtree")
2029 | select_subtree (Node tr, x::xs) =
2030 select_subtree (List.nth (tr, x), xs)
2031 (* select the correct subtree of the parameter corresponding to constructor 'c' *)
2032 val p_intr = select_subtree (List.nth (p_intrs, get_coffset idx + c), args)
2033 (* find the indices of the constructor's recursive arguments *)
2034 val (_, _, constrs) = (the o AList.lookup (op =) descr) idx
2035 val constr_args = (snd o List.nth) (constrs, c)
2036 val rec_args = List.filter (DatatypeAux.is_rec_type o fst) (constr_args ~~ args)
2037 val rec_args' = map (fn (dtyp, elem) => (DatatypeAux.dest_DtRec dtyp, elem)) rec_args
2038 (* apply 'p_intr' to recursively computed results *)
2039 val result = foldl (fn ((idx, elem), intr) =>
2040 interpretation_apply (intr, compute_array_entry idx elem)) p_intr rec_args'
2041 (* update 'INTRS' *)
2042 val _ = Array.update ((the o AList.lookup (op =) INTRS) idx, elem, SOME result)
2046 (* compute all entries in INTRS for the current datatype (given by 'index') *)
2047 (* TODO: we can use Array.modify instead once PolyML conforms to the ML standard *)
2048 (* (int * 'a -> 'a) -> 'a array -> unit *)
2051 val size = Array.length arr
2052 fun modifyi_loop i =
2054 Array.update (arr, i, f (i, Array.sub (arr, i)));
2061 val _ = modifyi (fn (i, _) => SOME (compute_array_entry index i)) ((the o AList.lookup (op =) INTRS) index)
2062 (* 'a Array.array -> 'a list *)
2064 Array.foldr op:: [] arr
2066 (* return the part of 'INTRS' that corresponds to the current datatype *)
2067 SOME ((Node o map the o toList o the o AList.lookup (op =) INTRS) index, model', args')
2071 NONE (* not a recursion operator of this datatype *)
2072 ) (DatatypePackage.get_datatypes thy) NONE
2073 | _ => (* head of term is not a constant *)
2076 (* theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option *)
2078 (* only an optimization: 'card' could in principle be interpreted with *)
2079 (* interpreters available already (using its definition), but the code *)
2080 (* below is more efficient *)
2082 fun Finite_Set_card_interpreter thy model args t =
2084 Const ("Finite_Set.card", Type ("fun", [Type ("set", [T]), Type ("nat", [])])) =>
2086 val (i_nat, _, _) = interpret thy model {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True} (Free ("dummy", Type ("nat", [])))
2087 val size_nat = size_of_type i_nat
2088 val (i_set, _, _) = interpret thy model {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True} (Free ("dummy", Type ("set", [T])))
2089 val constants = make_constants i_set
2090 (* interpretation -> int *)
2091 fun number_of_elements (Node xs) =
2092 Library.foldl (fn (n, x) =>
2098 raise REFUTE ("Finite_Set_card_interpreter", "interpretation for set type does not yield a Boolean")) (0, xs)
2099 | number_of_elements (Leaf _) =
2100 raise REFUTE ("Finite_Set_card_interpreter", "interpretation for set type is a leaf")
2101 (* takes an interpretation for a set and returns an interpretation for a 'nat' *)
2102 (* interpretation -> interpretation *)
2105 val n = number_of_elements i
2108 Leaf ((replicate n False) @ True :: (replicate (size_nat-n-1) False))
2110 Leaf (replicate size_nat False)
2113 SOME (Node (map card constants), model, args)
2118 (* theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option *)
2120 (* only an optimization: 'Finites' could in principle be interpreted with *)
2121 (* interpreters available already (using its definition), but the code *)
2122 (* below is more efficient *)
2124 fun Finite_Set_Finites_interpreter thy model args t =
2126 Const ("Finite_Set.Finites", Type ("set", [Type ("set", [T])])) =>
2128 val (i_set, _, _) = interpret thy model {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True} (Free ("dummy", Type ("set", [T])))
2129 val size_set = size_of_type i_set
2131 (* we only consider finite models anyway, hence EVERY set is in "Finites" *)
2132 SOME (Node (replicate size_set TT), model, args)
2137 (* theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option *)
2139 (* only an optimization: 'op <' could in principle be interpreted with *)
2140 (* interpreters available already (using its definition), but the code *)
2141 (* below is more efficient *)
2143 fun Nat_less_interpreter thy model args t =
2145 Const ("Orderings.less", Type ("fun", [Type ("nat", []), Type ("fun", [Type ("nat", []), Type ("bool", [])])])) =>
2147 val (i_nat, _, _) = interpret thy model {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True} (Free ("dummy", Type ("nat", [])))
2148 val size_nat = size_of_type i_nat
2149 (* int -> interpretation *)
2150 (* the 'n'-th nat is not less than the first 'n' nats, while it *)
2151 (* is less than the remaining 'size_nat - n' nats *)
2152 fun less n = Node ((replicate n FF) @ (replicate (size_nat - n) TT))
2154 SOME (Node (map less (1 upto size_nat)), model, args)
2159 (* theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option *)
2161 (* only an optimization: 'HOL.plus' could in principle be interpreted with*)
2162 (* interpreters available already (using its definition), but the code *)
2163 (* below is more efficient *)
2165 fun Nat_plus_interpreter thy model args t =
2167 Const ("HOL.plus", Type ("fun", [Type ("nat", []), Type ("fun", [Type ("nat", []), Type ("nat", [])])])) =>
2169 val (i_nat, _, _) = interpret thy model {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True} (Free ("dummy", Type ("nat", [])))
2170 val size_nat = size_of_type i_nat
2171 (* int -> int -> interpretation *)
2174 val element = (m+n)+1
2176 if element > size_nat then
2177 Leaf (replicate size_nat False)
2179 Leaf ((replicate (element-1) False) @ True :: (replicate (size_nat - element) False))
2182 SOME (Node (map (fn m => Node (map (plus m) (0 upto size_nat-1))) (0 upto size_nat-1)), model, args)
2187 (* theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option *)
2189 (* only an optimization: 'op -' could in principle be interpreted with *)
2190 (* interpreters available already (using its definition), but the code *)
2191 (* below is more efficient *)
2193 fun Nat_minus_interpreter thy model args t =
2195 Const ("HOL.minus", Type ("fun", [Type ("nat", []), Type ("fun", [Type ("nat", []), Type ("nat", [])])])) =>
2197 val (i_nat, _, _) = interpret thy model {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True} (Free ("dummy", Type ("nat", [])))
2198 val size_nat = size_of_type i_nat
2199 (* int -> int -> interpretation *)
2202 val element = Int.max (m-n, 0) + 1
2204 Leaf ((replicate (element-1) False) @ True :: (replicate (size_nat - element) False))
2207 SOME (Node (map (fn m => Node (map (minus m) (0 upto size_nat-1))) (0 upto size_nat-1)), model, args)
2212 (* theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option *)
2214 (* only an optimization: 'HOL.times' could in principle be interpreted with *)
2215 (* interpreters available already (using its definition), but the code *)
2216 (* below is more efficient *)
2218 fun Nat_mult_interpreter thy model args t =
2220 Const ("HOL.times", Type ("fun", [Type ("nat", []), Type ("fun", [Type ("nat", []), Type ("nat", [])])])) =>
2222 val (i_nat, _, _) = interpret thy model {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True} (Free ("dummy", Type ("nat", [])))
2223 val size_nat = size_of_type i_nat
2224 (* nat -> nat -> interpretation *)
2227 val element = (m*n)+1
2229 if element > size_nat then
2230 Leaf (replicate size_nat False)
2232 Leaf ((replicate (element-1) False) @ True :: (replicate (size_nat - element) False))
2235 SOME (Node (map (fn m => Node (map (mult m) (0 upto size_nat-1))) (0 upto size_nat-1)), model, args)
2240 (* theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option *)
2242 (* only an optimization: 'op @' could in principle be interpreted with *)
2243 (* interpreters available already (using its definition), but the code *)
2244 (* below is more efficient *)
2246 fun List_append_interpreter thy model args t =
2248 Const ("List.op @", Type ("fun", [Type ("List.list", [T]), Type ("fun", [Type ("List.list", [_]), Type ("List.list", [_])])])) =>
2250 val (i_elem, _, _) = interpret thy model {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True} (Free ("dummy", T))
2251 val size_elem = size_of_type i_elem
2252 val (i_list, _, _) = interpret thy model {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True} (Free ("dummy", Type ("List.list", [T])))
2253 val size_list = size_of_type i_list
2254 (* power (a, b) computes a^b, for a>=0, b>=0 *)
2255 (* int * int -> int *)
2256 fun power (a, 0) = 1
2258 | power (a, b) = let val ab = power(a, b div 2) in ab * ab * power(a, b mod 2) end
2259 (* log (a, b) computes floor(log_a(b)), i.e. the largest integer x s.t. a^x <= b, for a>=2, b>=1 *)
2260 (* int * int -> int *)
2263 fun logloop (ax, x) =
2264 if ax > b then x-1 else logloop (a * ax, x+1)
2268 (* nat -> nat -> interpretation *)
2271 (* The following formula depends on the order in which lists are *)
2272 (* enumerated by the 'IDT_constructor_interpreter'. It took me *)
2273 (* a while to come up with this formula. *)
2274 val element = n + m * (if size_elem = 1 then 1 else power (size_elem, log (size_elem, n+1))) + 1
2276 if element > size_list then
2277 Leaf (replicate size_list False)
2279 Leaf ((replicate (element-1) False) @ True :: (replicate (size_list - element) False))
2282 SOME (Node (map (fn m => Node (map (append m) (0 upto size_list-1))) (0 upto size_list-1)), model, args)
2287 (* theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option *)
2289 (* only an optimization: 'lfp' could in principle be interpreted with *)
2290 (* interpreters available already (using its definition), but the code *)
2291 (* below is more efficient *)
2293 fun Lfp_lfp_interpreter thy model args t =
2295 Const ("Lfp.lfp", Type ("fun", [Type ("fun", [Type ("set", [T]), Type ("set", [_])]), Type ("set", [_])])) =>
2297 val (i_elem, _, _) = interpret thy model {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True} (Free ("dummy", T))
2298 val size_elem = size_of_type i_elem
2299 (* the universe (i.e. the set that contains every element) *)
2300 val i_univ = Node (replicate size_elem TT)
2301 (* all sets with elements from type 'T' *)
2302 val (i_set, _, _) = interpret thy model {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True} (Free ("dummy", Type ("set", [T])))
2303 val i_sets = make_constants i_set
2304 (* all functions that map sets to sets *)
2305 val (i_fun, _, _) = interpret thy model {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True} (Free ("dummy", Type ("fun", [Type ("set", [T]), Type ("set", [T])])))
2306 val i_funs = make_constants i_fun
2307 (* "lfp(f) == Inter({u. f(u) <= u})" *)
2308 (* interpretation * interpretation -> bool *)
2309 fun is_subset (Node subs, Node sups) =
2310 List.all (fn (sub, sup) => (sub = FF) orelse (sup = TT)) (subs ~~ sups)
2311 | is_subset (_, _) =
2312 raise REFUTE ("Lfp_lfp_interpreter", "is_subset: interpretation for set is not a node")
2313 (* interpretation * interpretation -> interpretation *)
2314 fun intersection (Node xs, Node ys) =
2315 Node (map (fn (x, y) => if (x = TT) andalso (y = TT) then TT else FF) (xs ~~ ys))
2316 | intersection (_, _) =
2317 raise REFUTE ("Lfp_lfp_interpreter", "intersection: interpretation for set is not a node")
2318 (* interpretation -> interpretaion *)
2319 fun lfp (Node resultsets) =
2320 foldl (fn ((set, resultset), acc) =>
2321 if is_subset (resultset, set) then
2322 intersection (acc, set)
2324 acc) i_univ (i_sets ~~ resultsets)
2326 raise REFUTE ("Lfp_lfp_interpreter", "lfp: interpretation for function is not a node")
2328 SOME (Node (map lfp i_funs), model, args)
2333 (* theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option *)
2335 (* only an optimization: 'gfp' could in principle be interpreted with *)
2336 (* interpreters available already (using its definition), but the code *)
2337 (* below is more efficient *)
2339 fun Gfp_gfp_interpreter thy model args t =
2341 Const ("Gfp.gfp", Type ("fun", [Type ("fun", [Type ("set", [T]), Type ("set", [_])]), Type ("set", [_])])) =>
2342 let nonfix union (*because "union" is used below*)
2343 val (i_elem, _, _) = interpret thy model {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True} (Free ("dummy", T))
2344 val size_elem = size_of_type i_elem
2345 (* the universe (i.e. the set that contains every element) *)
2346 val i_univ = Node (replicate size_elem TT)
2347 (* all sets with elements from type 'T' *)
2348 val (i_set, _, _) = interpret thy model {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True} (Free ("dummy", Type ("set", [T])))
2349 val i_sets = make_constants i_set
2350 (* all functions that map sets to sets *)
2351 val (i_fun, _, _) = interpret thy model {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True} (Free ("dummy", Type ("fun", [Type ("set", [T]), Type ("set", [T])])))
2352 val i_funs = make_constants i_fun
2353 (* "gfp(f) == Union({u. u <= f(u)})" *)
2354 (* interpretation * interpretation -> bool *)
2355 fun is_subset (Node subs, Node sups) =
2356 List.all (fn (sub, sup) => (sub = FF) orelse (sup = TT)) (subs ~~ sups)
2357 | is_subset (_, _) =
2358 raise REFUTE ("Gfp_gfp_interpreter", "is_subset: interpretation for set is not a node")
2359 (* interpretation * interpretation -> interpretation *)
2360 fun union (Node xs, Node ys) =
2361 Node (map (fn (x,y) => if x=TT orelse y=TT then TT else FF)
2364 raise REFUTE ("Gfp_gfp_interpreter", "union: interpretation for set is not a node")
2365 (* interpretation -> interpretaion *)
2366 fun gfp (Node resultsets) =
2367 foldl (fn ((set, resultset), acc) =>
2368 if is_subset (set, resultset) then union (acc, set)
2370 i_univ (i_sets ~~ resultsets)
2372 raise REFUTE ("Gfp_gfp_interpreter", "gfp: interpretation for function is not a node")
2374 SOME (Node (map gfp i_funs), model, args)
2380 (* ------------------------------------------------------------------------- *)
2382 (* ------------------------------------------------------------------------- *)
2384 (* theory -> model -> Term.term -> interpretation -> (int -> bool) -> Term.term option *)
2386 fun stlc_printer thy model t intr assignment =
2388 (* Term.term -> Term.typ option *)
2389 fun typeof (Free (_, T)) = SOME T
2390 | typeof (Var (_, T)) = SOME T
2391 | typeof (Const (_, T)) = SOME T
2393 (* string -> string *)
2394 fun strip_leading_quote s =
2395 (implode o (fn ss => case ss of [] => [] | x::xs => if x="'" then xs else ss) o explode) s
2396 (* Term.typ -> string *)
2397 fun string_of_typ (Type (s, _)) = s
2398 | string_of_typ (TFree (x, _)) = strip_leading_quote x
2399 | string_of_typ (TVar ((x,i), _)) = strip_leading_quote x ^ string_of_int i
2400 (* interpretation -> int *)
2401 fun index_from_interpretation (Leaf xs) =
2402 find_index (PropLogic.eval assignment) xs
2403 | index_from_interpretation _ =
2404 raise REFUTE ("stlc_printer", "interpretation for ground type is not a leaf")
2409 Type ("fun", [T1, T2]) =>
2411 (* create all constants of type 'T1' *)
2412 val (i, _, _) = interpret thy model {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True} (Free ("dummy", T1))
2413 val constants = make_constants i
2414 (* interpretation list *)
2415 val results = (case intr of
2417 | _ => raise REFUTE ("stlc_printer", "interpretation for function type is a leaf"))
2418 (* Term.term list *)
2419 val pairs = map (fn (arg, result) =>
2421 (print thy model (Free ("dummy", T1)) arg assignment,
2422 print thy model (Free ("dummy", T2)) result assignment))
2423 (constants ~~ results)
2425 val HOLogic_prodT = HOLogic.mk_prodT (T1, T2)
2426 val HOLogic_setT = HOLogic.mk_setT HOLogic_prodT
2428 val HOLogic_empty_set = Const ("{}", HOLogic_setT)
2429 val HOLogic_insert = Const ("insert", HOLogic_prodT --> HOLogic_setT --> HOLogic_setT)
2431 SOME (foldr (fn (pair, acc) => HOLogic_insert $ pair $ acc) HOLogic_empty_set pairs)
2433 | Type ("prop", []) =>
2434 (case index_from_interpretation intr of
2435 (~1) => SOME (HOLogic.mk_Trueprop (Const ("arbitrary", HOLogic.boolT)))
2436 | 0 => SOME (HOLogic.mk_Trueprop HOLogic.true_const)
2437 | 1 => SOME (HOLogic.mk_Trueprop HOLogic.false_const)
2438 | _ => raise REFUTE ("stlc_interpreter", "illegal interpretation for a propositional value"))
2439 | Type _ => if index_from_interpretation intr = (~1) then
2440 SOME (Const ("arbitrary", T))
2442 SOME (Const (string_of_typ T ^ string_of_int (index_from_interpretation intr), T))
2443 | TFree _ => if index_from_interpretation intr = (~1) then
2444 SOME (Const ("arbitrary", T))
2446 SOME (Const (string_of_typ T ^ string_of_int (index_from_interpretation intr), T))
2447 | TVar _ => if index_from_interpretation intr = (~1) then
2448 SOME (Const ("arbitrary", T))
2450 SOME (Const (string_of_typ T ^ string_of_int (index_from_interpretation intr), T)))
2455 (* theory -> model -> Term.term -> interpretation -> (int -> bool) -> string option *)
2457 fun set_printer thy model t intr assignment =
2459 (* Term.term -> Term.typ option *)
2460 fun typeof (Free (_, T)) = SOME T
2461 | typeof (Var (_, T)) = SOME T
2462 | typeof (Const (_, T)) = SOME T
2466 SOME (Type ("set", [T])) =>
2468 (* create all constants of type 'T' *)
2469 val (i, _, _) = interpret thy model {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True} (Free ("dummy", T))
2470 val constants = make_constants i
2471 (* interpretation list *)
2472 val results = (case intr of
2474 | _ => raise REFUTE ("set_printer", "interpretation for set type is a leaf"))
2475 (* Term.term list *)
2476 val elements = List.mapPartial (fn (arg, result) =>
2478 Leaf [fmTrue, fmFalse] =>
2479 if PropLogic.eval assignment fmTrue then
2480 SOME (print thy model (Free ("dummy", T)) arg assignment)
2481 else (*if PropLogic.eval assignment fmFalse then*)
2484 raise REFUTE ("set_printer", "illegal interpretation for a Boolean value"))
2485 (constants ~~ results)
2487 val HOLogic_setT = HOLogic.mk_setT T
2489 val HOLogic_empty_set = Const ("{}", HOLogic_setT)
2490 val HOLogic_insert = Const ("insert", T --> HOLogic_setT --> HOLogic_setT)
2492 SOME (Library.foldl (fn (acc, elem) => HOLogic_insert $ elem $ acc) (HOLogic_empty_set, elements))
2498 (* theory -> model -> Term.term -> interpretation -> (int -> bool) -> Term.term option *)
2500 fun IDT_printer thy model t intr assignment =
2502 (* Term.term -> Term.typ option *)
2503 fun typeof (Free (_, T)) = SOME T
2504 | typeof (Var (_, T)) = SOME T
2505 | typeof (Const (_, T)) = SOME T
2509 SOME (Type (s, Ts)) =>
2510 (case DatatypePackage.get_datatype thy s of
2511 SOME info => (* inductive datatype *)
2513 val (typs, _) = model
2514 val index = #index info
2515 val descr = #descr info
2516 val (_, dtyps, constrs) = (the o AList.lookup (op =) descr) index
2517 val typ_assoc = dtyps ~~ Ts
2518 (* sanity check: every element in 'dtyps' must be a 'DtTFree' *)
2519 val _ = (if Library.exists (fn d =>
2520 case d of DatatypeAux.DtTFree _ => false | _ => true) dtyps
2522 raise REFUTE ("IDT_printer", "datatype argument (for type " ^ Sign.string_of_typ (sign_of thy) (Type (s, Ts)) ^ ") is not a variable")
2525 (* the index of the element in the datatype *)
2526 val element = (case intr of
2527 Leaf xs => find_index (PropLogic.eval assignment) xs
2528 | Node _ => raise REFUTE ("IDT_printer", "interpretation is not a leaf"))
2531 SOME (Const ("arbitrary", Type (s, Ts)))
2533 (* takes a datatype constructor, and if for some arguments this constructor *)
2534 (* generates the datatype's element that is given by 'element', returns the *)
2535 (* constructor (as a term) as well as the indices of the arguments *)
2536 (* string * DatatypeAux.dtyp list -> (Term.term * int list) option *)
2537 fun get_constr_args (cname, cargs) =
2539 val cTerm = Const (cname, (map (typ_of_dtyp descr typ_assoc) cargs) ---> Type (s, Ts))
2540 val (iC, _, _) = interpret thy (typs, []) {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True} cTerm
2541 (* interpretation -> int list option *)
2542 fun get_args (Leaf xs) =
2543 if find_index_eq True xs = element then
2547 | get_args (Node xs) =
2549 (* interpretation * int -> int list option *)
2550 fun search ([], _) =
2552 | search (x::xs, n) =
2554 SOME result => SOME (n::result)
2555 | NONE => search (xs, n+1))
2560 Option.map (fn args => (cTerm, cargs, args)) (get_args iC)
2562 (* Term.term * DatatypeAux.dtyp list * int list *)
2563 val (cTerm, cargs, args) = (case get_first get_constr_args constrs of
2565 | NONE => raise REFUTE ("IDT_printer", "no matching constructor found for element " ^ string_of_int element))
2566 val argsTerms = map (fn (d, n) =>
2568 val dT = typ_of_dtyp descr typ_assoc d
2569 val (i, _, _) = interpret thy (typs, []) {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True} (Free ("dummy", dT))
2570 val consts = make_constants i (* we only need the n-th element of this *)
2571 (* list, so there might be a more efficient implementation that does *)
2572 (* not generate all constants *)
2574 print thy (typs, []) (Free ("dummy", dT)) (List.nth (consts, n)) assignment
2575 end) (cargs ~~ args)
2577 SOME (Library.foldl op$ (cTerm, argsTerms))
2580 | NONE => (* not an inductive datatype *)
2582 | _ => (* a (free or schematic) type variable *)
2587 (* ------------------------------------------------------------------------- *)
2588 (* use 'setup Refute.setup' in an Isabelle theory to initialize the 'Refute' *)
2590 (* ------------------------------------------------------------------------- *)
2592 (* ------------------------------------------------------------------------- *)
2593 (* Note: the interpreters and printers are used in reverse order; however, *)
2594 (* an interpreter that can handle non-atomic terms ends up being *)
2595 (* applied before the 'stlc_interpreter' breaks the term apart into *)
2596 (* subterms that are then passed to other interpreters! *)
2597 (* ------------------------------------------------------------------------- *)
2599 (* (theory -> theory) list *)
2603 add_interpreter "stlc" stlc_interpreter #>
2604 add_interpreter "Pure" Pure_interpreter #>
2605 add_interpreter "HOLogic" HOLogic_interpreter #>
2606 add_interpreter "set" set_interpreter #>
2607 add_interpreter "IDT" IDT_interpreter #>
2608 add_interpreter "IDT_constructor" IDT_constructor_interpreter #>
2609 add_interpreter "IDT_recursion" IDT_recursion_interpreter #>
2610 add_interpreter "Finite_Set.card" Finite_Set_card_interpreter #>
2611 add_interpreter "Finite_Set.Finites" Finite_Set_Finites_interpreter #>
2612 add_interpreter "Nat.op <" Nat_less_interpreter #>
2613 add_interpreter "Nat.op +" Nat_plus_interpreter #>
2614 add_interpreter "Nat.op -" Nat_minus_interpreter #>
2615 add_interpreter "Nat.op *" Nat_mult_interpreter #>
2616 add_interpreter "List.op @" List_append_interpreter #>
2617 add_interpreter "Lfp.lfp" Lfp_lfp_interpreter #>
2618 add_interpreter "Gfp.gfp" Gfp_gfp_interpreter #>
2619 add_printer "stlc" stlc_printer #>
2620 add_printer "set" set_printer #>
2621 add_printer "IDT" IDT_printer;