webertj@14350
|
1 |
(* Title: HOL/Tools/refute.ML
|
webertj@14350
|
2 |
ID: $Id$
|
webertj@14350
|
3 |
Author: Tjark Weber
|
webertj@15547
|
4 |
Copyright 2003-2005
|
webertj@14350
|
5 |
|
webertj@14965
|
6 |
Finite model generation for HOL formulas, using a SAT solver.
|
webertj@14350
|
7 |
*)
|
webertj@14350
|
8 |
|
webertj@14456
|
9 |
(* ------------------------------------------------------------------------- *)
|
webertj@14456
|
10 |
(* Declares the 'REFUTE' signature as well as a structure 'Refute'. *)
|
webertj@14456
|
11 |
(* Documentation is available in the Isabelle/Isar theory 'HOL/Refute.thy'. *)
|
webertj@14350
|
12 |
(* ------------------------------------------------------------------------- *)
|
webertj@14350
|
13 |
|
webertj@14350
|
14 |
signature REFUTE =
|
webertj@14350
|
15 |
sig
|
webertj@14350
|
16 |
|
webertj@14456
|
17 |
exception REFUTE of string * string
|
webertj@14456
|
18 |
|
webertj@14456
|
19 |
(* ------------------------------------------------------------------------- *)
|
webertj@14807
|
20 |
(* Model/interpretation related code (translation HOL -> propositional logic *)
|
webertj@14456
|
21 |
(* ------------------------------------------------------------------------- *)
|
webertj@14456
|
22 |
|
webertj@14807
|
23 |
type params
|
webertj@14456
|
24 |
type interpretation
|
webertj@14456
|
25 |
type model
|
webertj@14456
|
26 |
type arguments
|
webertj@14456
|
27 |
|
webertj@14807
|
28 |
exception MAXVARS_EXCEEDED
|
webertj@14456
|
29 |
|
webertj@14807
|
30 |
val add_interpreter : string -> (theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option) -> theory -> theory
|
webertj@14807
|
31 |
val add_printer : string -> (theory -> model -> Term.term -> interpretation -> (int -> bool) -> Term.term option) -> theory -> theory
|
webertj@14456
|
32 |
|
webertj@15334
|
33 |
val interpret : theory -> model -> arguments -> Term.term -> (interpretation * model * arguments)
|
webertj@14807
|
34 |
|
webertj@14807
|
35 |
val print : theory -> model -> Term.term -> interpretation -> (int -> bool) -> Term.term
|
webertj@14456
|
36 |
val print_model : theory -> model -> (int -> bool) -> string
|
webertj@14456
|
37 |
|
webertj@14456
|
38 |
(* ------------------------------------------------------------------------- *)
|
webertj@14456
|
39 |
(* Interface *)
|
webertj@14456
|
40 |
(* ------------------------------------------------------------------------- *)
|
webertj@14456
|
41 |
|
webertj@14456
|
42 |
val set_default_param : (string * string) -> theory -> theory
|
webertj@14456
|
43 |
val get_default_param : theory -> string -> string option
|
webertj@14456
|
44 |
val get_default_params : theory -> (string * string) list
|
webertj@14807
|
45 |
val actual_params : theory -> (string * string) list -> params
|
webertj@14456
|
46 |
|
webertj@14807
|
47 |
val find_model : theory -> params -> Term.term -> bool -> unit
|
webertj@14456
|
48 |
|
webertj@14456
|
49 |
val satisfy_term : theory -> (string * string) list -> Term.term -> unit (* tries to find a model for a formula *)
|
webertj@14456
|
50 |
val refute_term : theory -> (string * string) list -> Term.term -> unit (* tries to find a model that refutes a formula *)
|
webertj@14456
|
51 |
val refute_subgoal : theory -> (string * string) list -> Thm.thm -> int -> unit
|
webertj@14456
|
52 |
|
wenzelm@18708
|
53 |
val setup : theory -> theory
|
webertj@14456
|
54 |
end;
|
webertj@14456
|
55 |
|
webertj@14456
|
56 |
structure Refute : REFUTE =
|
webertj@14456
|
57 |
struct
|
webertj@14456
|
58 |
|
webertj@14456
|
59 |
open PropLogic;
|
webertj@14456
|
60 |
|
webertj@14350
|
61 |
(* We use 'REFUTE' only for internal error conditions that should *)
|
webertj@14350
|
62 |
(* never occur in the first place (i.e. errors caused by bugs in our *)
|
webertj@14350
|
63 |
(* code). Otherwise (e.g. to indicate invalid input data) we use *)
|
webertj@14350
|
64 |
(* 'error'. *)
|
webertj@14456
|
65 |
exception REFUTE of string * string; (* ("in function", "cause") *)
|
webertj@14350
|
66 |
|
webertj@14807
|
67 |
(* should be raised by an interpreter when more variables would be *)
|
webertj@14807
|
68 |
(* required than allowed by 'maxvars' *)
|
webertj@14807
|
69 |
exception MAXVARS_EXCEEDED;
|
webertj@14350
|
70 |
|
webertj@14350
|
71 |
(* ------------------------------------------------------------------------- *)
|
webertj@14350
|
72 |
(* TREES *)
|
webertj@14350
|
73 |
(* ------------------------------------------------------------------------- *)
|
webertj@14350
|
74 |
|
webertj@14350
|
75 |
(* ------------------------------------------------------------------------- *)
|
webertj@14350
|
76 |
(* tree: implements an arbitrarily (but finitely) branching tree as a list *)
|
webertj@14350
|
77 |
(* of (lists of ...) elements *)
|
webertj@14350
|
78 |
(* ------------------------------------------------------------------------- *)
|
webertj@14350
|
79 |
|
webertj@14350
|
80 |
datatype 'a tree =
|
webertj@14350
|
81 |
Leaf of 'a
|
webertj@14350
|
82 |
| Node of ('a tree) list;
|
webertj@14350
|
83 |
|
webertj@14350
|
84 |
(* ('a -> 'b) -> 'a tree -> 'b tree *)
|
webertj@14350
|
85 |
|
webertj@14350
|
86 |
fun tree_map f tr =
|
webertj@14350
|
87 |
case tr of
|
webertj@14350
|
88 |
Leaf x => Leaf (f x)
|
webertj@14350
|
89 |
| Node xs => Node (map (tree_map f) xs);
|
webertj@14350
|
90 |
|
webertj@14350
|
91 |
(* ('a * 'b -> 'a) -> 'a * ('b tree) -> 'a *)
|
webertj@14350
|
92 |
|
webertj@14350
|
93 |
fun tree_foldl f =
|
webertj@14350
|
94 |
let
|
webertj@14350
|
95 |
fun itl (e, Leaf x) = f(e,x)
|
skalberg@15570
|
96 |
| itl (e, Node xs) = Library.foldl (tree_foldl f) (e,xs)
|
webertj@14350
|
97 |
in
|
webertj@14350
|
98 |
itl
|
webertj@14350
|
99 |
end;
|
webertj@14350
|
100 |
|
webertj@14350
|
101 |
(* 'a tree * 'b tree -> ('a * 'b) tree *)
|
webertj@14350
|
102 |
|
webertj@15611
|
103 |
fun tree_pair (t1, t2) =
|
webertj@14350
|
104 |
case t1 of
|
webertj@14350
|
105 |
Leaf x =>
|
webertj@14350
|
106 |
(case t2 of
|
webertj@14350
|
107 |
Leaf y => Leaf (x,y)
|
webertj@14350
|
108 |
| Node _ => raise REFUTE ("tree_pair", "trees are of different height (second tree is higher)"))
|
webertj@14350
|
109 |
| Node xs =>
|
webertj@14350
|
110 |
(case t2 of
|
webertj@14807
|
111 |
(* '~~' will raise an exception if the number of branches in *)
|
webertj@14807
|
112 |
(* both trees is different at the current node *)
|
webertj@14350
|
113 |
Node ys => Node (map tree_pair (xs ~~ ys))
|
webertj@14350
|
114 |
| Leaf _ => raise REFUTE ("tree_pair", "trees are of different height (first tree is higher)"));
|
webertj@14350
|
115 |
|
webertj@14350
|
116 |
(* ------------------------------------------------------------------------- *)
|
webertj@14807
|
117 |
(* params: parameters that control the translation into a propositional *)
|
webertj@14807
|
118 |
(* formula/model generation *)
|
webertj@14807
|
119 |
(* *)
|
webertj@14807
|
120 |
(* The following parameters are supported (and required (!), except for *)
|
webertj@14807
|
121 |
(* "sizes"): *)
|
webertj@14807
|
122 |
(* *)
|
webertj@14807
|
123 |
(* Name Type Description *)
|
webertj@14807
|
124 |
(* *)
|
webertj@14807
|
125 |
(* "sizes" (string * int) list *)
|
webertj@14807
|
126 |
(* Size of ground types (e.g. 'a=2), or depth of IDTs. *)
|
webertj@14807
|
127 |
(* "minsize" int If >0, minimal size of each ground type/IDT depth. *)
|
webertj@14807
|
128 |
(* "maxsize" int If >0, maximal size of each ground type/IDT depth. *)
|
webertj@14807
|
129 |
(* "maxvars" int If >0, use at most 'maxvars' Boolean variables *)
|
webertj@14807
|
130 |
(* when transforming the term into a propositional *)
|
webertj@14807
|
131 |
(* formula. *)
|
webertj@14807
|
132 |
(* "maxtime" int If >0, terminate after at most 'maxtime' seconds. *)
|
webertj@14807
|
133 |
(* "satsolver" string SAT solver to be used. *)
|
webertj@14807
|
134 |
(* ------------------------------------------------------------------------- *)
|
webertj@14807
|
135 |
|
webertj@14807
|
136 |
type params =
|
webertj@14807
|
137 |
{
|
webertj@14807
|
138 |
sizes : (string * int) list,
|
webertj@14807
|
139 |
minsize : int,
|
webertj@14807
|
140 |
maxsize : int,
|
webertj@14807
|
141 |
maxvars : int,
|
webertj@14807
|
142 |
maxtime : int,
|
webertj@14807
|
143 |
satsolver: string
|
webertj@14807
|
144 |
};
|
webertj@14807
|
145 |
|
webertj@14807
|
146 |
(* ------------------------------------------------------------------------- *)
|
webertj@14456
|
147 |
(* interpretation: a term's interpretation is given by a variable of type *)
|
webertj@14456
|
148 |
(* 'interpretation' *)
|
webertj@14350
|
149 |
(* ------------------------------------------------------------------------- *)
|
webertj@14350
|
150 |
|
webertj@14456
|
151 |
type interpretation =
|
webertj@14456
|
152 |
prop_formula list tree;
|
webertj@14350
|
153 |
|
webertj@14350
|
154 |
(* ------------------------------------------------------------------------- *)
|
webertj@14456
|
155 |
(* model: a model specifies the size of types and the interpretation of *)
|
webertj@14456
|
156 |
(* terms *)
|
webertj@14350
|
157 |
(* ------------------------------------------------------------------------- *)
|
webertj@14350
|
158 |
|
webertj@14456
|
159 |
type model =
|
webertj@14456
|
160 |
(Term.typ * int) list * (Term.term * interpretation) list;
|
webertj@14350
|
161 |
|
webertj@14456
|
162 |
(* ------------------------------------------------------------------------- *)
|
webertj@14456
|
163 |
(* arguments: additional arguments required during interpretation of terms *)
|
webertj@14456
|
164 |
(* ------------------------------------------------------------------------- *)
|
webertj@14807
|
165 |
|
webertj@14456
|
166 |
type arguments =
|
webertj@14807
|
167 |
{
|
webertj@15547
|
168 |
maxvars : int, (* just passed unchanged from 'params' *)
|
webertj@15547
|
169 |
def_eq : bool, (* whether to use 'make_equality' or 'make_def_equality' *)
|
webertj@15547
|
170 |
(* the following may change during the translation *)
|
webertj@14807
|
171 |
next_idx : int,
|
webertj@14807
|
172 |
bounds : interpretation list,
|
webertj@14807
|
173 |
wellformed: prop_formula
|
webertj@14807
|
174 |
};
|
webertj@14456
|
175 |
|
webertj@14350
|
176 |
|
webertj@14456
|
177 |
structure RefuteDataArgs =
|
webertj@14456
|
178 |
struct
|
webertj@14456
|
179 |
val name = "HOL/refute";
|
webertj@14456
|
180 |
type T =
|
webertj@14456
|
181 |
{interpreters: (string * (theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option)) list,
|
webertj@14807
|
182 |
printers: (string * (theory -> model -> Term.term -> interpretation -> (int -> bool) -> Term.term option)) list,
|
webertj@14456
|
183 |
parameters: string Symtab.table};
|
webertj@14456
|
184 |
val empty = {interpreters = [], printers = [], parameters = Symtab.empty};
|
webertj@14456
|
185 |
val copy = I;
|
wenzelm@16424
|
186 |
val extend = I;
|
wenzelm@16424
|
187 |
fun merge _
|
webertj@14456
|
188 |
({interpreters = in1, printers = pr1, parameters = pa1},
|
webertj@14456
|
189 |
{interpreters = in2, printers = pr2, parameters = pa2}) =
|
webertj@14456
|
190 |
{interpreters = rev (merge_alists (rev in1) (rev in2)),
|
webertj@14456
|
191 |
printers = rev (merge_alists (rev pr1) (rev pr2)),
|
webertj@14456
|
192 |
parameters = Symtab.merge (op=) (pa1, pa2)};
|
webertj@14456
|
193 |
fun print sg {interpreters, printers, parameters} =
|
webertj@14456
|
194 |
Pretty.writeln (Pretty.chunks
|
webertj@15767
|
195 |
[Pretty.strs ("default parameters:" :: List.concat (map (fn (name, value) => [name, "=", value]) (Symtab.dest parameters))),
|
webertj@14456
|
196 |
Pretty.strs ("interpreters:" :: map fst interpreters),
|
webertj@14456
|
197 |
Pretty.strs ("printers:" :: map fst printers)]);
|
webertj@14456
|
198 |
end;
|
webertj@14456
|
199 |
|
webertj@14456
|
200 |
structure RefuteData = TheoryDataFun(RefuteDataArgs);
|
webertj@14456
|
201 |
|
webertj@14350
|
202 |
|
webertj@14350
|
203 |
(* ------------------------------------------------------------------------- *)
|
webertj@15334
|
204 |
(* interpret: interprets the term 't' using a suitable interpreter; returns *)
|
webertj@15334
|
205 |
(* the interpretation and a (possibly extended) model that keeps *)
|
webertj@15334
|
206 |
(* track of the interpretation of subterms *)
|
webertj@14350
|
207 |
(* ------------------------------------------------------------------------- *)
|
webertj@14350
|
208 |
|
webertj@14807
|
209 |
(* theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) *)
|
webertj@14456
|
210 |
|
webertj@14456
|
211 |
fun interpret thy model args t =
|
webertj@14456
|
212 |
(case get_first (fn (_, f) => f thy model args t) (#interpreters (RefuteData.get thy)) of
|
webertj@15547
|
213 |
NONE => raise REFUTE ("interpret", "no interpreter for term " ^ quote (Sign.string_of_term (sign_of thy) t))
|
skalberg@15531
|
214 |
| SOME x => x);
|
webertj@14456
|
215 |
|
webertj@14456
|
216 |
(* ------------------------------------------------------------------------- *)
|
webertj@15547
|
217 |
(* print: converts the constant denoted by the term 't' into a term using a *)
|
webertj@15547
|
218 |
(* suitable printer *)
|
webertj@14456
|
219 |
(* ------------------------------------------------------------------------- *)
|
webertj@14350
|
220 |
|
webertj@14807
|
221 |
(* theory -> model -> Term.term -> interpretation -> (int -> bool) -> Term.term *)
|
webertj@14456
|
222 |
|
webertj@14456
|
223 |
fun print thy model t intr assignment =
|
webertj@14456
|
224 |
(case get_first (fn (_, f) => f thy model t intr assignment) (#printers (RefuteData.get thy)) of
|
webertj@15547
|
225 |
NONE => raise REFUTE ("print", "no printer for term " ^ quote (Sign.string_of_term (sign_of thy) t))
|
skalberg@15531
|
226 |
| SOME x => x);
|
webertj@14456
|
227 |
|
webertj@14456
|
228 |
(* ------------------------------------------------------------------------- *)
|
webertj@14456
|
229 |
(* print_model: turns the model into a string, using a fixed interpretation *)
|
webertj@14807
|
230 |
(* (given by an assignment for Boolean variables) and suitable *)
|
webertj@14456
|
231 |
(* printers *)
|
webertj@14456
|
232 |
(* ------------------------------------------------------------------------- *)
|
webertj@14456
|
233 |
|
webertj@14807
|
234 |
(* theory -> model -> (int -> bool) -> string *)
|
webertj@14807
|
235 |
|
webertj@14456
|
236 |
fun print_model thy model assignment =
|
webertj@14456
|
237 |
let
|
webertj@14456
|
238 |
val (typs, terms) = model
|
webertj@14807
|
239 |
val typs_msg =
|
webertj@14807
|
240 |
if null typs then
|
webertj@14807
|
241 |
"empty universe (no type variables in term)\n"
|
webertj@14807
|
242 |
else
|
webertj@15611
|
243 |
"Size of types: " ^ commas (map (fn (T, i) => Sign.string_of_typ (sign_of thy) T ^ ": " ^ string_of_int i) typs) ^ "\n"
|
webertj@14807
|
244 |
val show_consts_msg =
|
webertj@14807
|
245 |
if not (!show_consts) andalso Library.exists (is_Const o fst) terms then
|
webertj@14807
|
246 |
"set \"show_consts\" to show the interpretation of constants\n"
|
webertj@14807
|
247 |
else
|
webertj@14807
|
248 |
""
|
webertj@14807
|
249 |
val terms_msg =
|
webertj@14807
|
250 |
if null terms then
|
webertj@14807
|
251 |
"empty interpretation (no free variables in term)\n"
|
webertj@14807
|
252 |
else
|
webertj@15611
|
253 |
space_implode "\n" (List.mapPartial (fn (t, intr) =>
|
webertj@14807
|
254 |
(* print constants only if 'show_consts' is true *)
|
webertj@14807
|
255 |
if (!show_consts) orelse not (is_Const t) then
|
skalberg@15531
|
256 |
SOME (Sign.string_of_term (sign_of thy) t ^ ": " ^ Sign.string_of_term (sign_of thy) (print thy model t intr assignment))
|
webertj@14807
|
257 |
else
|
skalberg@15531
|
258 |
NONE) terms) ^ "\n"
|
webertj@14456
|
259 |
in
|
webertj@14807
|
260 |
typs_msg ^ show_consts_msg ^ terms_msg
|
webertj@14456
|
261 |
end;
|
webertj@14456
|
262 |
|
webertj@14456
|
263 |
|
webertj@14456
|
264 |
(* ------------------------------------------------------------------------- *)
|
webertj@14456
|
265 |
(* PARAMETER MANAGEMENT *)
|
webertj@14456
|
266 |
(* ------------------------------------------------------------------------- *)
|
webertj@14456
|
267 |
|
webertj@14456
|
268 |
(* string -> (theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option) -> theory -> theory *)
|
webertj@14456
|
269 |
|
webertj@14456
|
270 |
fun add_interpreter name f thy =
|
webertj@14456
|
271 |
let
|
webertj@14456
|
272 |
val {interpreters, printers, parameters} = RefuteData.get thy
|
webertj@14456
|
273 |
in
|
haftmann@17314
|
274 |
case AList.lookup (op =) interpreters name of
|
skalberg@15531
|
275 |
NONE => RefuteData.put {interpreters = (name, f) :: interpreters, printers = printers, parameters = parameters} thy
|
skalberg@15531
|
276 |
| SOME _ => error ("Interpreter " ^ name ^ " already declared")
|
webertj@14456
|
277 |
end;
|
webertj@14456
|
278 |
|
webertj@14807
|
279 |
(* string -> (theory -> model -> Term.term -> interpretation -> (int -> bool) -> Term.term option) -> theory -> theory *)
|
webertj@14456
|
280 |
|
webertj@14456
|
281 |
fun add_printer name f thy =
|
webertj@14456
|
282 |
let
|
webertj@14456
|
283 |
val {interpreters, printers, parameters} = RefuteData.get thy
|
webertj@14456
|
284 |
in
|
haftmann@17314
|
285 |
case AList.lookup (op =) printers name of
|
skalberg@15531
|
286 |
NONE => RefuteData.put {interpreters = interpreters, printers = (name, f) :: printers, parameters = parameters} thy
|
skalberg@15531
|
287 |
| SOME _ => error ("Printer " ^ name ^ " already declared")
|
webertj@14456
|
288 |
end;
|
webertj@14456
|
289 |
|
webertj@14456
|
290 |
(* ------------------------------------------------------------------------- *)
|
webertj@14456
|
291 |
(* set_default_param: stores the '(name, value)' pair in RefuteData's *)
|
webertj@14456
|
292 |
(* parameter table *)
|
webertj@14456
|
293 |
(* ------------------------------------------------------------------------- *)
|
webertj@14456
|
294 |
|
webertj@14456
|
295 |
(* (string * string) -> theory -> theory *)
|
webertj@14456
|
296 |
|
webertj@14456
|
297 |
fun set_default_param (name, value) thy =
|
webertj@14456
|
298 |
let
|
webertj@14456
|
299 |
val {interpreters, printers, parameters} = RefuteData.get thy
|
webertj@14456
|
300 |
in
|
wenzelm@17412
|
301 |
case Symtab.lookup parameters name of
|
skalberg@15531
|
302 |
NONE => RefuteData.put
|
webertj@14456
|
303 |
{interpreters = interpreters, printers = printers, parameters = Symtab.extend (parameters, [(name, value)])} thy
|
skalberg@15531
|
304 |
| SOME _ => RefuteData.put
|
wenzelm@17412
|
305 |
{interpreters = interpreters, printers = printers, parameters = Symtab.update (name, value) parameters} thy
|
webertj@14350
|
306 |
end;
|
webertj@14350
|
307 |
|
webertj@14350
|
308 |
(* ------------------------------------------------------------------------- *)
|
webertj@14456
|
309 |
(* get_default_param: retrieves the value associated with 'name' from *)
|
webertj@14456
|
310 |
(* RefuteData's parameter table *)
|
webertj@14456
|
311 |
(* ------------------------------------------------------------------------- *)
|
webertj@14456
|
312 |
|
webertj@14456
|
313 |
(* theory -> string -> string option *)
|
webertj@14456
|
314 |
|
wenzelm@17412
|
315 |
val get_default_param = Symtab.lookup o #parameters o RefuteData.get;
|
webertj@14456
|
316 |
|
webertj@14456
|
317 |
(* ------------------------------------------------------------------------- *)
|
webertj@14456
|
318 |
(* get_default_params: returns a list of all '(name, value)' pairs that are *)
|
webertj@14456
|
319 |
(* stored in RefuteData's parameter table *)
|
webertj@14456
|
320 |
(* ------------------------------------------------------------------------- *)
|
webertj@14456
|
321 |
|
webertj@14456
|
322 |
(* theory -> (string * string) list *)
|
webertj@14456
|
323 |
|
wenzelm@17261
|
324 |
val get_default_params = Symtab.dest o #parameters o RefuteData.get;
|
webertj@14456
|
325 |
|
webertj@14456
|
326 |
(* ------------------------------------------------------------------------- *)
|
webertj@14456
|
327 |
(* actual_params: takes a (possibly empty) list 'params' of parameters that *)
|
webertj@14456
|
328 |
(* override the default parameters currently specified in 'thy', and *)
|
webertj@14807
|
329 |
(* returns a record that can be passed to 'find_model'. *)
|
webertj@14456
|
330 |
(* ------------------------------------------------------------------------- *)
|
webertj@14456
|
331 |
|
webertj@14807
|
332 |
(* theory -> (string * string) list -> params *)
|
webertj@14456
|
333 |
|
webertj@14807
|
334 |
fun actual_params thy override =
|
webertj@14456
|
335 |
let
|
webertj@14456
|
336 |
(* (string * string) list * string -> int *)
|
webertj@14456
|
337 |
fun read_int (parms, name) =
|
haftmann@17314
|
338 |
case AList.lookup (op =) parms name of
|
skalberg@15531
|
339 |
SOME s => (case Int.fromString s of
|
webertj@14456
|
340 |
SOME i => i
|
webertj@14456
|
341 |
| NONE => error ("parameter " ^ quote name ^ " (value is " ^ quote s ^ ") must be an integer value"))
|
skalberg@15531
|
342 |
| NONE => error ("parameter " ^ quote name ^ " must be assigned a value")
|
webertj@14456
|
343 |
(* (string * string) list * string -> string *)
|
webertj@14456
|
344 |
fun read_string (parms, name) =
|
haftmann@17314
|
345 |
case AList.lookup (op =) parms name of
|
skalberg@15531
|
346 |
SOME s => s
|
skalberg@15531
|
347 |
| NONE => error ("parameter " ^ quote name ^ " must be assigned a value")
|
webertj@14456
|
348 |
(* (string * string) list *)
|
webertj@14807
|
349 |
val allparams = override @ (get_default_params thy) (* 'override' first, defaults last *)
|
webertj@14456
|
350 |
(* int *)
|
webertj@14456
|
351 |
val minsize = read_int (allparams, "minsize")
|
webertj@14456
|
352 |
val maxsize = read_int (allparams, "maxsize")
|
webertj@14456
|
353 |
val maxvars = read_int (allparams, "maxvars")
|
webertj@14807
|
354 |
val maxtime = read_int (allparams, "maxtime")
|
webertj@14456
|
355 |
(* string *)
|
webertj@14456
|
356 |
val satsolver = read_string (allparams, "satsolver")
|
webertj@14807
|
357 |
(* all remaining parameters of the form "string=int" are collected in *)
|
webertj@14807
|
358 |
(* 'sizes' *)
|
webertj@14807
|
359 |
(* TODO: it is currently not possible to specify a size for a type *)
|
webertj@14807
|
360 |
(* whose name is one of the other parameters (e.g. 'maxvars') *)
|
webertj@14807
|
361 |
(* (string * int) list *)
|
skalberg@15570
|
362 |
val sizes = List.mapPartial
|
skalberg@15531
|
363 |
(fn (name,value) => (case Int.fromString value of SOME i => SOME (name, i) | NONE => NONE))
|
skalberg@15570
|
364 |
(List.filter (fn (name,_) => name<>"minsize" andalso name<>"maxsize" andalso name<>"maxvars" andalso name<>"maxtime" andalso name<>"satsolver")
|
webertj@14807
|
365 |
allparams)
|
webertj@14456
|
366 |
in
|
webertj@14807
|
367 |
{sizes=sizes, minsize=minsize, maxsize=maxsize, maxvars=maxvars, maxtime=maxtime, satsolver=satsolver}
|
webertj@14807
|
368 |
end;
|
webertj@14807
|
369 |
|
webertj@14807
|
370 |
|
webertj@14807
|
371 |
(* ------------------------------------------------------------------------- *)
|
webertj@14807
|
372 |
(* TRANSLATION HOL -> PROPOSITIONAL LOGIC, BOOLEAN ASSIGNMENT -> MODEL *)
|
webertj@14807
|
373 |
(* ------------------------------------------------------------------------- *)
|
webertj@14807
|
374 |
|
webertj@14807
|
375 |
(* ------------------------------------------------------------------------- *)
|
webertj@15335
|
376 |
(* typ_of_dtyp: converts a data type ('DatatypeAux.dtyp') into a type *)
|
webertj@15335
|
377 |
(* ('Term.typ'), given type parameters for the data type's type *)
|
webertj@15335
|
378 |
(* arguments *)
|
webertj@15335
|
379 |
(* ------------------------------------------------------------------------- *)
|
webertj@15335
|
380 |
|
webertj@15335
|
381 |
(* DatatypeAux.descr -> (DatatypeAux.dtyp * Term.typ) list -> DatatypeAux.dtyp -> Term.typ *)
|
webertj@15335
|
382 |
|
webertj@15335
|
383 |
fun typ_of_dtyp descr typ_assoc (DatatypeAux.DtTFree a) =
|
webertj@15335
|
384 |
(* replace a 'DtTFree' variable by the associated type *)
|
haftmann@17314
|
385 |
(the o AList.lookup (op =) typ_assoc) (DatatypeAux.DtTFree a)
|
webertj@15547
|
386 |
| typ_of_dtyp descr typ_assoc (DatatypeAux.DtType (s, ds)) =
|
webertj@15547
|
387 |
Type (s, map (typ_of_dtyp descr typ_assoc) ds)
|
webertj@15335
|
388 |
| typ_of_dtyp descr typ_assoc (DatatypeAux.DtRec i) =
|
webertj@15335
|
389 |
let
|
haftmann@17314
|
390 |
val (s, ds, _) = (the o AList.lookup (op =) descr) i
|
webertj@15335
|
391 |
in
|
webertj@15335
|
392 |
Type (s, map (typ_of_dtyp descr typ_assoc) ds)
|
webertj@15547
|
393 |
end;
|
webertj@15335
|
394 |
|
webertj@15335
|
395 |
(* ------------------------------------------------------------------------- *)
|
webertj@14807
|
396 |
(* collect_axioms: collects (monomorphic, universally quantified versions *)
|
webertj@14807
|
397 |
(* of) all HOL axioms that are relevant w.r.t 't' *)
|
webertj@14807
|
398 |
(* ------------------------------------------------------------------------- *)
|
webertj@14807
|
399 |
|
webertj@15547
|
400 |
(* Note: to make the collection of axioms more easily extensible, this *)
|
webertj@14807
|
401 |
(* function could be based on user-supplied "axiom collectors", *)
|
webertj@14807
|
402 |
(* similar to 'interpret'/interpreters or 'print'/printers *)
|
webertj@14807
|
403 |
|
webertj@14807
|
404 |
(* theory -> Term.term -> Term.term list *)
|
webertj@14807
|
405 |
|
webertj@14807
|
406 |
(* Which axioms are "relevant" for a particular term/type goes hand in *)
|
webertj@14807
|
407 |
(* hand with the interpretation of that term/type by its interpreter (see *)
|
webertj@14807
|
408 |
(* way below): if the interpretation respects an axiom anyway, the axiom *)
|
webertj@14807
|
409 |
(* does not need to be added as a constraint here. *)
|
webertj@14807
|
410 |
|
webertj@14807
|
411 |
(* When an axiom is added as relevant, further axioms may need to be *)
|
webertj@14807
|
412 |
(* added as well (e.g. when a constant is defined in terms of other *)
|
webertj@14807
|
413 |
(* constants). To avoid infinite recursion (which should not happen for *)
|
webertj@14807
|
414 |
(* constants anyway, but it could happen for "typedef"-related axioms, *)
|
webertj@14807
|
415 |
(* since they contain the type again), we use an accumulator 'axs' and *)
|
webertj@14807
|
416 |
(* add a relevant axiom only if it is not in 'axs' yet. *)
|
webertj@14807
|
417 |
|
webertj@14807
|
418 |
fun collect_axioms thy t =
|
webertj@14807
|
419 |
let
|
wenzelm@14984
|
420 |
val _ = immediate_output "Adding axioms..."
|
webertj@14807
|
421 |
(* (string * Term.term) list *)
|
wenzelm@16331
|
422 |
val axioms = Theory.all_axioms_of thy;
|
webertj@15547
|
423 |
(* string list *)
|
webertj@15547
|
424 |
val rec_names = Symtab.foldl (fn (acc, (_, info)) =>
|
webertj@15547
|
425 |
#rec_names info @ acc) ([], DatatypePackage.get_datatypes thy)
|
webertj@15547
|
426 |
(* string list *)
|
wenzelm@18932
|
427 |
val const_of_class_names = map Logic.const_of_class (Sign.classes (sign_of thy))
|
webertj@14807
|
428 |
(* given a constant 's' of type 'T', which is a subterm of 't', where *)
|
webertj@14807
|
429 |
(* 't' has a (possibly) more general type, the schematic type *)
|
webertj@15547
|
430 |
(* variables in 't' are instantiated to match the type 'T' (may raise *)
|
webertj@15547
|
431 |
(* Type.TYPE_MATCH) *)
|
webertj@14807
|
432 |
(* (string * Term.typ) * Term.term -> Term.term *)
|
webertj@14807
|
433 |
fun specialize_type ((s, T), t) =
|
webertj@14807
|
434 |
let
|
webertj@14807
|
435 |
fun find_typeSubs (Const (s', T')) =
|
webertj@14807
|
436 |
(if s=s' then
|
wenzelm@16935
|
437 |
SOME (Sign.typ_match thy (T', T) Vartab.empty) handle Type.TYPE_MATCH => NONE
|
webertj@14807
|
438 |
else
|
webertj@15547
|
439 |
NONE)
|
skalberg@15531
|
440 |
| find_typeSubs (Free _) = NONE
|
skalberg@15531
|
441 |
| find_typeSubs (Var _) = NONE
|
skalberg@15531
|
442 |
| find_typeSubs (Bound _) = NONE
|
webertj@14807
|
443 |
| find_typeSubs (Abs (_, _, body)) = find_typeSubs body
|
skalberg@15531
|
444 |
| find_typeSubs (t1 $ t2) = (case find_typeSubs t1 of SOME x => SOME x | NONE => find_typeSubs t2)
|
webertj@14807
|
445 |
val typeSubs = (case find_typeSubs t of
|
skalberg@15531
|
446 |
SOME x => x
|
webertj@15547
|
447 |
| NONE => raise Type.TYPE_MATCH (* no match found - perhaps due to sort constraints *))
|
webertj@14807
|
448 |
in
|
webertj@14807
|
449 |
map_term_types
|
webertj@14807
|
450 |
(map_type_tvar
|
berghofe@15794
|
451 |
(fn v =>
|
berghofe@15794
|
452 |
case Type.lookup (typeSubs, v) of
|
skalberg@15531
|
453 |
NONE =>
|
webertj@14807
|
454 |
(* schematic type variable not instantiated *)
|
webertj@14807
|
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 ^ ")")
|
skalberg@15531
|
456 |
| SOME typ =>
|
webertj@14807
|
457 |
typ))
|
webertj@14807
|
458 |
t
|
webertj@14807
|
459 |
end
|
webertj@15280
|
460 |
(* applies a type substitution 'typeSubs' for all type variables in a *)
|
webertj@15280
|
461 |
(* term 't' *)
|
berghofe@15794
|
462 |
(* (Term.sort * Term.typ) Term.Vartab.table -> Term.term -> Term.term *)
|
webertj@15280
|
463 |
fun monomorphic_term typeSubs t =
|
webertj@15280
|
464 |
map_term_types (map_type_tvar
|
berghofe@15794
|
465 |
(fn v =>
|
berghofe@15794
|
466 |
case Type.lookup (typeSubs, v) of
|
skalberg@15531
|
467 |
NONE =>
|
webertj@15280
|
468 |
(* schematic type variable not instantiated *)
|
wenzelm@18678
|
469 |
error ""
|
skalberg@15531
|
470 |
| SOME typ =>
|
webertj@15280
|
471 |
typ)) t
|
webertj@14807
|
472 |
(* Term.term list * Term.typ -> Term.term list *)
|
webertj@15547
|
473 |
fun collect_sort_axioms (axs, T) =
|
webertj@15547
|
474 |
let
|
webertj@15547
|
475 |
(* collect the axioms for a single 'class' (but not for its superclasses) *)
|
webertj@15547
|
476 |
(* Term.term list * string -> Term.term list *)
|
webertj@15547
|
477 |
fun collect_class_axioms (axs, class) =
|
webertj@15547
|
478 |
let
|
webertj@15547
|
479 |
(* obtain the axioms generated by the "axclass" command *)
|
webertj@15547
|
480 |
(* (string * Term.term) list *)
|
wenzelm@18932
|
481 |
val class_axioms = List.filter (fn (s, _) => String.isPrefix (Logic.const_of_class class ^ ".axioms_") s) axioms
|
webertj@15547
|
482 |
(* replace the one schematic type variable in each axiom by the actual type 'T' *)
|
webertj@15547
|
483 |
(* (string * Term.term) list *)
|
webertj@15547
|
484 |
val monomorphic_class_axioms = map (fn (axname, ax) =>
|
webertj@15547
|
485 |
let
|
berghofe@15794
|
486 |
val (idx, S) = (case term_tvars ax of
|
webertj@15547
|
487 |
[is] => is
|
webertj@15547
|
488 |
| _ => raise REFUTE ("collect_axioms", "class axiom " ^ axname ^ " (" ^ Sign.string_of_term (sign_of thy) ax ^ ") does not contain exactly one type variable"))
|
webertj@15547
|
489 |
in
|
berghofe@15794
|
490 |
(axname, monomorphic_term (Vartab.make [(idx, (S, T))]) ax)
|
webertj@15547
|
491 |
end) class_axioms
|
webertj@15547
|
492 |
(* Term.term list * (string * Term.term) list -> Term.term list *)
|
webertj@15547
|
493 |
fun collect_axiom (axs, (axname, ax)) =
|
webertj@15547
|
494 |
if mem_term (ax, axs) then
|
webertj@15547
|
495 |
axs
|
webertj@15547
|
496 |
else (
|
webertj@15547
|
497 |
immediate_output (" " ^ axname);
|
webertj@15547
|
498 |
collect_term_axioms (ax :: axs, ax)
|
webertj@15547
|
499 |
)
|
webertj@15547
|
500 |
in
|
skalberg@15570
|
501 |
Library.foldl collect_axiom (axs, monomorphic_class_axioms)
|
webertj@15547
|
502 |
end
|
webertj@15547
|
503 |
(* string list *)
|
webertj@15547
|
504 |
val sort = (case T of
|
webertj@15547
|
505 |
TFree (_, sort) => sort
|
webertj@15547
|
506 |
| TVar (_, sort) => sort
|
webertj@15547
|
507 |
| _ => raise REFUTE ("collect_axioms", "type " ^ Sign.string_of_typ (sign_of thy) T ^ " is not a variable"))
|
webertj@15547
|
508 |
(* obtain all superclasses of classes in 'sort' *)
|
webertj@15547
|
509 |
(* string list *)
|
wenzelm@19642
|
510 |
val superclasses = distinct (op =) (sort @ maps (Sign.super_classes thy) sort)
|
webertj@15547
|
511 |
in
|
skalberg@15570
|
512 |
Library.foldl collect_class_axioms (axs, superclasses)
|
webertj@15547
|
513 |
end
|
webertj@15547
|
514 |
(* Term.term list * Term.typ -> Term.term list *)
|
webertj@15547
|
515 |
and collect_type_axioms (axs, T) =
|
webertj@14807
|
516 |
case T of
|
webertj@14807
|
517 |
(* simple types *)
|
webertj@14807
|
518 |
Type ("prop", []) => axs
|
webertj@14807
|
519 |
| Type ("fun", [T1, T2]) => collect_type_axioms (collect_type_axioms (axs, T1), T2)
|
webertj@14807
|
520 |
| Type ("set", [T1]) => collect_type_axioms (axs, T1)
|
webertj@15547
|
521 |
| Type ("itself", [T1]) => collect_type_axioms (axs, T1) (* axiomatic type classes *)
|
webertj@14807
|
522 |
| Type (s, Ts) =>
|
webertj@14807
|
523 |
let
|
webertj@14807
|
524 |
(* look up the definition of a type, as created by "typedef" *)
|
webertj@14807
|
525 |
(* (string * Term.term) list -> (string * Term.term) option *)
|
webertj@14807
|
526 |
fun get_typedefn [] =
|
skalberg@15531
|
527 |
NONE
|
webertj@14807
|
528 |
| get_typedefn ((axname,ax)::axms) =
|
webertj@14807
|
529 |
(let
|
webertj@14807
|
530 |
(* Term.term -> Term.typ option *)
|
webertj@14807
|
531 |
fun type_of_type_definition (Const (s', T')) =
|
webertj@14807
|
532 |
if s'="Typedef.type_definition" then
|
skalberg@15531
|
533 |
SOME T'
|
webertj@14807
|
534 |
else
|
skalberg@15531
|
535 |
NONE
|
skalberg@15531
|
536 |
| type_of_type_definition (Free _) = NONE
|
skalberg@15531
|
537 |
| type_of_type_definition (Var _) = NONE
|
skalberg@15531
|
538 |
| type_of_type_definition (Bound _) = NONE
|
webertj@14807
|
539 |
| type_of_type_definition (Abs (_, _, body)) = type_of_type_definition body
|
skalberg@15531
|
540 |
| type_of_type_definition (t1 $ t2) = (case type_of_type_definition t1 of SOME x => SOME x | NONE => type_of_type_definition t2)
|
webertj@14807
|
541 |
in
|
webertj@14807
|
542 |
case type_of_type_definition ax of
|
skalberg@15531
|
543 |
SOME T' =>
|
webertj@14807
|
544 |
let
|
webertj@14807
|
545 |
val T'' = (domain_type o domain_type) T'
|
wenzelm@16935
|
546 |
val typeSubs = Sign.typ_match thy (T'', T) Vartab.empty
|
webertj@14807
|
547 |
in
|
skalberg@15531
|
548 |
SOME (axname, monomorphic_term typeSubs ax)
|
webertj@14807
|
549 |
end
|
skalberg@15531
|
550 |
| NONE =>
|
webertj@14807
|
551 |
get_typedefn axms
|
webertj@14807
|
552 |
end
|
wenzelm@18678
|
553 |
handle ERROR _ => get_typedefn axms
|
webertj@14807
|
554 |
| MATCH => get_typedefn axms
|
webertj@14807
|
555 |
| Type.TYPE_MATCH => get_typedefn axms)
|
webertj@14807
|
556 |
in
|
haftmann@19346
|
557 |
case DatatypePackage.get_datatype thy s of
|
skalberg@15531
|
558 |
SOME info => (* inductive datatype *)
|
webertj@14807
|
559 |
(* only collect relevant type axioms for the argument types *)
|
skalberg@15570
|
560 |
Library.foldl collect_type_axioms (axs, Ts)
|
skalberg@15531
|
561 |
| NONE =>
|
webertj@14807
|
562 |
(case get_typedefn axioms of
|
skalberg@15531
|
563 |
SOME (axname, ax) =>
|
webertj@14807
|
564 |
if mem_term (ax, axs) then
|
webertj@15547
|
565 |
(* only collect relevant type axioms for the argument types *)
|
skalberg@15570
|
566 |
Library.foldl collect_type_axioms (axs, Ts)
|
webertj@14807
|
567 |
else
|
wenzelm@14984
|
568 |
(immediate_output (" " ^ axname);
|
webertj@14807
|
569 |
collect_term_axioms (ax :: axs, ax))
|
skalberg@15531
|
570 |
| NONE =>
|
webertj@15547
|
571 |
(* unspecified type, perhaps introduced with 'typedecl' *)
|
webertj@14807
|
572 |
(* at least collect relevant type axioms for the argument types *)
|
skalberg@15570
|
573 |
Library.foldl collect_type_axioms (axs, Ts))
|
webertj@14807
|
574 |
end
|
webertj@15547
|
575 |
| TFree _ => collect_sort_axioms (axs, T) (* axiomatic type classes *)
|
webertj@15547
|
576 |
| TVar _ => collect_sort_axioms (axs, T) (* axiomatic type classes *)
|
webertj@14807
|
577 |
(* Term.term list * Term.term -> Term.term list *)
|
webertj@14807
|
578 |
and collect_term_axioms (axs, t) =
|
webertj@14807
|
579 |
case t of
|
webertj@14807
|
580 |
(* Pure *)
|
webertj@14807
|
581 |
Const ("all", _) => axs
|
webertj@14807
|
582 |
| Const ("==", _) => axs
|
webertj@14807
|
583 |
| Const ("==>", _) => axs
|
webertj@15547
|
584 |
| Const ("TYPE", T) => collect_type_axioms (axs, T) (* axiomatic type classes *)
|
webertj@14807
|
585 |
(* HOL *)
|
webertj@14807
|
586 |
| Const ("Trueprop", _) => axs
|
webertj@14807
|
587 |
| Const ("Not", _) => axs
|
webertj@14807
|
588 |
| Const ("True", _) => axs (* redundant, since 'True' is also an IDT constructor *)
|
webertj@14807
|
589 |
| Const ("False", _) => axs (* redundant, since 'False' is also an IDT constructor *)
|
webertj@14807
|
590 |
| Const ("arbitrary", T) => collect_type_axioms (axs, T)
|
webertj@14807
|
591 |
| Const ("The", T) =>
|
webertj@14807
|
592 |
let
|
haftmann@17314
|
593 |
val ax = specialize_type (("The", T), (the o AList.lookup (op =) axioms) "HOL.the_eq_trivial")
|
webertj@14807
|
594 |
in
|
webertj@14807
|
595 |
if mem_term (ax, axs) then
|
webertj@14807
|
596 |
collect_type_axioms (axs, T)
|
webertj@14807
|
597 |
else
|
wenzelm@14984
|
598 |
(immediate_output " HOL.the_eq_trivial";
|
webertj@14807
|
599 |
collect_term_axioms (ax :: axs, ax))
|
webertj@14807
|
600 |
end
|
webertj@14807
|
601 |
| Const ("Hilbert_Choice.Eps", T) =>
|
webertj@14807
|
602 |
let
|
haftmann@17314
|
603 |
val ax = specialize_type (("Hilbert_Choice.Eps", T),
|
haftmann@17314
|
604 |
(the o AList.lookup (op =) axioms) "Hilbert_Choice.someI")
|
webertj@14807
|
605 |
in
|
webertj@14807
|
606 |
if mem_term (ax, axs) then
|
webertj@14807
|
607 |
collect_type_axioms (axs, T)
|
webertj@14807
|
608 |
else
|
wenzelm@14984
|
609 |
(immediate_output " Hilbert_Choice.someI";
|
webertj@14807
|
610 |
collect_term_axioms (ax :: axs, ax))
|
webertj@14807
|
611 |
end
|
webertj@14807
|
612 |
| Const ("All", _) $ t1 => collect_term_axioms (axs, t1)
|
webertj@14807
|
613 |
| Const ("Ex", _) $ t1 => collect_term_axioms (axs, t1)
|
webertj@14807
|
614 |
| Const ("op =", T) => collect_type_axioms (axs, T)
|
webertj@14807
|
615 |
| Const ("op &", _) => axs
|
webertj@14807
|
616 |
| Const ("op |", _) => axs
|
webertj@14807
|
617 |
| Const ("op -->", _) => axs
|
webertj@14807
|
618 |
(* sets *)
|
webertj@14807
|
619 |
| Const ("Collect", T) => collect_type_axioms (axs, T)
|
webertj@14807
|
620 |
| Const ("op :", T) => collect_type_axioms (axs, T)
|
webertj@14807
|
621 |
(* other optimizations *)
|
webertj@14807
|
622 |
| Const ("Finite_Set.card", T) => collect_type_axioms (axs, T)
|
webertj@15571
|
623 |
| Const ("Finite_Set.Finites", T) => collect_type_axioms (axs, T)
|
haftmann@19277
|
624 |
| Const ("Orderings.less", T as Type ("fun", [Type ("nat", []), Type ("fun", [Type ("nat", []), Type ("bool", [])])])) => collect_type_axioms (axs, T)
|
haftmann@19233
|
625 |
| Const ("HOL.plus", T as Type ("fun", [Type ("nat", []), Type ("fun", [Type ("nat", []), Type ("nat", [])])])) => collect_type_axioms (axs, T)
|
haftmann@19233
|
626 |
| Const ("HOL.minus", T as Type ("fun", [Type ("nat", []), Type ("fun", [Type ("nat", []), Type ("nat", [])])])) => collect_type_axioms (axs, T)
|
haftmann@19233
|
627 |
| Const ("HOL.times", T as Type ("fun", [Type ("nat", []), Type ("fun", [Type ("nat", []), Type ("nat", [])])])) => collect_type_axioms (axs, T)
|
webertj@15767
|
628 |
| Const ("List.op @", T) => collect_type_axioms (axs, T)
|
webertj@16050
|
629 |
| Const ("Lfp.lfp", T) => collect_type_axioms (axs, T)
|
webertj@16050
|
630 |
| Const ("Gfp.gfp", T) => collect_type_axioms (axs, T)
|
webertj@14807
|
631 |
(* simply-typed lambda calculus *)
|
webertj@14807
|
632 |
| Const (s, T) =>
|
webertj@14807
|
633 |
let
|
webertj@14807
|
634 |
(* look up the definition of a constant, as created by "constdefs" *)
|
webertj@14807
|
635 |
(* string -> Term.typ -> (string * Term.term) list -> (string * Term.term) option *)
|
webertj@14807
|
636 |
fun get_defn [] =
|
skalberg@15531
|
637 |
NONE
|
webertj@15547
|
638 |
| get_defn ((axname, ax)::axms) =
|
webertj@14807
|
639 |
(let
|
webertj@14807
|
640 |
val (lhs, _) = Logic.dest_equals ax (* equations only *)
|
webertj@14807
|
641 |
val c = head_of lhs
|
webertj@14807
|
642 |
val (s', T') = dest_Const c
|
webertj@14807
|
643 |
in
|
webertj@14807
|
644 |
if s=s' then
|
webertj@14807
|
645 |
let
|
wenzelm@16935
|
646 |
val typeSubs = Sign.typ_match thy (T', T) Vartab.empty
|
webertj@14807
|
647 |
in
|
skalberg@15531
|
648 |
SOME (axname, monomorphic_term typeSubs ax)
|
webertj@14807
|
649 |
end
|
webertj@14807
|
650 |
else
|
webertj@14807
|
651 |
get_defn axms
|
webertj@14807
|
652 |
end
|
wenzelm@18678
|
653 |
handle ERROR _ => get_defn axms
|
webertj@14807
|
654 |
| TERM _ => get_defn axms
|
webertj@14807
|
655 |
| Type.TYPE_MATCH => get_defn axms)
|
webertj@15547
|
656 |
(* axiomatic type classes *)
|
webertj@15547
|
657 |
(* unit -> bool *)
|
webertj@15547
|
658 |
fun is_const_of_class () =
|
webertj@15547
|
659 |
(* I'm not quite sure if checking the name 's' is sufficient, *)
|
webertj@15547
|
660 |
(* or if we should also check the type 'T' *)
|
webertj@15547
|
661 |
s mem const_of_class_names
|
webertj@15547
|
662 |
(* inductive data types *)
|
webertj@15547
|
663 |
(* unit -> bool *)
|
webertj@15547
|
664 |
fun is_IDT_constructor () =
|
webertj@15547
|
665 |
(case body_type T of
|
webertj@15547
|
666 |
Type (s', _) =>
|
haftmann@19346
|
667 |
(case DatatypePackage.get_datatype_constrs thy s' of
|
webertj@15547
|
668 |
SOME constrs =>
|
haftmann@19346
|
669 |
Library.exists (fn (cname, cty) =>
|
haftmann@19346
|
670 |
cname = s andalso Sign.typ_instance thy (T, cty))
|
webertj@15547
|
671 |
constrs
|
webertj@15547
|
672 |
| NONE =>
|
webertj@14807
|
673 |
false)
|
webertj@15547
|
674 |
| _ =>
|
webertj@15547
|
675 |
false)
|
webertj@15547
|
676 |
(* unit -> bool *)
|
webertj@15547
|
677 |
fun is_IDT_recursor () =
|
webertj@15547
|
678 |
(* I'm not quite sure if checking the name 's' is sufficient, *)
|
webertj@15547
|
679 |
(* or if we should also check the type 'T' *)
|
webertj@15547
|
680 |
s mem rec_names
|
webertj@14807
|
681 |
in
|
webertj@15547
|
682 |
if is_const_of_class () then
|
webertj@15547
|
683 |
(* axiomatic type classes: add "OFCLASS(?'a::c, c_class)" and *)
|
webertj@15547
|
684 |
(* the introduction rule "class.intro" as axioms *)
|
webertj@15547
|
685 |
let
|
wenzelm@18932
|
686 |
val class = Logic.class_of_const s
|
webertj@15547
|
687 |
val inclass = Logic.mk_inclass (TVar (("'a", 0), [class]), class)
|
webertj@15547
|
688 |
(* Term.term option *)
|
webertj@15547
|
689 |
val ofclass_ax = (SOME (specialize_type ((s, T), inclass)) handle Type.TYPE_MATCH => NONE)
|
haftmann@17314
|
690 |
val intro_ax = (Option.map (fn t => specialize_type ((s, T), t))
|
haftmann@17314
|
691 |
(AList.lookup (op =) axioms (class ^ ".intro")) handle Type.TYPE_MATCH => NONE)
|
webertj@15547
|
692 |
val axs' = (case ofclass_ax of NONE => axs | SOME ax => if mem_term (ax, axs) then
|
webertj@15547
|
693 |
(* collect relevant type axioms *)
|
webertj@15547
|
694 |
collect_type_axioms (axs, T)
|
webertj@15547
|
695 |
else
|
webertj@15547
|
696 |
(immediate_output (" " ^ Sign.string_of_term (sign_of thy) ax);
|
webertj@15547
|
697 |
collect_term_axioms (ax :: axs, ax)))
|
webertj@15547
|
698 |
val axs'' = (case intro_ax of NONE => axs' | SOME ax => if mem_term (ax, axs') then
|
webertj@15547
|
699 |
(* collect relevant type axioms *)
|
webertj@15547
|
700 |
collect_type_axioms (axs', T)
|
webertj@15547
|
701 |
else
|
wenzelm@17274
|
702 |
(immediate_output (" " ^ s ^ ".intro");
|
webertj@15547
|
703 |
collect_term_axioms (ax :: axs', ax)))
|
webertj@15547
|
704 |
in
|
webertj@15547
|
705 |
axs''
|
webertj@15547
|
706 |
end
|
webertj@15547
|
707 |
else if is_IDT_constructor () then
|
webertj@14807
|
708 |
(* only collect relevant type axioms *)
|
webertj@14807
|
709 |
collect_type_axioms (axs, T)
|
webertj@15547
|
710 |
else if is_IDT_recursor () then
|
webertj@15125
|
711 |
(* only collect relevant type axioms *)
|
webertj@15125
|
712 |
collect_type_axioms (axs, T)
|
webertj@15547
|
713 |
else (
|
webertj@15547
|
714 |
case get_defn axioms of
|
skalberg@15531
|
715 |
SOME (axname, ax) =>
|
webertj@14807
|
716 |
if mem_term (ax, axs) then
|
webertj@14807
|
717 |
(* collect relevant type axioms *)
|
webertj@14807
|
718 |
collect_type_axioms (axs, T)
|
webertj@14807
|
719 |
else
|
wenzelm@14984
|
720 |
(immediate_output (" " ^ axname);
|
webertj@14807
|
721 |
collect_term_axioms (ax :: axs, ax))
|
skalberg@15531
|
722 |
| NONE =>
|
webertj@14807
|
723 |
(* collect relevant type axioms *)
|
webertj@15547
|
724 |
collect_type_axioms (axs, T)
|
webertj@15547
|
725 |
)
|
webertj@14807
|
726 |
end
|
webertj@14807
|
727 |
| Free (_, T) => collect_type_axioms (axs, T)
|
webertj@14807
|
728 |
| Var (_, T) => collect_type_axioms (axs, T)
|
webertj@14807
|
729 |
| Bound i => axs
|
webertj@14807
|
730 |
| Abs (_, T, body) => collect_term_axioms (collect_type_axioms (axs, T), body)
|
webertj@14807
|
731 |
| t1 $ t2 => collect_term_axioms (collect_term_axioms (axs, t1), t2)
|
webertj@14807
|
732 |
(* universal closure over schematic variables *)
|
webertj@14807
|
733 |
(* Term.term -> Term.term *)
|
webertj@14807
|
734 |
fun close_form t =
|
webertj@14807
|
735 |
let
|
webertj@14807
|
736 |
(* (Term.indexname * Term.typ) list *)
|
webertj@14807
|
737 |
val vars = sort_wrt (fst o fst) (map dest_Var (term_vars t))
|
webertj@14807
|
738 |
in
|
skalberg@15570
|
739 |
Library.foldl
|
webertj@15547
|
740 |
(fn (t', ((x, i), T)) => (Term.all T) $ Abs (x, T, abstract_over (Var((x, i), T), t')))
|
webertj@14807
|
741 |
(t, vars)
|
webertj@14807
|
742 |
end
|
webertj@14807
|
743 |
(* Term.term list *)
|
webertj@14807
|
744 |
val result = map close_form (collect_term_axioms ([], t))
|
webertj@14807
|
745 |
val _ = writeln " ...done."
|
webertj@14807
|
746 |
in
|
webertj@14807
|
747 |
result
|
webertj@14456
|
748 |
end;
|
webertj@14456
|
749 |
|
webertj@14456
|
750 |
(* ------------------------------------------------------------------------- *)
|
webertj@14807
|
751 |
(* ground_types: collects all ground types in a term (including argument *)
|
webertj@14807
|
752 |
(* types of other types), suppressing duplicates. Does not *)
|
webertj@14807
|
753 |
(* return function types, set types, non-recursive IDTs, or *)
|
webertj@14807
|
754 |
(* 'propT'. For IDTs, also the argument types of constructors *)
|
webertj@14807
|
755 |
(* are considered. *)
|
webertj@14807
|
756 |
(* ------------------------------------------------------------------------- *)
|
webertj@14807
|
757 |
|
webertj@14807
|
758 |
(* theory -> Term.term -> Term.typ list *)
|
webertj@14807
|
759 |
|
webertj@14807
|
760 |
fun ground_types thy t =
|
webertj@14807
|
761 |
let
|
webertj@14807
|
762 |
(* Term.typ * Term.typ list -> Term.typ list *)
|
webertj@14807
|
763 |
fun collect_types (T, acc) =
|
webertj@14807
|
764 |
if T mem acc then
|
webertj@14807
|
765 |
acc (* prevent infinite recursion (for IDTs) *)
|
webertj@14807
|
766 |
else
|
webertj@14807
|
767 |
(case T of
|
webertj@14807
|
768 |
Type ("fun", [T1, T2]) => collect_types (T1, collect_types (T2, acc))
|
webertj@14807
|
769 |
| Type ("prop", []) => acc
|
webertj@14807
|
770 |
| Type ("set", [T1]) => collect_types (T1, acc)
|
webertj@14807
|
771 |
| Type (s, Ts) =>
|
haftmann@19346
|
772 |
(case DatatypePackage.get_datatype thy s of
|
skalberg@15531
|
773 |
SOME info => (* inductive datatype *)
|
webertj@14807
|
774 |
let
|
webertj@14807
|
775 |
val index = #index info
|
webertj@14807
|
776 |
val descr = #descr info
|
haftmann@17314
|
777 |
val (_, dtyps, constrs) = (the o AList.lookup (op =) descr) index
|
webertj@14807
|
778 |
val typ_assoc = dtyps ~~ Ts
|
webertj@14807
|
779 |
(* sanity check: every element in 'dtyps' must be a 'DtTFree' *)
|
webertj@14807
|
780 |
val _ = (if Library.exists (fn d =>
|
webertj@14807
|
781 |
case d of DatatypeAux.DtTFree _ => false | _ => true) dtyps
|
webertj@14807
|
782 |
then
|
webertj@14807
|
783 |
raise REFUTE ("ground_types", "datatype argument (for type " ^ Sign.string_of_typ (sign_of thy) (Type (s, Ts)) ^ ") is not a variable")
|
webertj@14807
|
784 |
else
|
webertj@14807
|
785 |
())
|
webertj@14807
|
786 |
(* if the current type is a recursive IDT (i.e. a depth is required), add it to 'acc' *)
|
webertj@14807
|
787 |
val acc' = (if Library.exists (fn (_, ds) => Library.exists DatatypeAux.is_rec_type ds) constrs then
|
webertj@14807
|
788 |
T ins acc
|
webertj@14807
|
789 |
else
|
webertj@14807
|
790 |
acc)
|
webertj@14807
|
791 |
(* collect argument types *)
|
skalberg@15574
|
792 |
val acc_args = foldr collect_types acc' Ts
|
webertj@14807
|
793 |
(* collect constructor types *)
|
skalberg@15574
|
794 |
val acc_constrs = foldr collect_types acc_args (List.concat (map (fn (_, ds) => map (typ_of_dtyp descr typ_assoc) ds) constrs))
|
webertj@14807
|
795 |
in
|
webertj@14807
|
796 |
acc_constrs
|
webertj@14807
|
797 |
end
|
skalberg@15531
|
798 |
| NONE => (* not an inductive datatype, e.g. defined via "typedef" or "typedecl" *)
|
skalberg@15574
|
799 |
T ins (foldr collect_types acc Ts))
|
webertj@14807
|
800 |
| TFree _ => T ins acc
|
webertj@14807
|
801 |
| TVar _ => T ins acc)
|
webertj@14807
|
802 |
in
|
webertj@14807
|
803 |
it_term_types collect_types (t, [])
|
webertj@14807
|
804 |
end;
|
webertj@14807
|
805 |
|
webertj@14807
|
806 |
(* ------------------------------------------------------------------------- *)
|
webertj@14807
|
807 |
(* string_of_typ: (rather naive) conversion from types to strings, used to *)
|
webertj@14807
|
808 |
(* look up the size of a type in 'sizes'. Parameterized *)
|
webertj@14807
|
809 |
(* types with different parameters (e.g. "'a list" vs. "bool *)
|
webertj@14807
|
810 |
(* list") are identified. *)
|
webertj@14807
|
811 |
(* ------------------------------------------------------------------------- *)
|
webertj@14807
|
812 |
|
webertj@14807
|
813 |
(* Term.typ -> string *)
|
webertj@14807
|
814 |
|
webertj@14807
|
815 |
fun string_of_typ (Type (s, _)) = s
|
webertj@14807
|
816 |
| string_of_typ (TFree (s, _)) = s
|
webertj@14807
|
817 |
| string_of_typ (TVar ((s,_), _)) = s;
|
webertj@14807
|
818 |
|
webertj@14807
|
819 |
(* ------------------------------------------------------------------------- *)
|
webertj@14807
|
820 |
(* first_universe: returns the "first" (i.e. smallest) universe by assigning *)
|
webertj@14807
|
821 |
(* 'minsize' to every type for which no size is specified in *)
|
webertj@14807
|
822 |
(* 'sizes' *)
|
webertj@14807
|
823 |
(* ------------------------------------------------------------------------- *)
|
webertj@14807
|
824 |
|
webertj@14807
|
825 |
(* Term.typ list -> (string * int) list -> int -> (Term.typ * int) list *)
|
webertj@14807
|
826 |
|
webertj@14807
|
827 |
fun first_universe xs sizes minsize =
|
webertj@14807
|
828 |
let
|
webertj@14807
|
829 |
fun size_of_typ T =
|
haftmann@17314
|
830 |
case AList.lookup (op =) sizes (string_of_typ T) of
|
skalberg@15531
|
831 |
SOME n => n
|
skalberg@15531
|
832 |
| NONE => minsize
|
webertj@14807
|
833 |
in
|
webertj@14807
|
834 |
map (fn T => (T, size_of_typ T)) xs
|
webertj@14807
|
835 |
end;
|
webertj@14807
|
836 |
|
webertj@14807
|
837 |
(* ------------------------------------------------------------------------- *)
|
webertj@14807
|
838 |
(* next_universe: enumerates all universes (i.e. assignments of sizes to *)
|
webertj@14807
|
839 |
(* types), where the minimal size of a type is given by *)
|
webertj@14807
|
840 |
(* 'minsize', the maximal size is given by 'maxsize', and a *)
|
webertj@14807
|
841 |
(* type may have a fixed size given in 'sizes' *)
|
webertj@14456
|
842 |
(* ------------------------------------------------------------------------- *)
|
webertj@14456
|
843 |
|
webertj@14807
|
844 |
(* (Term.typ * int) list -> (string * int) list -> int -> int -> (Term.typ * int) list option *)
|
webertj@14456
|
845 |
|
webertj@14807
|
846 |
fun next_universe xs sizes minsize maxsize =
|
webertj@14456
|
847 |
let
|
webertj@14807
|
848 |
(* creates the "first" list of length 'len', where the sum of all list *)
|
webertj@14807
|
849 |
(* elements is 'sum', and the length of the list is 'len' *)
|
webertj@14807
|
850 |
(* int -> int -> int -> int list option *)
|
webertj@15547
|
851 |
fun make_first _ 0 sum =
|
webertj@14807
|
852 |
if sum=0 then
|
skalberg@15531
|
853 |
SOME []
|
webertj@14807
|
854 |
else
|
skalberg@15531
|
855 |
NONE
|
webertj@15547
|
856 |
| make_first max len sum =
|
webertj@14807
|
857 |
if sum<=max orelse max<0 then
|
skalberg@15570
|
858 |
Option.map (fn xs' => sum :: xs') (make_first max (len-1) 0)
|
webertj@14807
|
859 |
else
|
skalberg@15570
|
860 |
Option.map (fn xs' => max :: xs') (make_first max (len-1) (sum-max))
|
webertj@14807
|
861 |
(* enumerates all int lists with a fixed length, where 0<=x<='max' for *)
|
webertj@14807
|
862 |
(* all list elements x (unless 'max'<0) *)
|
webertj@15547
|
863 |
(* int -> int -> int -> int list -> int list option *)
|
webertj@15547
|
864 |
fun next max len sum [] =
|
webertj@15547
|
865 |
NONE
|
webertj@15547
|
866 |
| next max len sum [x] =
|
webertj@15547
|
867 |
(* we've reached the last list element, so there's no shift possible *)
|
webertj@15547
|
868 |
make_first max (len+1) (sum+x+1) (* increment 'sum' by 1 *)
|
webertj@15547
|
869 |
| next max len sum (x1::x2::xs) =
|
webertj@15547
|
870 |
if x1>0 andalso (x2<max orelse max<0) then
|
webertj@15547
|
871 |
(* we can shift *)
|
skalberg@15570
|
872 |
SOME (valOf (make_first max (len+1) (sum+x1-1)) @ (x2+1) :: xs)
|
webertj@15547
|
873 |
else
|
webertj@15547
|
874 |
(* continue search *)
|
webertj@15547
|
875 |
next max (len+1) (sum+x1) (x2::xs)
|
webertj@14807
|
876 |
(* only consider those types for which the size is not fixed *)
|
haftmann@17314
|
877 |
val mutables = List.filter (not o (AList.defined (op =) sizes) o string_of_typ o fst) xs
|
webertj@14807
|
878 |
(* subtract 'minsize' from every size (will be added again at the end) *)
|
webertj@14807
|
879 |
val diffs = map (fn (_, n) => n-minsize) mutables
|
webertj@14807
|
880 |
in
|
webertj@15547
|
881 |
case next (maxsize-minsize) 0 0 diffs of
|
skalberg@15531
|
882 |
SOME diffs' =>
|
webertj@14807
|
883 |
(* merge with those types for which the size is fixed *)
|
skalberg@15531
|
884 |
SOME (snd (foldl_map (fn (ds, (T, _)) =>
|
haftmann@17314
|
885 |
case AList.lookup (op =) sizes (string_of_typ T) of
|
webertj@15547
|
886 |
SOME n => (ds, (T, n)) (* return the fixed size *)
|
webertj@15547
|
887 |
| NONE => (tl ds, (T, minsize + hd ds))) (* consume the head of 'ds', add 'minsize' *)
|
webertj@14807
|
888 |
(diffs', xs)))
|
skalberg@15531
|
889 |
| NONE =>
|
skalberg@15531
|
890 |
NONE
|
webertj@14807
|
891 |
end;
|
webertj@14807
|
892 |
|
webertj@14807
|
893 |
(* ------------------------------------------------------------------------- *)
|
webertj@14807
|
894 |
(* toTrue: converts the interpretation of a Boolean value to a propositional *)
|
webertj@14807
|
895 |
(* formula that is true iff the interpretation denotes "true" *)
|
webertj@14807
|
896 |
(* ------------------------------------------------------------------------- *)
|
webertj@14807
|
897 |
|
webertj@14807
|
898 |
(* interpretation -> prop_formula *)
|
webertj@14807
|
899 |
|
webertj@15767
|
900 |
fun toTrue (Leaf [fm, _]) = fm
|
webertj@15767
|
901 |
| toTrue _ = raise REFUTE ("toTrue", "interpretation does not denote a Boolean value");
|
webertj@14807
|
902 |
|
webertj@14807
|
903 |
(* ------------------------------------------------------------------------- *)
|
webertj@14807
|
904 |
(* toFalse: converts the interpretation of a Boolean value to a *)
|
webertj@14807
|
905 |
(* propositional formula that is true iff the interpretation *)
|
webertj@14807
|
906 |
(* denotes "false" *)
|
webertj@14807
|
907 |
(* ------------------------------------------------------------------------- *)
|
webertj@14807
|
908 |
|
webertj@14807
|
909 |
(* interpretation -> prop_formula *)
|
webertj@14807
|
910 |
|
webertj@15767
|
911 |
fun toFalse (Leaf [_, fm]) = fm
|
webertj@15767
|
912 |
| toFalse _ = raise REFUTE ("toFalse", "interpretation does not denote a Boolean value");
|
webertj@14807
|
913 |
|
webertj@14807
|
914 |
(* ------------------------------------------------------------------------- *)
|
webertj@14807
|
915 |
(* find_model: repeatedly calls 'interpret' with appropriate parameters, *)
|
webertj@14807
|
916 |
(* applies a SAT solver, and (in case a model is found) displays *)
|
webertj@14807
|
917 |
(* the model to the user by calling 'print_model' *)
|
webertj@14807
|
918 |
(* thy : the current theory *)
|
webertj@14807
|
919 |
(* {...} : parameters that control the translation/model generation *)
|
webertj@14807
|
920 |
(* t : term to be translated into a propositional formula *)
|
webertj@14807
|
921 |
(* negate : if true, find a model that makes 't' false (rather than true) *)
|
webertj@14807
|
922 |
(* ------------------------------------------------------------------------- *)
|
webertj@14807
|
923 |
|
webertj@14807
|
924 |
(* theory -> params -> Term.term -> bool -> unit *)
|
webertj@14807
|
925 |
|
webertj@14807
|
926 |
fun find_model thy {sizes, minsize, maxsize, maxvars, maxtime, satsolver} t negate =
|
webertj@14807
|
927 |
let
|
webertj@14807
|
928 |
(* unit -> unit *)
|
webertj@14807
|
929 |
fun wrapper () =
|
webertj@14807
|
930 |
let
|
webertj@14807
|
931 |
(* Term.term list *)
|
webertj@14807
|
932 |
val axioms = collect_axioms thy t
|
webertj@14807
|
933 |
(* Term.typ list *)
|
skalberg@15570
|
934 |
val types = Library.foldl (fn (acc, t') => acc union (ground_types thy t')) ([], t :: axioms)
|
webertj@14807
|
935 |
val _ = writeln ("Ground types: "
|
webertj@14807
|
936 |
^ (if null types then "none."
|
webertj@14807
|
937 |
else commas (map (Sign.string_of_typ (sign_of thy)) types)))
|
webertj@15547
|
938 |
(* we can only consider fragments of recursive IDTs, so we issue a *)
|
webertj@15547
|
939 |
(* warning if the formula contains a recursive IDT *)
|
webertj@15547
|
940 |
(* TODO: no warning needed for /positive/ occurrences of IDTs *)
|
webertj@15547
|
941 |
val _ = if Library.exists (fn
|
webertj@15547
|
942 |
Type (s, _) =>
|
haftmann@19346
|
943 |
(case DatatypePackage.get_datatype thy s of
|
webertj@15547
|
944 |
SOME info => (* inductive datatype *)
|
webertj@15547
|
945 |
let
|
webertj@15547
|
946 |
val index = #index info
|
webertj@15547
|
947 |
val descr = #descr info
|
haftmann@17314
|
948 |
val (_, _, constrs) = (the o AList.lookup (op =) descr) index
|
webertj@15547
|
949 |
in
|
webertj@15547
|
950 |
(* recursive datatype? *)
|
webertj@15547
|
951 |
Library.exists (fn (_, ds) => Library.exists DatatypeAux.is_rec_type ds) constrs
|
webertj@15547
|
952 |
end
|
webertj@15547
|
953 |
| NONE => false)
|
webertj@15547
|
954 |
| _ => false) types then
|
webertj@15547
|
955 |
warning "Term contains a recursive datatype; countermodel(s) may be spurious!"
|
webertj@15547
|
956 |
else
|
webertj@15547
|
957 |
()
|
webertj@14807
|
958 |
(* (Term.typ * int) list -> unit *)
|
webertj@14807
|
959 |
fun find_model_loop universe =
|
webertj@15334
|
960 |
let
|
webertj@14807
|
961 |
val init_model = (universe, [])
|
webertj@15547
|
962 |
val init_args = {maxvars = maxvars, def_eq = false, next_idx = 1, bounds = [], wellformed = True}
|
wenzelm@14984
|
963 |
val _ = immediate_output ("Translating term (sizes: " ^ commas (map (fn (_, n) => string_of_int n) universe) ^ ") ...")
|
webertj@14807
|
964 |
(* translate 't' and all axioms *)
|
webertj@14807
|
965 |
val ((model, args), intrs) = foldl_map (fn ((m, a), t') =>
|
webertj@14807
|
966 |
let
|
webertj@14807
|
967 |
val (i, m', a') = interpret thy m a t'
|
webertj@14807
|
968 |
in
|
webertj@15547
|
969 |
(* set 'def_eq' to 'true' *)
|
webertj@15547
|
970 |
((m', {maxvars = #maxvars a', def_eq = true, next_idx = #next_idx a', bounds = #bounds a', wellformed = #wellformed a'}), i)
|
webertj@14807
|
971 |
end) ((init_model, init_args), t :: axioms)
|
webertj@14807
|
972 |
(* make 't' either true or false, and make all axioms true, and *)
|
webertj@14807
|
973 |
(* add the well-formedness side condition *)
|
webertj@14807
|
974 |
val fm_t = (if negate then toFalse else toTrue) (hd intrs)
|
webertj@14807
|
975 |
val fm_ax = PropLogic.all (map toTrue (tl intrs))
|
webertj@14807
|
976 |
val fm = PropLogic.all [#wellformed args, fm_ax, fm_t]
|
webertj@14456
|
977 |
in
|
wenzelm@14984
|
978 |
immediate_output " invoking SAT solver...";
|
webertj@14965
|
979 |
(case SatSolver.invoke_solver satsolver fm of
|
webertj@14965
|
980 |
SatSolver.SATISFIABLE assignment =>
|
webertj@15547
|
981 |
(writeln " model found!";
|
webertj@15547
|
982 |
writeln ("*** Model found: ***\n" ^ print_model thy model (fn i => case assignment i of SOME b => b | NONE => true)))
|
webertj@17493
|
983 |
| SatSolver.UNSATISFIABLE _ =>
|
webertj@15547
|
984 |
(immediate_output " no model exists.\n";
|
webertj@15547
|
985 |
case next_universe universe sizes minsize maxsize of
|
webertj@15547
|
986 |
SOME universe' => find_model_loop universe'
|
webertj@15547
|
987 |
| NONE => writeln "Search terminated, no larger universe within the given limits.")
|
webertj@15547
|
988 |
| SatSolver.UNKNOWN =>
|
wenzelm@14984
|
989 |
(immediate_output " no model found.\n";
|
webertj@14807
|
990 |
case next_universe universe sizes minsize maxsize of
|
skalberg@15531
|
991 |
SOME universe' => find_model_loop universe'
|
webertj@15547
|
992 |
| NONE => writeln "Search terminated, no larger universe within the given limits.")
|
webertj@15547
|
993 |
) handle SatSolver.NOT_CONFIGURED =>
|
webertj@14965
|
994 |
error ("SAT solver " ^ quote satsolver ^ " is not configured.")
|
webertj@14807
|
995 |
end handle MAXVARS_EXCEEDED =>
|
webertj@14807
|
996 |
writeln ("\nSearch terminated, number of Boolean variables (" ^ string_of_int maxvars ^ " allowed) exceeded.")
|
webertj@14807
|
997 |
in
|
webertj@14807
|
998 |
find_model_loop (first_universe types sizes minsize)
|
webertj@14456
|
999 |
end
|
webertj@14807
|
1000 |
in
|
webertj@14807
|
1001 |
(* some parameter sanity checks *)
|
webertj@14807
|
1002 |
assert (minsize>=1) ("\"minsize\" is " ^ string_of_int minsize ^ ", must be at least 1");
|
webertj@14807
|
1003 |
assert (maxsize>=1) ("\"maxsize\" is " ^ string_of_int maxsize ^ ", must be at least 1");
|
webertj@14807
|
1004 |
assert (maxsize>=minsize) ("\"maxsize\" (=" ^ string_of_int maxsize ^ ") is less than \"minsize\" (=" ^ string_of_int minsize ^ ").");
|
webertj@14807
|
1005 |
assert (maxvars>=0) ("\"maxvars\" is " ^ string_of_int maxvars ^ ", must be at least 0");
|
webertj@14807
|
1006 |
assert (maxtime>=0) ("\"maxtime\" is " ^ string_of_int maxtime ^ ", must be at least 0");
|
webertj@15547
|
1007 |
(* enter loop with or without time limit *)
|
webertj@14807
|
1008 |
writeln ("Trying to find a model that " ^ (if negate then "refutes" else "satisfies") ^ ": "
|
webertj@14807
|
1009 |
^ Sign.string_of_term (sign_of thy) t);
|
webertj@15547
|
1010 |
if maxtime>0 then (
|
webertj@18760
|
1011 |
interrupt_timeout (Time.fromSeconds (Int.toLarge maxtime))
|
webertj@14807
|
1012 |
wrapper ()
|
webertj@18760
|
1013 |
handle Interrupt =>
|
webertj@14807
|
1014 |
writeln ("\nSearch terminated, time limit ("
|
webertj@14965
|
1015 |
^ string_of_int maxtime ^ (if maxtime=1 then " second" else " seconds")
|
webertj@15547
|
1016 |
^ ") exceeded.")
|
webertj@15547
|
1017 |
) else
|
webertj@14807
|
1018 |
wrapper ()
|
webertj@14807
|
1019 |
end;
|
webertj@14456
|
1020 |
|
webertj@14456
|
1021 |
|
webertj@14456
|
1022 |
(* ------------------------------------------------------------------------- *)
|
webertj@14456
|
1023 |
(* INTERFACE, PART 2: FINDING A MODEL *)
|
webertj@14350
|
1024 |
(* ------------------------------------------------------------------------- *)
|
webertj@14350
|
1025 |
|
webertj@14350
|
1026 |
(* ------------------------------------------------------------------------- *)
|
webertj@14456
|
1027 |
(* satisfy_term: calls 'find_model' to find a model that satisfies 't' *)
|
webertj@14456
|
1028 |
(* params : list of '(name, value)' pairs used to override default *)
|
webertj@14456
|
1029 |
(* parameters *)
|
webertj@14350
|
1030 |
(* ------------------------------------------------------------------------- *)
|
webertj@14350
|
1031 |
|
webertj@14456
|
1032 |
(* theory -> (string * string) list -> Term.term -> unit *)
|
webertj@14350
|
1033 |
|
webertj@14456
|
1034 |
fun satisfy_term thy params t =
|
webertj@14807
|
1035 |
find_model thy (actual_params thy params) t false;
|
webertj@14350
|
1036 |
|
webertj@14350
|
1037 |
(* ------------------------------------------------------------------------- *)
|
webertj@14456
|
1038 |
(* refute_term: calls 'find_model' to find a model that refutes 't' *)
|
webertj@14456
|
1039 |
(* params : list of '(name, value)' pairs used to override default *)
|
webertj@14456
|
1040 |
(* parameters *)
|
webertj@14350
|
1041 |
(* ------------------------------------------------------------------------- *)
|
webertj@14350
|
1042 |
|
webertj@14456
|
1043 |
(* theory -> (string * string) list -> Term.term -> unit *)
|
webertj@14350
|
1044 |
|
webertj@14456
|
1045 |
fun refute_term thy params t =
|
webertj@14350
|
1046 |
let
|
webertj@14807
|
1047 |
(* disallow schematic type variables, since we cannot properly negate *)
|
webertj@14807
|
1048 |
(* terms containing them (their logical meaning is that there EXISTS a *)
|
webertj@14807
|
1049 |
(* type s.t. ...; to refute such a formula, we would have to show that *)
|
webertj@14807
|
1050 |
(* for ALL types, not ...) *)
|
webertj@14456
|
1051 |
val _ = assert (null (term_tvars t)) "Term to be refuted contains schematic type variables"
|
webertj@14456
|
1052 |
(* existential closure over schematic variables *)
|
webertj@14456
|
1053 |
(* (Term.indexname * Term.typ) list *)
|
webertj@14456
|
1054 |
val vars = sort_wrt (fst o fst) (map dest_Var (term_vars t))
|
webertj@14456
|
1055 |
(* Term.term *)
|
skalberg@15570
|
1056 |
val ex_closure = Library.foldl
|
webertj@15611
|
1057 |
(fn (t', ((x, i), T)) => (HOLogic.exists_const T) $ Abs (x, T, abstract_over (Var ((x, i), T), t')))
|
webertj@14456
|
1058 |
(t, vars)
|
webertj@14456
|
1059 |
(* If 't' is of type 'propT' (rather than 'boolT'), applying *)
|
webertj@14456
|
1060 |
(* 'HOLogic.exists_const' is not type-correct. However, this *)
|
webertj@14807
|
1061 |
(* is not really a problem as long as 'find_model' still *)
|
webertj@14456
|
1062 |
(* interprets the resulting term correctly, without checking *)
|
webertj@14456
|
1063 |
(* its type. *)
|
webertj@14350
|
1064 |
in
|
webertj@14807
|
1065 |
find_model thy (actual_params thy params) ex_closure true
|
webertj@14350
|
1066 |
end;
|
webertj@14350
|
1067 |
|
webertj@14350
|
1068 |
(* ------------------------------------------------------------------------- *)
|
webertj@14456
|
1069 |
(* refute_subgoal: calls 'refute_term' on a specific subgoal *)
|
webertj@14456
|
1070 |
(* params : list of '(name, value)' pairs used to override default *)
|
webertj@14456
|
1071 |
(* parameters *)
|
webertj@14456
|
1072 |
(* subgoal : 0-based index specifying the subgoal number *)
|
webertj@14350
|
1073 |
(* ------------------------------------------------------------------------- *)
|
webertj@14350
|
1074 |
|
webertj@14456
|
1075 |
(* theory -> (string * string) list -> Thm.thm -> int -> unit *)
|
webertj@14350
|
1076 |
|
webertj@14456
|
1077 |
fun refute_subgoal thy params thm subgoal =
|
skalberg@15570
|
1078 |
refute_term thy params (List.nth (prems_of thm, subgoal));
|
webertj@14350
|
1079 |
|
webertj@14350
|
1080 |
|
webertj@14350
|
1081 |
(* ------------------------------------------------------------------------- *)
|
webertj@15292
|
1082 |
(* INTERPRETERS: Auxiliary Functions *)
|
webertj@14350
|
1083 |
(* ------------------------------------------------------------------------- *)
|
webertj@14350
|
1084 |
|
webertj@14350
|
1085 |
(* ------------------------------------------------------------------------- *)
|
webertj@14807
|
1086 |
(* make_constants: returns all interpretations that have the same tree *)
|
webertj@14807
|
1087 |
(* structure as 'intr', but consist of unit vectors with *)
|
webertj@14807
|
1088 |
(* 'True'/'False' only (no Boolean variables) *)
|
webertj@14350
|
1089 |
(* ------------------------------------------------------------------------- *)
|
webertj@14350
|
1090 |
|
webertj@14807
|
1091 |
(* interpretation -> interpretation list *)
|
webertj@14350
|
1092 |
|
webertj@14807
|
1093 |
fun make_constants intr =
|
webertj@14456
|
1094 |
let
|
webertj@14350
|
1095 |
(* returns a list with all unit vectors of length n *)
|
webertj@14456
|
1096 |
(* int -> interpretation list *)
|
webertj@14350
|
1097 |
fun unit_vectors n =
|
webertj@14350
|
1098 |
let
|
webertj@14350
|
1099 |
(* returns the k-th unit vector of length n *)
|
webertj@14456
|
1100 |
(* int * int -> interpretation *)
|
webertj@14350
|
1101 |
fun unit_vector (k,n) =
|
webertj@14350
|
1102 |
Leaf ((replicate (k-1) False) @ (True :: (replicate (n-k) False)))
|
webertj@14456
|
1103 |
(* int -> interpretation list -> interpretation list *)
|
webertj@14350
|
1104 |
fun unit_vectors_acc k vs =
|
webertj@14350
|
1105 |
if k>n then [] else (unit_vector (k,n))::(unit_vectors_acc (k+1) vs)
|
webertj@14350
|
1106 |
in
|
webertj@14350
|
1107 |
unit_vectors_acc 1 []
|
webertj@14350
|
1108 |
end
|
webertj@14350
|
1109 |
(* concatenates 'x' with every list in 'xss', returning a new list of lists *)
|
webertj@14350
|
1110 |
(* 'a -> 'a list list -> 'a list list *)
|
webertj@14350
|
1111 |
fun cons_list x xss =
|
webertj@14350
|
1112 |
map (fn xs => x::xs) xss
|
webertj@14350
|
1113 |
(* returns a list of lists, each one consisting of n (possibly identical) elements from 'xs' *)
|
webertj@14350
|
1114 |
(* int -> 'a list -> 'a list list *)
|
webertj@14350
|
1115 |
fun pick_all 1 xs =
|
webertj@14350
|
1116 |
map (fn x => [x]) xs
|
webertj@14350
|
1117 |
| pick_all n xs =
|
webertj@14350
|
1118 |
let val rec_pick = pick_all (n-1) xs in
|
webertj@15611
|
1119 |
Library.foldl (fn (acc, x) => (cons_list x rec_pick) @ acc) ([], xs)
|
webertj@14350
|
1120 |
end
|
webertj@14807
|
1121 |
in
|
webertj@14807
|
1122 |
case intr of
|
webertj@14807
|
1123 |
Leaf xs => unit_vectors (length xs)
|
webertj@14807
|
1124 |
| Node xs => map (fn xs' => Node xs') (pick_all (length xs) (make_constants (hd xs)))
|
webertj@14807
|
1125 |
end;
|
webertj@14807
|
1126 |
|
webertj@14807
|
1127 |
(* ------------------------------------------------------------------------- *)
|
webertj@14807
|
1128 |
(* size_of_type: returns the number of constants in a type (i.e. 'length *)
|
webertj@14807
|
1129 |
(* (make_constants intr)', but implemented more efficiently) *)
|
webertj@14807
|
1130 |
(* ------------------------------------------------------------------------- *)
|
webertj@14807
|
1131 |
|
webertj@14807
|
1132 |
(* interpretation -> int *)
|
webertj@14807
|
1133 |
|
webertj@14807
|
1134 |
fun size_of_type intr =
|
webertj@14807
|
1135 |
let
|
webertj@15611
|
1136 |
(* power (a, b) computes a^b, for a>=0, b>=0 *)
|
webertj@14807
|
1137 |
(* int * int -> int *)
|
webertj@15611
|
1138 |
fun power (a, 0) = 1
|
webertj@15611
|
1139 |
| power (a, 1) = a
|
webertj@15611
|
1140 |
| power (a, b) = let val ab = power(a, b div 2) in ab * ab * power(a, b mod 2) end
|
webertj@14807
|
1141 |
in
|
webertj@14807
|
1142 |
case intr of
|
webertj@14807
|
1143 |
Leaf xs => length xs
|
webertj@14807
|
1144 |
| Node xs => power (size_of_type (hd xs), length xs)
|
webertj@14807
|
1145 |
end;
|
webertj@14807
|
1146 |
|
webertj@14807
|
1147 |
(* ------------------------------------------------------------------------- *)
|
webertj@14807
|
1148 |
(* TT/FF: interpretations that denote "true" or "false", respectively *)
|
webertj@14807
|
1149 |
(* ------------------------------------------------------------------------- *)
|
webertj@14807
|
1150 |
|
webertj@14807
|
1151 |
(* interpretation *)
|
webertj@14807
|
1152 |
|
webertj@14807
|
1153 |
val TT = Leaf [True, False];
|
webertj@14807
|
1154 |
|
webertj@14807
|
1155 |
val FF = Leaf [False, True];
|
webertj@14807
|
1156 |
|
webertj@14807
|
1157 |
(* ------------------------------------------------------------------------- *)
|
webertj@14807
|
1158 |
(* make_equality: returns an interpretation that denotes (extensional) *)
|
webertj@14807
|
1159 |
(* equality of two interpretations *)
|
webertj@15547
|
1160 |
(* - two interpretations are 'equal' iff they are both defined and denote *)
|
webertj@15547
|
1161 |
(* the same value *)
|
webertj@15547
|
1162 |
(* - two interpretations are 'not_equal' iff they are both defined at least *)
|
webertj@15547
|
1163 |
(* partially, and a defined part denotes different values *)
|
webertj@15547
|
1164 |
(* - a completely undefined interpretation is neither 'equal' nor *)
|
webertj@15547
|
1165 |
(* 'not_equal' to another interpretation *)
|
webertj@14807
|
1166 |
(* ------------------------------------------------------------------------- *)
|
webertj@14807
|
1167 |
|
webertj@14807
|
1168 |
(* We could in principle represent '=' on a type T by a particular *)
|
webertj@14807
|
1169 |
(* interpretation. However, the size of that interpretation is quadratic *)
|
webertj@14807
|
1170 |
(* in the size of T. Therefore comparing the interpretations 'i1' and *)
|
webertj@14807
|
1171 |
(* 'i2' directly is more efficient than constructing the interpretation *)
|
webertj@14807
|
1172 |
(* for equality on T first, and "applying" this interpretation to 'i1' *)
|
webertj@14807
|
1173 |
(* and 'i2' in the usual way (cf. 'interpretation_apply') then. *)
|
webertj@14807
|
1174 |
|
webertj@14807
|
1175 |
(* interpretation * interpretation -> interpretation *)
|
webertj@14807
|
1176 |
|
webertj@14807
|
1177 |
fun make_equality (i1, i2) =
|
webertj@14807
|
1178 |
let
|
webertj@14807
|
1179 |
(* interpretation * interpretation -> prop_formula *)
|
webertj@14807
|
1180 |
fun equal (i1, i2) =
|
webertj@14807
|
1181 |
(case i1 of
|
webertj@14807
|
1182 |
Leaf xs =>
|
webertj@14807
|
1183 |
(case i2 of
|
webertj@15547
|
1184 |
Leaf ys => PropLogic.dot_product (xs, ys) (* defined and equal *)
|
webertj@14807
|
1185 |
| Node _ => raise REFUTE ("make_equality", "second interpretation is higher"))
|
webertj@14807
|
1186 |
| Node xs =>
|
webertj@14807
|
1187 |
(case i2 of
|
webertj@14807
|
1188 |
Leaf _ => raise REFUTE ("make_equality", "first interpretation is higher")
|
webertj@14807
|
1189 |
| Node ys => PropLogic.all (map equal (xs ~~ ys))))
|
webertj@14807
|
1190 |
(* interpretation * interpretation -> prop_formula *)
|
webertj@14807
|
1191 |
fun not_equal (i1, i2) =
|
webertj@14807
|
1192 |
(case i1 of
|
webertj@14807
|
1193 |
Leaf xs =>
|
webertj@14807
|
1194 |
(case i2 of
|
webertj@14807
|
1195 |
Leaf ys => PropLogic.all ((PropLogic.exists xs) :: (PropLogic.exists ys) ::
|
webertj@14807
|
1196 |
(map (fn (x,y) => SOr (SNot x, SNot y)) (xs ~~ ys))) (* defined and not equal *)
|
webertj@14807
|
1197 |
| Node _ => raise REFUTE ("make_equality", "second interpretation is higher"))
|
webertj@14807
|
1198 |
| Node xs =>
|
webertj@14807
|
1199 |
(case i2 of
|
webertj@14807
|
1200 |
Leaf _ => raise REFUTE ("make_equality", "first interpretation is higher")
|
webertj@14807
|
1201 |
| Node ys => PropLogic.exists (map not_equal (xs ~~ ys))))
|
webertj@14350
|
1202 |
in
|
webertj@14807
|
1203 |
(* a value may be undefined; therefore 'not_equal' is not just the *)
|
webertj@15547
|
1204 |
(* negation of 'equal' *)
|
webertj@14807
|
1205 |
Leaf [equal (i1, i2), not_equal (i1, i2)]
|
webertj@14807
|
1206 |
end;
|
webertj@14807
|
1207 |
|
webertj@15292
|
1208 |
(* ------------------------------------------------------------------------- *)
|
webertj@15547
|
1209 |
(* make_def_equality: returns an interpretation that denotes (extensional) *)
|
webertj@15547
|
1210 |
(* equality of two interpretations *)
|
webertj@15547
|
1211 |
(* This function treats undefined/partially defined interpretations *)
|
webertj@15547
|
1212 |
(* different from 'make_equality': two undefined interpretations are *)
|
webertj@15547
|
1213 |
(* considered equal, while a defined interpretation is considered not equal *)
|
webertj@15547
|
1214 |
(* to an undefined interpretation. *)
|
webertj@15547
|
1215 |
(* ------------------------------------------------------------------------- *)
|
webertj@15547
|
1216 |
|
webertj@15547
|
1217 |
(* interpretation * interpretation -> interpretation *)
|
webertj@15547
|
1218 |
|
webertj@15547
|
1219 |
fun make_def_equality (i1, i2) =
|
webertj@15547
|
1220 |
let
|
webertj@15547
|
1221 |
(* interpretation * interpretation -> prop_formula *)
|
webertj@15547
|
1222 |
fun equal (i1, i2) =
|
webertj@15547
|
1223 |
(case i1 of
|
webertj@15547
|
1224 |
Leaf xs =>
|
webertj@15547
|
1225 |
(case i2 of
|
webertj@15547
|
1226 |
Leaf ys => SOr (PropLogic.dot_product (xs, ys), (* defined and equal, or both undefined *)
|
webertj@15547
|
1227 |
SAnd (PropLogic.all (map SNot xs), PropLogic.all (map SNot ys)))
|
webertj@15547
|
1228 |
| Node _ => raise REFUTE ("make_def_equality", "second interpretation is higher"))
|
webertj@15547
|
1229 |
| Node xs =>
|
webertj@15547
|
1230 |
(case i2 of
|
webertj@15547
|
1231 |
Leaf _ => raise REFUTE ("make_def_equality", "first interpretation is higher")
|
webertj@15547
|
1232 |
| Node ys => PropLogic.all (map equal (xs ~~ ys))))
|
webertj@15547
|
1233 |
(* interpretation *)
|
webertj@15547
|
1234 |
val eq = equal (i1, i2)
|
webertj@15547
|
1235 |
in
|
webertj@15547
|
1236 |
Leaf [eq, SNot eq]
|
webertj@15547
|
1237 |
end;
|
webertj@15547
|
1238 |
|
webertj@15547
|
1239 |
(* ------------------------------------------------------------------------- *)
|
webertj@15547
|
1240 |
(* interpretation_apply: returns an interpretation that denotes the result *)
|
webertj@15547
|
1241 |
(* of applying the function denoted by 'i2' to the *)
|
webertj@15547
|
1242 |
(* argument denoted by 'i2' *)
|
webertj@15547
|
1243 |
(* ------------------------------------------------------------------------- *)
|
webertj@15547
|
1244 |
|
webertj@15547
|
1245 |
(* interpretation * interpretation -> interpretation *)
|
webertj@15547
|
1246 |
|
webertj@15547
|
1247 |
fun interpretation_apply (i1, i2) =
|
webertj@15547
|
1248 |
let
|
webertj@15547
|
1249 |
(* interpretation * interpretation -> interpretation *)
|
webertj@15547
|
1250 |
fun interpretation_disjunction (tr1,tr2) =
|
webertj@15547
|
1251 |
tree_map (fn (xs,ys) => map (fn (x,y) => SOr(x,y)) (xs ~~ ys)) (tree_pair (tr1,tr2))
|
webertj@15547
|
1252 |
(* prop_formula * interpretation -> interpretation *)
|
webertj@15547
|
1253 |
fun prop_formula_times_interpretation (fm,tr) =
|
webertj@15547
|
1254 |
tree_map (map (fn x => SAnd (fm,x))) tr
|
webertj@15547
|
1255 |
(* prop_formula list * interpretation list -> interpretation *)
|
webertj@15547
|
1256 |
fun prop_formula_list_dot_product_interpretation_list ([fm],[tr]) =
|
webertj@15547
|
1257 |
prop_formula_times_interpretation (fm,tr)
|
webertj@15547
|
1258 |
| prop_formula_list_dot_product_interpretation_list (fm::fms,tr::trees) =
|
webertj@15547
|
1259 |
interpretation_disjunction (prop_formula_times_interpretation (fm,tr), prop_formula_list_dot_product_interpretation_list (fms,trees))
|
webertj@15547
|
1260 |
| prop_formula_list_dot_product_interpretation_list (_,_) =
|
webertj@15547
|
1261 |
raise REFUTE ("interpretation_apply", "empty list (in dot product)")
|
webertj@15547
|
1262 |
(* concatenates 'x' with every list in 'xss', returning a new list of lists *)
|
webertj@15547
|
1263 |
(* 'a -> 'a list list -> 'a list list *)
|
webertj@15547
|
1264 |
fun cons_list x xss =
|
webertj@15547
|
1265 |
map (fn xs => x::xs) xss
|
webertj@15547
|
1266 |
(* returns a list of lists, each one consisting of one element from each element of 'xss' *)
|
webertj@15547
|
1267 |
(* 'a list list -> 'a list list *)
|
webertj@15547
|
1268 |
fun pick_all [xs] =
|
webertj@15547
|
1269 |
map (fn x => [x]) xs
|
webertj@15547
|
1270 |
| pick_all (xs::xss) =
|
webertj@15547
|
1271 |
let val rec_pick = pick_all xss in
|
webertj@15611
|
1272 |
Library.foldl (fn (acc, x) => (cons_list x rec_pick) @ acc) ([], xs)
|
webertj@15547
|
1273 |
end
|
webertj@15547
|
1274 |
| pick_all _ =
|
webertj@15547
|
1275 |
raise REFUTE ("interpretation_apply", "empty list (in pick_all)")
|
webertj@15547
|
1276 |
(* interpretation -> prop_formula list *)
|
webertj@15547
|
1277 |
fun interpretation_to_prop_formula_list (Leaf xs) =
|
webertj@15547
|
1278 |
xs
|
webertj@15547
|
1279 |
| interpretation_to_prop_formula_list (Node trees) =
|
webertj@15547
|
1280 |
map PropLogic.all (pick_all (map interpretation_to_prop_formula_list trees))
|
webertj@15547
|
1281 |
in
|
webertj@15547
|
1282 |
case i1 of
|
webertj@15547
|
1283 |
Leaf _ =>
|
webertj@15547
|
1284 |
raise REFUTE ("interpretation_apply", "first interpretation is a leaf")
|
webertj@15547
|
1285 |
| Node xs =>
|
webertj@15547
|
1286 |
prop_formula_list_dot_product_interpretation_list (interpretation_to_prop_formula_list i2, xs)
|
webertj@15547
|
1287 |
end;
|
webertj@15547
|
1288 |
|
webertj@15547
|
1289 |
(* ------------------------------------------------------------------------- *)
|
webertj@15292
|
1290 |
(* eta_expand: eta-expands a term 't' by adding 'i' lambda abstractions *)
|
webertj@15292
|
1291 |
(* ------------------------------------------------------------------------- *)
|
webertj@15292
|
1292 |
|
webertj@15292
|
1293 |
(* Term.term -> int -> Term.term *)
|
webertj@15292
|
1294 |
|
webertj@15292
|
1295 |
fun eta_expand t i =
|
webertj@15292
|
1296 |
let
|
webertj@15292
|
1297 |
val Ts = binder_types (fastype_of t)
|
webertj@15292
|
1298 |
in
|
skalberg@15574
|
1299 |
foldr (fn (T, t) => Abs ("<eta_expand>", T, t))
|
skalberg@15574
|
1300 |
(list_comb (t, map Bound (i-1 downto 0))) (Library.take (i, Ts))
|
webertj@15292
|
1301 |
end;
|
webertj@15292
|
1302 |
|
webertj@15335
|
1303 |
(* ------------------------------------------------------------------------- *)
|
webertj@15335
|
1304 |
(* sum: returns the sum of a list 'xs' of integers *)
|
webertj@15335
|
1305 |
(* ------------------------------------------------------------------------- *)
|
webertj@15335
|
1306 |
|
webertj@15335
|
1307 |
(* int list -> int *)
|
webertj@15335
|
1308 |
|
webertj@15611
|
1309 |
fun sum xs = foldl op+ 0 xs;
|
webertj@15335
|
1310 |
|
webertj@15335
|
1311 |
(* ------------------------------------------------------------------------- *)
|
webertj@15335
|
1312 |
(* product: returns the product of a list 'xs' of integers *)
|
webertj@15335
|
1313 |
(* ------------------------------------------------------------------------- *)
|
webertj@15335
|
1314 |
|
webertj@15335
|
1315 |
(* int list -> int *)
|
webertj@15335
|
1316 |
|
webertj@15611
|
1317 |
fun product xs = foldl op* 1 xs;
|
webertj@15335
|
1318 |
|
webertj@15335
|
1319 |
(* ------------------------------------------------------------------------- *)
|
webertj@15547
|
1320 |
(* size_of_dtyp: the size of (an initial fragment of) an inductive data type *)
|
webertj@15547
|
1321 |
(* is the sum (over its constructors) of the product (over *)
|
webertj@15547
|
1322 |
(* their arguments) of the size of the argument types *)
|
webertj@15335
|
1323 |
(* ------------------------------------------------------------------------- *)
|
webertj@15335
|
1324 |
|
webertj@15335
|
1325 |
(* theory -> (Term.typ * int) list -> DatatypeAux.descr -> (DatatypeAux.dtyp * Term.typ) list -> (string * DatatypeAux.dtyp list) list -> int *)
|
webertj@15335
|
1326 |
|
webertj@15335
|
1327 |
fun size_of_dtyp thy typ_sizes descr typ_assoc constructors =
|
webertj@15335
|
1328 |
sum (map (fn (_, dtyps) =>
|
webertj@15335
|
1329 |
product (map (fn dtyp =>
|
webertj@15335
|
1330 |
let
|
webertj@15335
|
1331 |
val T = typ_of_dtyp descr typ_assoc dtyp
|
webertj@15547
|
1332 |
val (i, _, _) = interpret thy (typ_sizes, []) {maxvars=0, def_eq = false, next_idx=1, bounds=[], wellformed=True} (Free ("dummy", T))
|
webertj@15335
|
1333 |
in
|
webertj@15335
|
1334 |
size_of_type i
|
webertj@15335
|
1335 |
end) dtyps)) constructors);
|
webertj@15335
|
1336 |
|
webertj@15292
|
1337 |
|
webertj@15292
|
1338 |
(* ------------------------------------------------------------------------- *)
|
webertj@15292
|
1339 |
(* INTERPRETERS: Actual Interpreters *)
|
webertj@15292
|
1340 |
(* ------------------------------------------------------------------------- *)
|
webertj@14807
|
1341 |
|
webertj@14807
|
1342 |
(* theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option *)
|
webertj@14807
|
1343 |
|
webertj@14807
|
1344 |
(* simply typed lambda calculus: Isabelle's basic term syntax, with type *)
|
webertj@14807
|
1345 |
(* variables, function types, and propT *)
|
webertj@14807
|
1346 |
|
webertj@14807
|
1347 |
fun stlc_interpreter thy model args t =
|
webertj@14807
|
1348 |
let
|
webertj@15547
|
1349 |
val (typs, terms) = model
|
webertj@15547
|
1350 |
val {maxvars, def_eq, next_idx, bounds, wellformed} = args
|
webertj@14807
|
1351 |
(* Term.typ -> (interpretation * model * arguments) option *)
|
webertj@14807
|
1352 |
fun interpret_groundterm T =
|
webertj@14807
|
1353 |
let
|
webertj@14807
|
1354 |
(* unit -> (interpretation * model * arguments) option *)
|
webertj@14807
|
1355 |
fun interpret_groundtype () =
|
webertj@14807
|
1356 |
let
|
haftmann@17314
|
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 *)
|
webertj@15547
|
1358 |
val next = next_idx+size
|
webertj@15547
|
1359 |
val _ = (if next-1>maxvars andalso maxvars>0 then raise MAXVARS_EXCEEDED else ()) (* check if 'maxvars' is large enough *)
|
webertj@14807
|
1360 |
(* prop_formula list *)
|
webertj@15547
|
1361 |
val fms = map BoolVar (next_idx upto (next_idx+size-1))
|
webertj@14807
|
1362 |
(* interpretation *)
|
webertj@14807
|
1363 |
val intr = Leaf fms
|
webertj@14807
|
1364 |
(* prop_formula list -> prop_formula *)
|
webertj@14807
|
1365 |
fun one_of_two_false [] = True
|
webertj@14807
|
1366 |
| one_of_two_false (x::xs) = SAnd (PropLogic.all (map (fn x' => SOr (SNot x, SNot x')) xs), one_of_two_false xs)
|
webertj@14807
|
1367 |
(* prop_formula *)
|
webertj@15547
|
1368 |
val wf = one_of_two_false fms
|
webertj@14807
|
1369 |
in
|
webertj@14807
|
1370 |
(* extend the model, increase 'next_idx', add well-formedness condition *)
|
webertj@15547
|
1371 |
SOME (intr, (typs, (t, intr)::terms), {maxvars = maxvars, def_eq = def_eq, next_idx = next, bounds = bounds, wellformed = SAnd (wellformed, wf)})
|
webertj@14807
|
1372 |
end
|
webertj@14807
|
1373 |
in
|
webertj@14807
|
1374 |
case T of
|
webertj@14807
|
1375 |
Type ("fun", [T1, T2]) =>
|
webertj@14807
|
1376 |
let
|
webertj@14807
|
1377 |
(* we create 'size_of_type (interpret (... T1))' different copies *)
|
webertj@14807
|
1378 |
(* of the interpretation for 'T2', which are then combined into a *)
|
webertj@14807
|
1379 |
(* single new interpretation *)
|
webertj@15547
|
1380 |
val (i1, _, _) = interpret thy model {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True} (Free ("dummy", T1))
|
webertj@14807
|
1381 |
(* make fresh copies, with different variable indices *)
|
webertj@14807
|
1382 |
(* 'idx': next variable index *)
|
webertj@14807
|
1383 |
(* 'n' : number of copies *)
|
webertj@14807
|
1384 |
(* int -> int -> (int * interpretation list * prop_formula *)
|
webertj@14807
|
1385 |
fun make_copies idx 0 =
|
webertj@14807
|
1386 |
(idx, [], True)
|
webertj@14807
|
1387 |
| make_copies idx n =
|
webertj@14807
|
1388 |
let
|
webertj@15547
|
1389 |
val (copy, _, new_args) = interpret thy (typs, []) {maxvars = maxvars, def_eq = false, next_idx = idx, bounds = [], wellformed = True} (Free ("dummy", T2))
|
webertj@14807
|
1390 |
val (idx', copies, wf') = make_copies (#next_idx new_args) (n-1)
|
webertj@14807
|
1391 |
in
|
webertj@14807
|
1392 |
(idx', copy :: copies, SAnd (#wellformed new_args, wf'))
|
webertj@14807
|
1393 |
end
|
webertj@14807
|
1394 |
val (next, copies, wf) = make_copies next_idx (size_of_type i1)
|
webertj@14807
|
1395 |
(* combine copies into a single interpretation *)
|
webertj@14807
|
1396 |
val intr = Node copies
|
webertj@14807
|
1397 |
in
|
webertj@14807
|
1398 |
(* extend the model, increase 'next_idx', add well-formedness condition *)
|
webertj@15547
|
1399 |
SOME (intr, (typs, (t, intr)::terms), {maxvars = maxvars, def_eq = def_eq, next_idx = next, bounds = bounds, wellformed = SAnd (wellformed, wf)})
|
webertj@14807
|
1400 |
end
|
webertj@14807
|
1401 |
| Type _ => interpret_groundtype ()
|
webertj@14807
|
1402 |
| TFree _ => interpret_groundtype ()
|
webertj@14807
|
1403 |
| TVar _ => interpret_groundtype ()
|
webertj@14807
|
1404 |
end
|
webertj@14807
|
1405 |
in
|
haftmann@17314
|
1406 |
case AList.lookup (op =) terms t of
|
skalberg@15531
|
1407 |
SOME intr =>
|
webertj@14807
|
1408 |
(* return an existing interpretation *)
|
skalberg@15531
|
1409 |
SOME (intr, model, args)
|
skalberg@15531
|
1410 |
| NONE =>
|
webertj@14807
|
1411 |
(case t of
|
webertj@14807
|
1412 |
Const (_, T) =>
|
webertj@14807
|
1413 |
interpret_groundterm T
|
webertj@14807
|
1414 |
| Free (_, T) =>
|
webertj@14807
|
1415 |
interpret_groundterm T
|
webertj@14807
|
1416 |
| Var (_, T) =>
|
webertj@14807
|
1417 |
interpret_groundterm T
|
webertj@14807
|
1418 |
| Bound i =>
|
skalberg@15570
|
1419 |
SOME (List.nth (#bounds args, i), model, args)
|
webertj@14807
|
1420 |
| Abs (x, T, body) =>
|
webertj@14807
|
1421 |
let
|
webertj@14807
|
1422 |
(* create all constants of type 'T' *)
|
webertj@15547
|
1423 |
val (i, _, _) = interpret thy model {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True} (Free ("dummy", T))
|
webertj@14807
|
1424 |
val constants = make_constants i
|
webertj@14807
|
1425 |
(* interpret the 'body' separately for each constant *)
|
webertj@14807
|
1426 |
val ((model', args'), bodies) = foldl_map
|
webertj@15547
|
1427 |
(fn ((m, a), c) =>
|
webertj@14807
|
1428 |
let
|
webertj@14807
|
1429 |
(* add 'c' to 'bounds' *)
|
webertj@15547
|
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
|
webertj@14807
|
1431 |
in
|
webertj@14807
|
1432 |
(* keep the new model m' and 'next_idx' and 'wellformed', but use old 'bounds' *)
|
webertj@15547
|
1433 |
((m', {maxvars = maxvars, def_eq = def_eq, next_idx = #next_idx a', bounds = bounds, wellformed = #wellformed a'}), i')
|
webertj@14807
|
1434 |
end)
|
webertj@14807
|
1435 |
((model, args), constants)
|
webertj@14807
|
1436 |
in
|
skalberg@15531
|
1437 |
SOME (Node bodies, model', args')
|
webertj@14807
|
1438 |
end
|
webertj@14807
|
1439 |
| t1 $ t2 =>
|
webertj@14807
|
1440 |
let
|
webertj@14807
|
1441 |
(* interpret 't1' and 't2' separately *)
|
webertj@14807
|
1442 |
val (intr1, model1, args1) = interpret thy model args t1
|
webertj@14807
|
1443 |
val (intr2, model2, args2) = interpret thy model1 args1 t2
|
webertj@14807
|
1444 |
in
|
webertj@15547
|
1445 |
SOME (interpretation_apply (intr1, intr2), model2, args2)
|
webertj@14807
|
1446 |
end)
|
webertj@14807
|
1447 |
end;
|
webertj@14807
|
1448 |
|
webertj@14807
|
1449 |
(* theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option *)
|
webertj@14807
|
1450 |
|
webertj@14807
|
1451 |
fun Pure_interpreter thy model args t =
|
webertj@14456
|
1452 |
case t of
|
webertj@14807
|
1453 |
Const ("all", _) $ t1 => (* in the meta-logic, 'all' MUST be followed by an argument term *)
|
webertj@14807
|
1454 |
let
|
webertj@14807
|
1455 |
val (i, m, a) = interpret thy model args t1
|
webertj@14807
|
1456 |
in
|
webertj@14807
|
1457 |
case i of
|
webertj@14807
|
1458 |
Node xs =>
|
webertj@14807
|
1459 |
let
|
webertj@14807
|
1460 |
val fmTrue = PropLogic.all (map toTrue xs)
|
webertj@14807
|
1461 |
val fmFalse = PropLogic.exists (map toFalse xs)
|
webertj@14807
|
1462 |
in
|
skalberg@15531
|
1463 |
SOME (Leaf [fmTrue, fmFalse], m, a)
|
webertj@14807
|
1464 |
end
|
webertj@14807
|
1465 |
| _ =>
|
webertj@14807
|
1466 |
raise REFUTE ("Pure_interpreter", "\"all\" is not followed by a function")
|
webertj@14807
|
1467 |
end
|
webertj@14807
|
1468 |
| Const ("==", _) $ t1 $ t2 =>
|
webertj@14807
|
1469 |
let
|
webertj@14807
|
1470 |
val (i1, m1, a1) = interpret thy model args t1
|
webertj@14807
|
1471 |
val (i2, m2, a2) = interpret thy m1 a1 t2
|
webertj@14807
|
1472 |
in
|
webertj@15547
|
1473 |
(* we use either 'make_def_equality' or 'make_equality' *)
|
webertj@15547
|
1474 |
SOME ((if #def_eq args then make_def_equality else make_equality) (i1, i2), m2, a2)
|
webertj@14807
|
1475 |
end
|
webertj@14807
|
1476 |
| Const ("==>", _) => (* simpler than translating 'Const ("==>", _) $ t1 $ t2' *)
|
skalberg@15531
|
1477 |
SOME (Node [Node [TT, FF], Node [TT, TT]], model, args)
|
skalberg@15531
|
1478 |
| _ => NONE;
|
webertj@14807
|
1479 |
|
webertj@14807
|
1480 |
(* theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option *)
|
webertj@14807
|
1481 |
|
webertj@14807
|
1482 |
fun HOLogic_interpreter thy model args t =
|
webertj@14807
|
1483 |
(* ------------------------------------------------------------------------- *)
|
webertj@14807
|
1484 |
(* Providing interpretations directly is more efficient than unfolding the *)
|
webertj@15283
|
1485 |
(* logical constants. In HOL however, logical constants can themselves be *)
|
webertj@14807
|
1486 |
(* arguments. "All" and "Ex" are then translated just like any other *)
|
webertj@14807
|
1487 |
(* constant, with the relevant axiom being added by 'collect_axioms'. *)
|
webertj@14807
|
1488 |
(* ------------------------------------------------------------------------- *)
|
webertj@14807
|
1489 |
case t of
|
webertj@14807
|
1490 |
Const ("Trueprop", _) =>
|
skalberg@15531
|
1491 |
SOME (Node [TT, FF], model, args)
|
webertj@14807
|
1492 |
| Const ("Not", _) =>
|
skalberg@15531
|
1493 |
SOME (Node [FF, TT], model, args)
|
webertj@14807
|
1494 |
| Const ("True", _) => (* redundant, since 'True' is also an IDT constructor *)
|
skalberg@15531
|
1495 |
SOME (TT, model, args)
|
webertj@14807
|
1496 |
| Const ("False", _) => (* redundant, since 'False' is also an IDT constructor *)
|
skalberg@15531
|
1497 |
SOME (FF, model, args)
|
webertj@15333
|
1498 |
| Const ("All", _) $ t1 =>
|
webertj@15333
|
1499 |
(* if "All" occurs without an argument (i.e. as argument to a higher-order *)
|
webertj@15547
|
1500 |
(* function or predicate), it is handled by the 'stlc_interpreter' (i.e. *)
|
webertj@15333
|
1501 |
(* by unfolding its definition) *)
|
webertj@14350
|
1502 |
let
|
webertj@14807
|
1503 |
val (i, m, a) = interpret thy model args t1
|
webertj@14807
|
1504 |
in
|
webertj@14807
|
1505 |
case i of
|
webertj@14807
|
1506 |
Node xs =>
|
webertj@14807
|
1507 |
let
|
webertj@14807
|
1508 |
val fmTrue = PropLogic.all (map toTrue xs)
|
webertj@14807
|
1509 |
val fmFalse = PropLogic.exists (map toFalse xs)
|
webertj@14807
|
1510 |
in
|
skalberg@15531
|
1511 |
SOME (Leaf [fmTrue, fmFalse], m, a)
|
webertj@14807
|
1512 |
end
|
webertj@14807
|
1513 |
| _ =>
|
webertj@15292
|
1514 |
raise REFUTE ("HOLogic_interpreter", "\"All\" is followed by a non-function")
|
webertj@14807
|
1515 |
end
|
webertj@15333
|
1516 |
| Const ("Ex", _) $ t1 =>
|
webertj@15333
|
1517 |
(* if "Ex" occurs without an argument (i.e. as argument to a higher-order *)
|
webertj@15547
|
1518 |
(* function or predicate), it is handled by the 'stlc_interpreter' (i.e. *)
|
webertj@15333
|
1519 |
(* by unfolding its definition) *)
|
webertj@14807
|
1520 |
let
|
webertj@14807
|
1521 |
val (i, m, a) = interpret thy model args t1
|
webertj@14807
|
1522 |
in
|
webertj@14807
|
1523 |
case i of
|
webertj@14807
|
1524 |
Node xs =>
|
webertj@14807
|
1525 |
let
|
webertj@14807
|
1526 |
val fmTrue = PropLogic.exists (map toTrue xs)
|
webertj@14807
|
1527 |
val fmFalse = PropLogic.all (map toFalse xs)
|
webertj@14807
|
1528 |
in
|
skalberg@15531
|
1529 |
SOME (Leaf [fmTrue, fmFalse], m, a)
|
webertj@14807
|
1530 |
end
|
webertj@14807
|
1531 |
| _ =>
|
webertj@15292
|
1532 |
raise REFUTE ("HOLogic_interpreter", "\"Ex\" is followed by a non-function")
|
webertj@14807
|
1533 |
end
|
webertj@14807
|
1534 |
| Const ("op =", _) $ t1 $ t2 =>
|
webertj@14807
|
1535 |
let
|
webertj@14807
|
1536 |
val (i1, m1, a1) = interpret thy model args t1
|
webertj@14807
|
1537 |
val (i2, m2, a2) = interpret thy m1 a1 t2
|
webertj@14807
|
1538 |
in
|
skalberg@15531
|
1539 |
SOME (make_equality (i1, i2), m2, a2)
|
webertj@14807
|
1540 |
end
|
webertj@14807
|
1541 |
| Const ("op =", _) $ t1 =>
|
skalberg@15531
|
1542 |
SOME (interpret thy model args (eta_expand t 1))
|
webertj@14807
|
1543 |
| Const ("op =", _) =>
|
skalberg@15531
|
1544 |
SOME (interpret thy model args (eta_expand t 2))
|
webertj@15547
|
1545 |
| Const ("op &", _) $ t1 $ t2 =>
|
webertj@15547
|
1546 |
(* 3-valued logic *)
|
webertj@15547
|
1547 |
let
|
webertj@15547
|
1548 |
val (i1, m1, a1) = interpret thy model args t1
|
webertj@15547
|
1549 |
val (i2, m2, a2) = interpret thy m1 a1 t2
|
webertj@15547
|
1550 |
val fmTrue = PropLogic.SAnd (toTrue i1, toTrue i2)
|
webertj@15547
|
1551 |
val fmFalse = PropLogic.SOr (toFalse i1, toFalse i2)
|
webertj@15547
|
1552 |
in
|
webertj@15547
|
1553 |
SOME (Leaf [fmTrue, fmFalse], m2, a2)
|
webertj@15547
|
1554 |
end
|
webertj@15547
|
1555 |
| Const ("op &", _) $ t1 =>
|
webertj@15547
|
1556 |
SOME (interpret thy model args (eta_expand t 1))
|
webertj@14807
|
1557 |
| Const ("op &", _) =>
|
webertj@15547
|
1558 |
SOME (interpret thy model args (eta_expand t 2))
|
webertj@15547
|
1559 |
(* SOME (Node [Node [TT, FF], Node [FF, FF]], model, args) *)
|
webertj@15547
|
1560 |
| Const ("op |", _) $ t1 $ t2 =>
|
webertj@15547
|
1561 |
(* 3-valued logic *)
|
webertj@15547
|
1562 |
let
|
webertj@15547
|
1563 |
val (i1, m1, a1) = interpret thy model args t1
|
webertj@15547
|
1564 |
val (i2, m2, a2) = interpret thy m1 a1 t2
|
webertj@15547
|
1565 |
val fmTrue = PropLogic.SOr (toTrue i1, toTrue i2)
|
webertj@15547
|
1566 |
val fmFalse = PropLogic.SAnd (toFalse i1, toFalse i2)
|
webertj@15547
|
1567 |
in
|
webertj@15547
|
1568 |
SOME (Leaf [fmTrue, fmFalse], m2, a2)
|
webertj@15547
|
1569 |
end
|
webertj@15547
|
1570 |
| Const ("op |", _) $ t1 =>
|
webertj@15547
|
1571 |
SOME (interpret thy model args (eta_expand t 1))
|
webertj@14807
|
1572 |
| Const ("op |", _) =>
|
webertj@15547
|
1573 |
SOME (interpret thy model args (eta_expand t 2))
|
webertj@15547
|
1574 |
(* SOME (Node [Node [TT, TT], Node [TT, FF]], model, args) *)
|
webertj@15547
|
1575 |
| Const ("op -->", _) $ t1 $ t2 =>
|
webertj@15547
|
1576 |
(* 3-valued logic *)
|
webertj@15547
|
1577 |
let
|
webertj@15547
|
1578 |
val (i1, m1, a1) = interpret thy model args t1
|
webertj@15547
|
1579 |
val (i2, m2, a2) = interpret thy m1 a1 t2
|
webertj@15547
|
1580 |
val fmTrue = PropLogic.SOr (toFalse i1, toTrue i2)
|
webertj@15547
|
1581 |
val fmFalse = PropLogic.SAnd (toTrue i1, toFalse i2)
|
webertj@15547
|
1582 |
in
|
webertj@15547
|
1583 |
SOME (Leaf [fmTrue, fmFalse], m2, a2)
|
webertj@15547
|
1584 |
end
|
webertj@14807
|
1585 |
| Const ("op -->", _) =>
|
webertj@15547
|
1586 |
(* SOME (Node [Node [TT, FF], Node [TT, TT]], model, args) *)
|
webertj@15547
|
1587 |
SOME (interpret thy model args (eta_expand t 2))
|
skalberg@15531
|
1588 |
| _ => NONE;
|
webertj@14807
|
1589 |
|
webertj@14807
|
1590 |
(* theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option *)
|
webertj@14807
|
1591 |
|
webertj@14807
|
1592 |
fun set_interpreter thy model args t =
|
webertj@14807
|
1593 |
(* "T set" is isomorphic to "T --> bool" *)
|
webertj@14807
|
1594 |
let
|
webertj@14807
|
1595 |
val (typs, terms) = model
|
webertj@14807
|
1596 |
in
|
haftmann@17314
|
1597 |
case AList.lookup (op =) terms t of
|
skalberg@15531
|
1598 |
SOME intr =>
|
webertj@14807
|
1599 |
(* return an existing interpretation *)
|
skalberg@15531
|
1600 |
SOME (intr, model, args)
|
skalberg@15531
|
1601 |
| NONE =>
|
webertj@14807
|
1602 |
(case t of
|
webertj@14807
|
1603 |
Free (x, Type ("set", [T])) =>
|
webertj@15334
|
1604 |
let
|
webertj@14807
|
1605 |
val (intr, _, args') = interpret thy (typs, []) args (Free (x, T --> HOLogic.boolT))
|
webertj@14807
|
1606 |
in
|
skalberg@15531
|
1607 |
SOME (intr, (typs, (t, intr)::terms), args')
|
webertj@15334
|
1608 |
end
|
webertj@15767
|
1609 |
| Var ((x, i), Type ("set", [T])) =>
|
webertj@15334
|
1610 |
let
|
webertj@14807
|
1611 |
val (intr, _, args') = interpret thy (typs, []) args (Var ((x,i), T --> HOLogic.boolT))
|
webertj@14807
|
1612 |
in
|
skalberg@15531
|
1613 |
SOME (intr, (typs, (t, intr)::terms), args')
|
webertj@15334
|
1614 |
end
|
webertj@14807
|
1615 |
| Const (s, Type ("set", [T])) =>
|
webertj@15334
|
1616 |
let
|
webertj@14807
|
1617 |
val (intr, _, args') = interpret thy (typs, []) args (Const (s, T --> HOLogic.boolT))
|
webertj@14807
|
1618 |
in
|
skalberg@15531
|
1619 |
SOME (intr, (typs, (t, intr)::terms), args')
|
webertj@15334
|
1620 |
end
|
webertj@14807
|
1621 |
(* 'Collect' == identity *)
|
webertj@14807
|
1622 |
| Const ("Collect", _) $ t1 =>
|
skalberg@15531
|
1623 |
SOME (interpret thy model args t1)
|
webertj@14807
|
1624 |
| Const ("Collect", _) =>
|
skalberg@15531
|
1625 |
SOME (interpret thy model args (eta_expand t 1))
|
webertj@14807
|
1626 |
(* 'op :' == application *)
|
webertj@14807
|
1627 |
| Const ("op :", _) $ t1 $ t2 =>
|
skalberg@15531
|
1628 |
SOME (interpret thy model args (t2 $ t1))
|
webertj@14807
|
1629 |
| Const ("op :", _) $ t1 =>
|
skalberg@15531
|
1630 |
SOME (interpret thy model args (eta_expand t 1))
|
webertj@14807
|
1631 |
| Const ("op :", _) =>
|
skalberg@15531
|
1632 |
SOME (interpret thy model args (eta_expand t 2))
|
skalberg@15531
|
1633 |
| _ => NONE)
|
webertj@14807
|
1634 |
end;
|
webertj@14807
|
1635 |
|
webertj@14807
|
1636 |
(* theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option *)
|
webertj@14807
|
1637 |
|
webertj@15547
|
1638 |
(* interprets variables and constants whose type is an IDT; constructors of *)
|
webertj@15547
|
1639 |
(* IDTs are properly interpreted by 'IDT_constructor_interpreter' however *)
|
webertj@15547
|
1640 |
|
webertj@14807
|
1641 |
fun IDT_interpreter thy model args t =
|
webertj@14807
|
1642 |
let
|
webertj@14807
|
1643 |
val (typs, terms) = model
|
webertj@14807
|
1644 |
(* Term.typ -> (interpretation * model * arguments) option *)
|
webertj@15547
|
1645 |
fun interpret_term (Type (s, Ts)) =
|
haftmann@19346
|
1646 |
(case DatatypePackage.get_datatype thy s of
|
skalberg@15531
|
1647 |
SOME info => (* inductive datatype *)
|
webertj@14807
|
1648 |
let
|
webertj@14807
|
1649 |
(* int option -- only recursive IDTs have an associated depth *)
|
haftmann@17314
|
1650 |
val depth = AList.lookup (op =) typs (Type (s, Ts))
|
webertj@14807
|
1651 |
in
|
skalberg@15531
|
1652 |
if depth = (SOME 0) then (* termination condition to avoid infinite recursion *)
|
webertj@14807
|
1653 |
(* return a leaf of size 0 *)
|
skalberg@15531
|
1654 |
SOME (Leaf [], model, args)
|
webertj@14807
|
1655 |
else
|
webertj@14807
|
1656 |
let
|
webertj@14807
|
1657 |
val index = #index info
|
webertj@14807
|
1658 |
val descr = #descr info
|
haftmann@17314
|
1659 |
val (_, dtyps, constrs) = (the o AList.lookup (op =) descr) index
|
webertj@14807
|
1660 |
val typ_assoc = dtyps ~~ Ts
|
webertj@14807
|
1661 |
(* sanity check: every element in 'dtyps' must be a 'DtTFree' *)
|
webertj@14807
|
1662 |
val _ = (if Library.exists (fn d =>
|
webertj@14807
|
1663 |
case d of DatatypeAux.DtTFree _ => false | _ => true) dtyps
|
webertj@14807
|
1664 |
then
|
webertj@14807
|
1665 |
raise REFUTE ("IDT_interpreter", "datatype argument (for type " ^ Sign.string_of_typ (sign_of thy) (Type (s, Ts)) ^ ") is not a variable")
|
webertj@14807
|
1666 |
else
|
webertj@14807
|
1667 |
())
|
webertj@14807
|
1668 |
(* if the model specifies a depth for the current type, decrement it to avoid infinite recursion *)
|
haftmann@17314
|
1669 |
val typs' = case depth of NONE => typs | SOME n =>
|
haftmann@17314
|
1670 |
AList.update (op =) (Type (s, Ts), n-1) typs
|
webertj@14807
|
1671 |
(* recursively compute the size of the datatype *)
|
webertj@15335
|
1672 |
val size = size_of_dtyp thy typs' descr typ_assoc constrs
|
webertj@14807
|
1673 |
val next_idx = #next_idx args
|
webertj@15547
|
1674 |
val next = next_idx+size
|
webertj@14807
|
1675 |
val _ = (if next-1>(#maxvars args) andalso (#maxvars args)>0 then raise MAXVARS_EXCEEDED else ()) (* check if 'maxvars' is large enough *)
|
webertj@14807
|
1676 |
(* prop_formula list *)
|
webertj@15547
|
1677 |
val fms = map BoolVar (next_idx upto (next_idx+size-1))
|
webertj@14807
|
1678 |
(* interpretation *)
|
webertj@14807
|
1679 |
val intr = Leaf fms
|
webertj@14807
|
1680 |
(* prop_formula list -> prop_formula *)
|
webertj@14807
|
1681 |
fun one_of_two_false [] = True
|
webertj@14807
|
1682 |
| one_of_two_false (x::xs) = SAnd (PropLogic.all (map (fn x' => SOr (SNot x, SNot x')) xs), one_of_two_false xs)
|
webertj@14807
|
1683 |
(* prop_formula *)
|
webertj@15547
|
1684 |
val wf = one_of_two_false fms
|
webertj@14807
|
1685 |
in
|
webertj@14807
|
1686 |
(* extend the model, increase 'next_idx', add well-formedness condition *)
|
webertj@15547
|
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)})
|
webertj@14807
|
1688 |
end
|
webertj@14807
|
1689 |
end
|
skalberg@15531
|
1690 |
| NONE => (* not an inductive datatype *)
|
skalberg@15531
|
1691 |
NONE)
|
webertj@15547
|
1692 |
| interpret_term _ = (* a (free or schematic) type variable *)
|
skalberg@15531
|
1693 |
NONE
|
webertj@14807
|
1694 |
in
|
haftmann@17314
|
1695 |
case AList.lookup (op =) terms t of
|
skalberg@15531
|
1696 |
SOME intr =>
|
webertj@14807
|
1697 |
(* return an existing interpretation *)
|
skalberg@15531
|
1698 |
SOME (intr, model, args)
|
skalberg@15531
|
1699 |
| NONE =>
|
webertj@14807
|
1700 |
(case t of
|
webertj@15547
|
1701 |
Free (_, T) => interpret_term T
|
webertj@15547
|
1702 |
| Var (_, T) => interpret_term T
|
webertj@15547
|
1703 |
| Const (_, T) => interpret_term T
|
webertj@15547
|
1704 |
| _ => NONE)
|
webertj@15547
|
1705 |
end;
|
webertj@15547
|
1706 |
|
webertj@15547
|
1707 |
(* theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option *)
|
webertj@15547
|
1708 |
|
webertj@15547
|
1709 |
fun IDT_constructor_interpreter thy model args t =
|
webertj@15547
|
1710 |
let
|
webertj@15547
|
1711 |
val (typs, terms) = model
|
webertj@15547
|
1712 |
in
|
haftmann@17314
|
1713 |
case AList.lookup (op =) terms t of
|
webertj@15547
|
1714 |
SOME intr =>
|
webertj@15547
|
1715 |
(* return an existing interpretation *)
|
webertj@15547
|
1716 |
SOME (intr, model, args)
|
webertj@15547
|
1717 |
| NONE =>
|
webertj@15547
|
1718 |
(case t of
|
webertj@15547
|
1719 |
Const (s, T) =>
|
webertj@15547
|
1720 |
(case body_type T of
|
webertj@15547
|
1721 |
Type (s', Ts') =>
|
haftmann@19346
|
1722 |
(case DatatypePackage.get_datatype thy s' of
|
webertj@15547
|
1723 |
SOME info => (* body type is an inductive datatype *)
|
webertj@15547
|
1724 |
let
|
webertj@15547
|
1725 |
val index = #index info
|
webertj@15547
|
1726 |
val descr = #descr info
|
haftmann@17314
|
1727 |
val (_, dtyps, constrs) = (the o AList.lookup (op =) descr) index
|
webertj@15547
|
1728 |
val typ_assoc = dtyps ~~ Ts'
|
webertj@15547
|
1729 |
(* sanity check: every element in 'dtyps' must be a 'DtTFree' *)
|
webertj@15547
|
1730 |
val _ = (if Library.exists (fn d =>
|
webertj@15547
|
1731 |
case d of DatatypeAux.DtTFree _ => false | _ => true) dtyps
|
webertj@15547
|
1732 |
then
|
webertj@15547
|
1733 |
raise REFUTE ("IDT_constructor_interpreter", "datatype argument (for type " ^ Sign.string_of_typ (sign_of thy) (Type (s', Ts')) ^ ") is not a variable")
|
webertj@15547
|
1734 |
else
|
webertj@15547
|
1735 |
())
|
webertj@15547
|
1736 |
(* split the constructors into those occuring before/after 'Const (s, T)' *)
|
webertj@15547
|
1737 |
val (constrs1, constrs2) = take_prefix (fn (cname, ctypes) =>
|
wenzelm@16935
|
1738 |
not (cname = s andalso Sign.typ_instance thy (T,
|
webertj@15547
|
1739 |
map (typ_of_dtyp descr typ_assoc) ctypes ---> Type (s', Ts')))) constrs
|
webertj@15547
|
1740 |
in
|
webertj@15547
|
1741 |
case constrs2 of
|
webertj@15547
|
1742 |
[] =>
|
webertj@15547
|
1743 |
(* 'Const (s, T)' is not a constructor of this datatype *)
|
webertj@15547
|
1744 |
NONE
|
webertj@15547
|
1745 |
| (_, ctypes)::cs =>
|
webertj@14807
|
1746 |
let
|
webertj@15547
|
1747 |
(* compute the total size of the datatype (with the current depth) *)
|
webertj@15547
|
1748 |
val (i, _, _) = interpret thy (typs, []) {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True} (Free ("dummy", Type (s', Ts')))
|
webertj@15547
|
1749 |
val total = size_of_type i
|
webertj@15547
|
1750 |
(* int option -- only recursive IDTs have an associated depth *)
|
haftmann@17314
|
1751 |
val depth = AList.lookup (op =) typs (Type (s', Ts'))
|
haftmann@17314
|
1752 |
val typs' = (case depth of NONE => typs | SOME n =>
|
haftmann@17314
|
1753 |
AList.update (op =) (Type (s', Ts'), n-1) typs)
|
webertj@15611
|
1754 |
(* returns an interpretation where everything is mapped to "undefined" *)
|
webertj@15547
|
1755 |
(* DatatypeAux.dtyp list -> interpretation *)
|
webertj@15547
|
1756 |
fun make_undef [] =
|
webertj@15547
|
1757 |
Leaf (replicate total False)
|
webertj@15547
|
1758 |
| make_undef (d::ds) =
|
webertj@15547
|
1759 |
let
|
webertj@15547
|
1760 |
(* compute the current size of the type 'd' *)
|
webertj@15547
|
1761 |
val T = typ_of_dtyp descr typ_assoc d
|
webertj@15547
|
1762 |
val (i, _, _) = interpret thy (typs, []) {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True} (Free ("dummy", T))
|
webertj@15547
|
1763 |
val size = size_of_type i
|
webertj@15547
|
1764 |
in
|
webertj@15547
|
1765 |
Node (replicate size (make_undef ds))
|
webertj@15547
|
1766 |
end
|
webertj@15547
|
1767 |
(* returns the interpretation for a constructor at depth 1 *)
|
webertj@15547
|
1768 |
(* int * DatatypeAux.dtyp list -> int * interpretation *)
|
webertj@15547
|
1769 |
fun make_constr (offset, []) =
|
webertj@15547
|
1770 |
if offset<total then
|
webertj@15547
|
1771 |
(offset+1, Leaf ((replicate offset False) @ True :: (replicate (total-offset-1) False)))
|
webertj@14807
|
1772 |
else
|
webertj@15547
|
1773 |
raise REFUTE ("IDT_constructor_interpreter", "offset >= total")
|
webertj@15547
|
1774 |
| make_constr (offset, d::ds) =
|
webertj@14807
|
1775 |
let
|
webertj@15547
|
1776 |
(* compute the current and the old size of the type 'd' *)
|
webertj@15547
|
1777 |
val T = typ_of_dtyp descr typ_assoc d
|
webertj@15547
|
1778 |
val (i, _, _) = interpret thy (typs, []) {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True} (Free ("dummy", T))
|
webertj@15547
|
1779 |
val size = size_of_type i
|
webertj@15547
|
1780 |
val (i', _, _) = interpret thy (typs', []) {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True} (Free ("dummy", T))
|
webertj@15547
|
1781 |
val size' = size_of_type i'
|
webertj@15547
|
1782 |
(* sanity check *)
|
webertj@15547
|
1783 |
val _ = if size < size' then
|
webertj@15547
|
1784 |
raise REFUTE ("IDT_constructor_interpreter", "current size is less than old size")
|
webertj@14807
|
1785 |
else
|
webertj@15547
|
1786 |
()
|
webertj@15547
|
1787 |
(* int * interpretation list *)
|
webertj@15547
|
1788 |
val (new_offset, intrs) = foldl_map make_constr (offset, replicate size' ds)
|
webertj@15547
|
1789 |
(* interpretation list *)
|
webertj@15547
|
1790 |
val undefs = replicate (size - size') (make_undef ds)
|
webertj@15547
|
1791 |
in
|
webertj@15611
|
1792 |
(* elements that exist at the previous depth are mapped to a defined *)
|
webertj@15611
|
1793 |
(* value, while new elements are mapped to "undefined" by the *)
|
webertj@15611
|
1794 |
(* recursive constructor *)
|
webertj@15547
|
1795 |
(new_offset, Node (intrs @ undefs))
|
webertj@15547
|
1796 |
end
|
webertj@15547
|
1797 |
(* extends the interpretation for a constructor (both recursive *)
|
webertj@15547
|
1798 |
(* and non-recursive) obtained at depth n (n>=1) to depth n+1 *)
|
webertj@15547
|
1799 |
(* int * DatatypeAux.dtyp list * interpretation -> int * interpretation *)
|
webertj@15547
|
1800 |
fun extend_constr (offset, [], Leaf xs) =
|
webertj@15547
|
1801 |
let
|
webertj@15547
|
1802 |
(* returns the k-th unit vector of length n *)
|
webertj@15547
|
1803 |
(* int * int -> interpretation *)
|
webertj@15611
|
1804 |
fun unit_vector (k, n) =
|
webertj@15547
|
1805 |
Leaf ((replicate (k-1) False) @ (True :: (replicate (n-k) False)))
|
webertj@15547
|
1806 |
(* int *)
|
webertj@15547
|
1807 |
val k = find_index_eq True xs
|
webertj@14807
|
1808 |
in
|
webertj@15547
|
1809 |
if k=(~1) then
|
webertj@15547
|
1810 |
(* if the element was mapped to "undefined" before, map it to *)
|
webertj@15547
|
1811 |
(* the value given by 'offset' now (and extend the length of *)
|
webertj@15547
|
1812 |
(* the leaf) *)
|
webertj@15547
|
1813 |
(offset+1, unit_vector (offset+1, total))
|
webertj@15547
|
1814 |
else
|
webertj@15547
|
1815 |
(* if the element was already mapped to a defined value, map it *)
|
webertj@15547
|
1816 |
(* to the same value again, just extend the length of the leaf, *)
|
webertj@15611
|
1817 |
(* do not increment the 'offset' *)
|
webertj@15547
|
1818 |
(offset, unit_vector (k+1, total))
|
webertj@15547
|
1819 |
end
|
webertj@15547
|
1820 |
| extend_constr (_, [], Node _) =
|
webertj@15547
|
1821 |
raise REFUTE ("IDT_constructor_interpreter", "interpretation for constructor (with no arguments left) is a node")
|
webertj@15547
|
1822 |
| extend_constr (offset, d::ds, Node xs) =
|
webertj@15547
|
1823 |
let
|
webertj@15547
|
1824 |
(* compute the size of the type 'd' *)
|
webertj@15547
|
1825 |
val T = typ_of_dtyp descr typ_assoc d
|
webertj@15547
|
1826 |
val (i, _, _) = interpret thy (typs, []) {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True} (Free ("dummy", T))
|
webertj@15547
|
1827 |
val size = size_of_type i
|
webertj@15547
|
1828 |
(* sanity check *)
|
webertj@15547
|
1829 |
val _ = if size < length xs then
|
webertj@15547
|
1830 |
raise REFUTE ("IDT_constructor_interpreter", "new size of type is less than old size")
|
webertj@15547
|
1831 |
else
|
webertj@15547
|
1832 |
()
|
webertj@15547
|
1833 |
(* extend the existing interpretations *)
|
webertj@15547
|
1834 |
(* int * interpretation list *)
|
webertj@15547
|
1835 |
val (new_offset, intrs) = foldl_map (fn (off, i) => extend_constr (off, ds, i)) (offset, xs)
|
webertj@15547
|
1836 |
(* new elements of the type 'd' are mapped to "undefined" *)
|
webertj@15547
|
1837 |
val undefs = replicate (size - length xs) (make_undef ds)
|
webertj@15547
|
1838 |
in
|
webertj@15547
|
1839 |
(new_offset, Node (intrs @ undefs))
|
webertj@15547
|
1840 |
end
|
webertj@15547
|
1841 |
| extend_constr (_, d::ds, Leaf _) =
|
webertj@15547
|
1842 |
raise REFUTE ("IDT_constructor_interpreter", "interpretation for constructor (with arguments left) is a leaf")
|
webertj@15547
|
1843 |
(* returns 'true' iff the constructor has a recursive argument *)
|
webertj@15547
|
1844 |
(* DatatypeAux.dtyp list -> bool *)
|
webertj@15547
|
1845 |
fun is_rec_constr ds =
|
webertj@15547
|
1846 |
Library.exists DatatypeAux.is_rec_type ds
|
webertj@15611
|
1847 |
(* constructors before 'Const (s, T)' generate elements of the datatype *)
|
webertj@15611
|
1848 |
val offset = size_of_dtyp thy typs' descr typ_assoc constrs1
|
webertj@15547
|
1849 |
in
|
webertj@15547
|
1850 |
case depth of
|
webertj@15547
|
1851 |
NONE => (* equivalent to a depth of 1 *)
|
webertj@15547
|
1852 |
SOME (snd (make_constr (offset, ctypes)), model, args)
|
webertj@15547
|
1853 |
| SOME 0 =>
|
webertj@15547
|
1854 |
raise REFUTE ("IDT_constructor_interpreter", "depth is 0")
|
webertj@15547
|
1855 |
| SOME 1 =>
|
webertj@15547
|
1856 |
SOME (snd (make_constr (offset, ctypes)), model, args)
|
webertj@15547
|
1857 |
| SOME n => (* n > 1 *)
|
webertj@15547
|
1858 |
let
|
webertj@15547
|
1859 |
(* interpret the constructor at depth-1 *)
|
webertj@15547
|
1860 |
val (iC, _, _) = interpret thy (typs', []) {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True} (Const (s, T))
|
webertj@15547
|
1861 |
(* elements generated by the constructor at depth-1 must be added to 'offset' *)
|
webertj@15547
|
1862 |
(* interpretation -> int *)
|
webertj@15547
|
1863 |
fun number_of_defined_elements (Leaf xs) =
|
webertj@15547
|
1864 |
if find_index_eq True xs = (~1) then 0 else 1
|
webertj@15547
|
1865 |
| number_of_defined_elements (Node xs) =
|
webertj@15547
|
1866 |
sum (map number_of_defined_elements xs)
|
webertj@15547
|
1867 |
(* int *)
|
webertj@15547
|
1868 |
val offset' = offset + number_of_defined_elements iC
|
webertj@15547
|
1869 |
in
|
webertj@15547
|
1870 |
SOME (snd (extend_constr (offset', ctypes, iC)), model, args)
|
webertj@14807
|
1871 |
end
|
webertj@14807
|
1872 |
end
|
webertj@15547
|
1873 |
end
|
webertj@15547
|
1874 |
| NONE => (* body type is not an inductive datatype *)
|
webertj@15547
|
1875 |
NONE)
|
webertj@15547
|
1876 |
| _ => (* body type is a (free or schematic) type variable *)
|
webertj@15547
|
1877 |
NONE)
|
webertj@15547
|
1878 |
| _ => (* term is not a constant *)
|
webertj@15547
|
1879 |
NONE)
|
webertj@14807
|
1880 |
end;
|
webertj@14807
|
1881 |
|
webertj@14807
|
1882 |
(* theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option *)
|
webertj@14807
|
1883 |
|
webertj@15547
|
1884 |
(* Difficult code ahead. Make sure you understand the 'IDT_constructor_interpreter' *)
|
webertj@15547
|
1885 |
(* and the order in which it enumerates elements of an IDT before you try to *)
|
webertj@15547
|
1886 |
(* understand this function. *)
|
webertj@15547
|
1887 |
|
webertj@15547
|
1888 |
fun IDT_recursion_interpreter thy model args t =
|
webertj@15547
|
1889 |
case strip_comb t of (* careful: here we descend arbitrarily deep into 't', *)
|
webertj@15547
|
1890 |
(* possibly before any other interpreter for atomic *)
|
webertj@15547
|
1891 |
(* terms has had a chance to look at 't' *)
|
webertj@15547
|
1892 |
(Const (s, T), params) =>
|
webertj@15547
|
1893 |
(* iterate over all datatypes in 'thy' *)
|
webertj@15547
|
1894 |
Symtab.foldl (fn (result, (_, info)) =>
|
webertj@15547
|
1895 |
case result of
|
webertj@15547
|
1896 |
SOME _ =>
|
webertj@15547
|
1897 |
result (* just keep 'result' *)
|
webertj@15547
|
1898 |
| NONE =>
|
webertj@15547
|
1899 |
if s mem (#rec_names info) then
|
webertj@15767
|
1900 |
(* we do have a recursion operator of the datatype given by 'info', *)
|
webertj@15767
|
1901 |
(* or of a mutually recursive datatype *)
|
webertj@15547
|
1902 |
let
|
webertj@15767
|
1903 |
val index = #index info
|
webertj@15767
|
1904 |
val descr = #descr info
|
haftmann@17314
|
1905 |
val (dtname, dtyps, _) = (the o AList.lookup (op =) descr) index
|
webertj@15767
|
1906 |
(* number of all constructors, including those of different *)
|
webertj@15767
|
1907 |
(* (mutually recursive) datatypes within the same descriptor 'descr' *)
|
webertj@15767
|
1908 |
val mconstrs_count = sum (map (fn (_, (_, _, cs)) => length cs) descr)
|
webertj@15767
|
1909 |
val params_count = length params
|
webertj@15767
|
1910 |
(* the type of a recursion operator: [T1, ..., Tn, IDT] ---> Tresult *)
|
webertj@15767
|
1911 |
val IDT = List.nth (binder_types T, mconstrs_count)
|
webertj@15547
|
1912 |
in
|
webertj@15767
|
1913 |
if (fst o dest_Type) IDT <> dtname then
|
webertj@15767
|
1914 |
(* recursion operator of a mutually recursive datatype *)
|
webertj@15767
|
1915 |
NONE
|
webertj@15767
|
1916 |
else if mconstrs_count < params_count then
|
webertj@15547
|
1917 |
(* too many actual parameters; for now we'll use the *)
|
webertj@15547
|
1918 |
(* 'stlc_interpreter' to strip off one application *)
|
webertj@15547
|
1919 |
NONE
|
webertj@15767
|
1920 |
else if mconstrs_count > params_count then
|
webertj@15547
|
1921 |
(* too few actual parameters; we use eta expansion *)
|
webertj@15547
|
1922 |
(* Note that the resulting expansion of lambda abstractions *)
|
webertj@15547
|
1923 |
(* by the 'stlc_interpreter' may be rather slow (depending on *)
|
webertj@15547
|
1924 |
(* the argument types and the size of the IDT, of course). *)
|
webertj@15767
|
1925 |
SOME (interpret thy model args (eta_expand t (mconstrs_count - params_count)))
|
webertj@15767
|
1926 |
else (* mconstrs_count = params_count *)
|
webertj@15547
|
1927 |
let
|
webertj@15547
|
1928 |
(* interpret each parameter separately *)
|
webertj@15547
|
1929 |
val ((model', args'), p_intrs) = foldl_map (fn ((m, a), p) =>
|
webertj@15547
|
1930 |
let
|
webertj@15547
|
1931 |
val (i, m', a') = interpret thy m a p
|
webertj@15547
|
1932 |
in
|
webertj@15547
|
1933 |
((m', a'), i)
|
webertj@15547
|
1934 |
end) ((model, args), params)
|
webertj@15767
|
1935 |
val (typs, _) = model'
|
webertj@15547
|
1936 |
val typ_assoc = dtyps ~~ (snd o dest_Type) IDT
|
webertj@15767
|
1937 |
(* interpret each constructor in the descriptor (including *)
|
webertj@15767
|
1938 |
(* those of mutually recursive datatypes) *)
|
webertj@15767
|
1939 |
(* (int * interpretation list) list *)
|
webertj@15767
|
1940 |
val mc_intrs = map (fn (idx, (_, _, cs)) =>
|
webertj@15767
|
1941 |
let
|
webertj@15767
|
1942 |
val c_return_typ = typ_of_dtyp descr typ_assoc (DatatypeAux.DtRec idx)
|
webertj@15767
|
1943 |
in
|
webertj@15767
|
1944 |
(idx, map (fn (cname, cargs) =>
|
webertj@15767
|
1945 |
(#1 o interpret thy (typs, []) {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True})
|
webertj@15767
|
1946 |
(Const (cname, map (typ_of_dtyp descr typ_assoc) cargs ---> c_return_typ))) cs)
|
webertj@15767
|
1947 |
end) descr
|
webertj@15547
|
1948 |
(* the recursion operator is a function that maps every element of *)
|
webertj@15767
|
1949 |
(* the inductive datatype (and of mutually recursive types) to an *)
|
webertj@15783
|
1950 |
(* element of some result type; an array entry of NONE means that *)
|
webertj@15783
|
1951 |
(* the actual result has not been computed yet *)
|
webertj@15767
|
1952 |
(* (int * interpretation option Array.array) list *)
|
webertj@15767
|
1953 |
val INTRS = map (fn (idx, _) =>
|
webertj@15767
|
1954 |
let
|
webertj@15767
|
1955 |
val T = typ_of_dtyp descr typ_assoc (DatatypeAux.DtRec idx)
|
webertj@15767
|
1956 |
val (i, _, _) = interpret thy (typs, []) {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True} (Free ("dummy", T))
|
webertj@15767
|
1957 |
val size = size_of_type i
|
webertj@15767
|
1958 |
in
|
webertj@15767
|
1959 |
(idx, Array.array (size, NONE))
|
webertj@15767
|
1960 |
end) descr
|
webertj@15767
|
1961 |
(* takes an interpretation, and if some leaf of this interpretation *)
|
webertj@15767
|
1962 |
(* is the 'elem'-th element of the type, the indices of the arguments *)
|
webertj@15767
|
1963 |
(* leading to this leaf are returned *)
|
webertj@15547
|
1964 |
(* interpretation -> int -> int list option *)
|
webertj@15547
|
1965 |
fun get_args (Leaf xs) elem =
|
webertj@15547
|
1966 |
if find_index_eq True xs = elem then
|
webertj@15547
|
1967 |
SOME []
|
webertj@15547
|
1968 |
else
|
webertj@15547
|
1969 |
NONE
|
webertj@15547
|
1970 |
| get_args (Node xs) elem =
|
webertj@15547
|
1971 |
let
|
webertj@15547
|
1972 |
(* interpretation * int -> int list option *)
|
webertj@15547
|
1973 |
fun search ([], _) =
|
webertj@15547
|
1974 |
NONE
|
webertj@15547
|
1975 |
| search (x::xs, n) =
|
webertj@15547
|
1976 |
(case get_args x elem of
|
webertj@15547
|
1977 |
SOME result => SOME (n::result)
|
webertj@15547
|
1978 |
| NONE => search (xs, n+1))
|
webertj@15547
|
1979 |
in
|
webertj@15547
|
1980 |
search (xs, 0)
|
webertj@15547
|
1981 |
end
|
webertj@15547
|
1982 |
(* returns the index of the constructor and indices for its *)
|
webertj@15547
|
1983 |
(* arguments that generate the 'elem'-th element of the datatype *)
|
webertj@15767
|
1984 |
(* given by 'idx' *)
|
webertj@15767
|
1985 |
(* int -> int -> int * int list *)
|
webertj@15767
|
1986 |
fun get_cargs idx elem =
|
webertj@15547
|
1987 |
let
|
webertj@15547
|
1988 |
(* int * interpretation list -> int * int list *)
|
webertj@15547
|
1989 |
fun get_cargs_rec (_, []) =
|
webertj@15767
|
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 ^ ")")
|
webertj@15547
|
1991 |
| get_cargs_rec (n, x::xs) =
|
webertj@15547
|
1992 |
(case get_args x elem of
|
webertj@15547
|
1993 |
SOME args => (n, args)
|
webertj@15547
|
1994 |
| NONE => get_cargs_rec (n+1, xs))
|
webertj@15547
|
1995 |
in
|
haftmann@17314
|
1996 |
get_cargs_rec (0, (the o AList.lookup (op =) mc_intrs) idx)
|
webertj@15547
|
1997 |
end
|
webertj@15767
|
1998 |
(* returns the number of constructors in datatypes that occur in *)
|
webertj@15767
|
1999 |
(* the descriptor 'descr' before the datatype given by 'idx' *)
|
webertj@15767
|
2000 |
fun get_coffset idx =
|
webertj@15767
|
2001 |
let
|
webertj@15767
|
2002 |
fun get_coffset_acc _ [] =
|
webertj@15767
|
2003 |
raise REFUTE ("IDT_recursion_interpreter", "index " ^ string_of_int idx ^ " not found in descriptor")
|
webertj@15767
|
2004 |
| get_coffset_acc sum ((i, (_, _, cs))::descr') =
|
webertj@15767
|
2005 |
if i=idx then
|
webertj@15767
|
2006 |
sum
|
webertj@15767
|
2007 |
else
|
webertj@15767
|
2008 |
get_coffset_acc (sum + length cs) descr'
|
webertj@15767
|
2009 |
in
|
webertj@15767
|
2010 |
get_coffset_acc 0 descr
|
webertj@15767
|
2011 |
end
|
webertj@15767
|
2012 |
(* computes one entry in INTRS, and recursively all entries needed for it, *)
|
webertj@15767
|
2013 |
(* where 'idx' gives the datatype and 'elem' the element of it *)
|
webertj@15767
|
2014 |
(* int -> int -> interpretation *)
|
webertj@15767
|
2015 |
fun compute_array_entry idx elem =
|
haftmann@17314
|
2016 |
case Array.sub ((the o AList.lookup (op =) INTRS) idx, elem) of
|
webertj@15767
|
2017 |
SOME result =>
|
webertj@15767
|
2018 |
(* simply return the previously computed result *)
|
webertj@15767
|
2019 |
result
|
webertj@15767
|
2020 |
| NONE =>
|
webertj@15547
|
2021 |
let
|
webertj@15547
|
2022 |
(* int * int list *)
|
webertj@15767
|
2023 |
val (c, args) = get_cargs idx elem
|
webertj@15547
|
2024 |
(* interpretation * int list -> interpretation *)
|
webertj@15547
|
2025 |
fun select_subtree (tr, []) =
|
webertj@15547
|
2026 |
tr (* return the whole tree *)
|
webertj@15547
|
2027 |
| select_subtree (Leaf _, _) =
|
webertj@15547
|
2028 |
raise REFUTE ("IDT_recursion_interpreter", "interpretation for parameter is a leaf; cannot select a subtree")
|
webertj@15547
|
2029 |
| select_subtree (Node tr, x::xs) =
|
skalberg@15570
|
2030 |
select_subtree (List.nth (tr, x), xs)
|
webertj@15547
|
2031 |
(* select the correct subtree of the parameter corresponding to constructor 'c' *)
|
webertj@15767
|
2032 |
val p_intr = select_subtree (List.nth (p_intrs, get_coffset idx + c), args)
|
webertj@15767
|
2033 |
(* find the indices of the constructor's recursive arguments *)
|
haftmann@17314
|
2034 |
val (_, _, constrs) = (the o AList.lookup (op =) descr) idx
|
webertj@15767
|
2035 |
val constr_args = (snd o List.nth) (constrs, c)
|
webertj@15767
|
2036 |
val rec_args = List.filter (DatatypeAux.is_rec_type o fst) (constr_args ~~ args)
|
webertj@15767
|
2037 |
val rec_args' = map (fn (dtyp, elem) => (DatatypeAux.dest_DtRec dtyp, elem)) rec_args
|
webertj@15547
|
2038 |
(* apply 'p_intr' to recursively computed results *)
|
webertj@15767
|
2039 |
val result = foldl (fn ((idx, elem), intr) =>
|
webertj@15767
|
2040 |
interpretation_apply (intr, compute_array_entry idx elem)) p_intr rec_args'
|
webertj@15547
|
2041 |
(* update 'INTRS' *)
|
haftmann@17314
|
2042 |
val _ = Array.update ((the o AList.lookup (op =) INTRS) idx, elem, SOME result)
|
webertj@15547
|
2043 |
in
|
webertj@15767
|
2044 |
result
|
webertj@15547
|
2045 |
end
|
webertj@15767
|
2046 |
(* compute all entries in INTRS for the current datatype (given by 'index') *)
|
webertj@15783
|
2047 |
(* TODO: we can use Array.modify instead once PolyML conforms to the ML standard *)
|
webertj@15783
|
2048 |
(* (int * 'a -> 'a) -> 'a array -> unit *)
|
webertj@15783
|
2049 |
fun modifyi f arr =
|
webertj@15783
|
2050 |
let
|
webertj@15783
|
2051 |
val size = Array.length arr
|
webertj@15783
|
2052 |
fun modifyi_loop i =
|
webertj@15783
|
2053 |
if i < size then (
|
webertj@15783
|
2054 |
Array.update (arr, i, f (i, Array.sub (arr, i)));
|
webertj@15783
|
2055 |
modifyi_loop (i+1)
|
webertj@15783
|
2056 |
) else
|
webertj@15783
|
2057 |
()
|
webertj@15783
|
2058 |
in
|
webertj@15783
|
2059 |
modifyi_loop 0
|
webertj@15783
|
2060 |
end
|
haftmann@17314
|
2061 |
val _ = modifyi (fn (i, _) => SOME (compute_array_entry index i)) ((the o AList.lookup (op =) INTRS) index)
|
webertj@15547
|
2062 |
(* 'a Array.array -> 'a list *)
|
webertj@15547
|
2063 |
fun toList arr =
|
webertj@15547
|
2064 |
Array.foldr op:: [] arr
|
webertj@15547
|
2065 |
in
|
webertj@15767
|
2066 |
(* return the part of 'INTRS' that corresponds to the current datatype *)
|
haftmann@17314
|
2067 |
SOME ((Node o map the o toList o the o AList.lookup (op =) INTRS) index, model', args')
|
webertj@15547
|
2068 |
end
|
webertj@15547
|
2069 |
end
|
webertj@15547
|
2070 |
else
|
webertj@15547
|
2071 |
NONE (* not a recursion operator of this datatype *)
|
webertj@15547
|
2072 |
) (NONE, DatatypePackage.get_datatypes thy)
|
webertj@15547
|
2073 |
| _ => (* head of term is not a constant *)
|
webertj@15547
|
2074 |
NONE;
|
webertj@15547
|
2075 |
|
webertj@15547
|
2076 |
(* theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option *)
|
webertj@15547
|
2077 |
|
webertj@14807
|
2078 |
(* only an optimization: 'card' could in principle be interpreted with *)
|
webertj@14807
|
2079 |
(* interpreters available already (using its definition), but the code *)
|
webertj@15547
|
2080 |
(* below is more efficient *)
|
webertj@14807
|
2081 |
|
webertj@14807
|
2082 |
fun Finite_Set_card_interpreter thy model args t =
|
webertj@14807
|
2083 |
case t of
|
webertj@14807
|
2084 |
Const ("Finite_Set.card", Type ("fun", [Type ("set", [T]), Type ("nat", [])])) =>
|
webertj@14807
|
2085 |
let
|
webertj@15547
|
2086 |
val (i_nat, _, _) = interpret thy model {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True} (Free ("dummy", Type ("nat", [])))
|
webertj@14807
|
2087 |
val size_nat = size_of_type i_nat
|
webertj@15547
|
2088 |
val (i_set, _, _) = interpret thy model {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True} (Free ("dummy", Type ("set", [T])))
|
webertj@14807
|
2089 |
val constants = make_constants i_set
|
webertj@14807
|
2090 |
(* interpretation -> int *)
|
webertj@14807
|
2091 |
fun number_of_elements (Node xs) =
|
skalberg@15570
|
2092 |
Library.foldl (fn (n, x) =>
|
webertj@15547
|
2093 |
if x=TT then
|
webertj@15547
|
2094 |
n+1
|
webertj@15547
|
2095 |
else if x=FF then
|
webertj@15547
|
2096 |
n
|
webertj@15547
|
2097 |
else
|
webertj@15547
|
2098 |
raise REFUTE ("Finite_Set_card_interpreter", "interpretation for set type does not yield a Boolean")) (0, xs)
|
webertj@14807
|
2099 |
| number_of_elements (Leaf _) =
|
webertj@14807
|
2100 |
raise REFUTE ("Finite_Set_card_interpreter", "interpretation for set type is a leaf")
|
webertj@14807
|
2101 |
(* takes an interpretation for a set and returns an interpretation for a 'nat' *)
|
webertj@14807
|
2102 |
(* interpretation -> interpretation *)
|
webertj@14807
|
2103 |
fun card i =
|
webertj@14807
|
2104 |
let
|
webertj@14807
|
2105 |
val n = number_of_elements i
|
webertj@14807
|
2106 |
in
|
webertj@14807
|
2107 |
if n<size_nat then
|
webertj@14807
|
2108 |
Leaf ((replicate n False) @ True :: (replicate (size_nat-n-1) False))
|
webertj@14456
|
2109 |
else
|
webertj@14807
|
2110 |
Leaf (replicate size_nat False)
|
webertj@14807
|
2111 |
end
|
webertj@14350
|
2112 |
in
|
skalberg@15531
|
2113 |
SOME (Node (map card constants), model, args)
|
webertj@14350
|
2114 |
end
|
webertj@14807
|
2115 |
| _ =>
|
skalberg@15531
|
2116 |
NONE;
|
webertj@14807
|
2117 |
|
webertj@15547
|
2118 |
(* theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option *)
|
webertj@15547
|
2119 |
|
webertj@15571
|
2120 |
(* only an optimization: 'Finites' could in principle be interpreted with *)
|
webertj@15571
|
2121 |
(* interpreters available already (using its definition), but the code *)
|
webertj@15571
|
2122 |
(* below is more efficient *)
|
webertj@15571
|
2123 |
|
webertj@15571
|
2124 |
fun Finite_Set_Finites_interpreter thy model args t =
|
webertj@15571
|
2125 |
case t of
|
webertj@15571
|
2126 |
Const ("Finite_Set.Finites", Type ("set", [Type ("set", [T])])) =>
|
webertj@15571
|
2127 |
let
|
webertj@15571
|
2128 |
val (i_set, _, _) = interpret thy model {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True} (Free ("dummy", Type ("set", [T])))
|
webertj@15571
|
2129 |
val size_set = size_of_type i_set
|
webertj@15571
|
2130 |
in
|
webertj@15571
|
2131 |
(* we only consider finite models anyway, hence EVERY set is in "Finites" *)
|
webertj@15571
|
2132 |
SOME (Node (replicate size_set TT), model, args)
|
webertj@15571
|
2133 |
end
|
webertj@15571
|
2134 |
| _ =>
|
webertj@15571
|
2135 |
NONE;
|
webertj@15571
|
2136 |
|
webertj@15571
|
2137 |
(* theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option *)
|
webertj@15571
|
2138 |
|
webertj@15547
|
2139 |
(* only an optimization: 'op <' could in principle be interpreted with *)
|
webertj@15547
|
2140 |
(* interpreters available already (using its definition), but the code *)
|
webertj@15547
|
2141 |
(* below is more efficient *)
|
webertj@15547
|
2142 |
|
webertj@15547
|
2143 |
fun Nat_less_interpreter thy model args t =
|
webertj@15547
|
2144 |
case t of
|
haftmann@19277
|
2145 |
Const ("Orderings.less", Type ("fun", [Type ("nat", []), Type ("fun", [Type ("nat", []), Type ("bool", [])])])) =>
|
webertj@15547
|
2146 |
let
|
webertj@15547
|
2147 |
val (i_nat, _, _) = interpret thy model {maxvars=0, def_eq=false, next_idx=1, bounds=[], wellformed=True} (Free ("dummy", Type ("nat", [])))
|
webertj@15547
|
2148 |
val size_nat = size_of_type i_nat
|
webertj@15547
|
2149 |
(* int -> interpretation *)
|
webertj@15547
|
2150 |
(* the 'n'-th nat is not less than the first 'n' nats, while it *)
|
webertj@15547
|
2151 |
(* is less than the remaining 'size_nat - n' nats *)
|
webertj@15547
|
2152 |
fun less n = Node ((replicate n FF) @ (replicate (size_nat - n) TT))
|
webertj@15547
|
2153 |
in
|
webertj@15547
|
2154 |
SOME (Node (map less (1 upto size_nat)), model, args)
|
webertj@15547
|
2155 |
end
|
webertj@15547
|
2156 |
| _ =>
|
webertj@15547
|
2157 |
NONE;
|
webertj@15547
|
2158 |
|
webertj@15547
|
2159 |
(* theory -> model -> arguments -> Term.term -> (interpretation * model * arguments) option *)
|
webertj@15547
|
2160 |
|
haftmann@19233
|
2161 |
(* only an optimization: 'HOL.plus' could in principle be interpreted with*)
|
webertj@15547
|
2162 |
(* interpreters available already (using its definition), but the code *)
|
webertj@15547
|
2163 |
(* below is more efficient *)
|
webertj@15547
|
2164 |
|
webertj@15547
|
2165 |
fun Nat_plus_interpreter thy model args t =
|
|