author | blanchet |
Mon, 06 Dec 2010 13:18:25 +0100 | |
changeset 40987 | 7d52af07bff1 |
parent 40986 | cfd91aa80701 |
child 40988 | f7af68bfa66d |
permissions | -rw-r--r-- |
33982 | 1 |
(* Title: HOL/Tools/Nitpick/nitpick_mono.ML |
33192 | 2 |
Author: Jasmin Blanchette, TU Muenchen |
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
3 |
Copyright 2009, 2010 |
33192 | 4 |
|
35384 | 5 |
Monotonicity inference for higher-order logic. |
33192 | 6 |
*) |
7 |
||
8 |
signature NITPICK_MONO = |
|
9 |
sig |
|
35070
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
blanchet
parents:
34998
diff
changeset
|
10 |
type hol_context = Nitpick_HOL.hol_context |
33192 | 11 |
|
38179 | 12 |
val trace : bool Unsynchronized.ref |
33192 | 13 |
val formulas_monotonic : |
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
14 |
hol_context -> bool -> typ -> term list * term list -> bool |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
15 |
val finitize_funs : |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
16 |
hol_context -> bool -> (typ option * bool option) list -> typ |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
17 |
-> term list * term list -> term list * term list |
33192 | 18 |
end; |
19 |
||
33232
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33192
diff
changeset
|
20 |
structure Nitpick_Mono : NITPICK_MONO = |
33192 | 21 |
struct |
22 |
||
33232
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33192
diff
changeset
|
23 |
open Nitpick_Util |
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33192
diff
changeset
|
24 |
open Nitpick_HOL |
33192 | 25 |
|
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
26 |
datatype sign = Plus | Minus |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
27 |
|
33192 | 28 |
type var = int |
29 |
||
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
30 |
datatype annotation = Gen | New | Fls | Tru |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
31 |
datatype annotation_atom = A of annotation | V of var |
33192 | 32 |
|
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
33 |
type literal = var * annotation |
33192 | 34 |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
35 |
datatype mtyp = |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
36 |
MAlpha | |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
37 |
MFun of mtyp * annotation_atom * mtyp | |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
38 |
MPair of mtyp * mtyp | |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
39 |
MType of string * mtyp list | |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
40 |
MRec of string * typ list |
33192 | 41 |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
42 |
datatype mterm = |
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
43 |
MRaw of term * mtyp | |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
44 |
MAbs of string * typ * mtyp * annotation_atom * mterm | |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
45 |
MApp of mterm * mterm |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
46 |
|
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
47 |
type mdata = |
35070
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
blanchet
parents:
34998
diff
changeset
|
48 |
{hol_ctxt: hol_context, |
35190
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents:
35079
diff
changeset
|
49 |
binarize: bool, |
33192 | 50 |
alpha_T: typ, |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
51 |
no_harmless: bool, |
33192 | 52 |
max_fresh: int Unsynchronized.ref, |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
53 |
datatype_mcache: ((string * typ list) * mtyp) list Unsynchronized.ref, |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
54 |
constr_mcache: (styp * mtyp) list Unsynchronized.ref} |
33192 | 55 |
|
35812
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
56 |
exception UNSOLVABLE of unit |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
57 |
exception MTYPE of string * mtyp list * typ list |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
58 |
exception MTERM of string * mterm list |
33192 | 59 |
|
38179 | 60 |
val trace = Unsynchronized.ref false |
61 |
fun trace_msg msg = if !trace then tracing (msg ()) else () |
|
33192 | 62 |
|
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
63 |
fun string_for_sign Plus = "+" |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
64 |
| string_for_sign Minus = "-" |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
65 |
|
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
66 |
fun negate_sign Plus = Minus |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
67 |
| negate_sign Minus = Plus |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
68 |
|
33192 | 69 |
val string_for_var = signed_string_of_int |
70 |
fun string_for_vars sep [] = "0\<^bsub>" ^ sep ^ "\<^esub>" |
|
71 |
| string_for_vars sep xs = space_implode sep (map string_for_var xs) |
|
72 |
fun subscript_string_for_vars sep xs = |
|
73 |
if null xs then "" else "\<^bsub>" ^ string_for_vars sep xs ^ "\<^esub>" |
|
74 |
||
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
75 |
fun string_for_annotation Gen = "G" |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
76 |
| string_for_annotation New = "N" |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
77 |
| string_for_annotation Fls = "F" |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
78 |
| string_for_annotation Tru = "T" |
33192 | 79 |
|
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
80 |
fun string_for_annotation_atom (A a) = string_for_annotation a |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
81 |
| string_for_annotation_atom (V x) = string_for_var x |
33192 | 82 |
|
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
83 |
fun string_for_literal (x, a) = |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
84 |
string_for_var x ^ " = " ^ string_for_annotation a |
33192 | 85 |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
86 |
val bool_M = MType (@{type_name bool}, []) |
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
87 |
val dummy_M = MType (nitpick_prefix ^ "dummy", []) |
33192 | 88 |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
89 |
fun is_MRec (MRec _) = true |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
90 |
| is_MRec _ = false |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
91 |
fun dest_MFun (MFun z) = z |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
92 |
| dest_MFun M = raise MTYPE ("Nitpick_Mono.dest_MFun", [M], []) |
33192 | 93 |
|
94 |
val no_prec = 100 |
|
95 |
||
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
96 |
fun precedence_of_mtype (MFun _) = 1 |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
97 |
| precedence_of_mtype (MPair _) = 2 |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
98 |
| precedence_of_mtype _ = no_prec |
33192 | 99 |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
100 |
val string_for_mtype = |
33192 | 101 |
let |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
102 |
fun aux outer_prec M = |
33192 | 103 |
let |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
104 |
val prec = precedence_of_mtype M |
33192 | 105 |
val need_parens = (prec < outer_prec) |
106 |
in |
|
107 |
(if need_parens then "(" else "") ^ |
|
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
108 |
(if M = dummy_M then |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
109 |
"_" |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
110 |
else case M of |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
111 |
MAlpha => "\<alpha>" |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
112 |
| MFun (M1, a, M2) => |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
113 |
aux (prec + 1) M1 ^ " \<Rightarrow>\<^bsup>" ^ |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
114 |
string_for_annotation_atom a ^ "\<^esup> " ^ aux prec M2 |
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
115 |
| MPair (M1, M2) => aux (prec + 1) M1 ^ " \<times> " ^ aux prec M2 |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
116 |
| MType (s, []) => |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
117 |
if s = @{type_name prop} orelse s = @{type_name bool} then "o" |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
118 |
else s |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
119 |
| MType (s, Ms) => "(" ^ commas (map (aux 0) Ms) ^ ") " ^ s |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
120 |
| MRec (s, _) => "[" ^ s ^ "]") ^ |
33192 | 121 |
(if need_parens then ")" else "") |
122 |
end |
|
123 |
in aux 0 end |
|
124 |
||
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
125 |
fun flatten_mtype (MPair (M1, M2)) = maps flatten_mtype [M1, M2] |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
126 |
| flatten_mtype (MType (_, Ms)) = maps flatten_mtype Ms |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
127 |
| flatten_mtype M = [M] |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
128 |
|
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
129 |
fun precedence_of_mterm (MRaw _) = no_prec |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
130 |
| precedence_of_mterm (MAbs _) = 1 |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
131 |
| precedence_of_mterm (MApp _) = 2 |
33192 | 132 |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
133 |
fun string_for_mterm ctxt = |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
134 |
let |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
135 |
fun mtype_annotation M = "\<^bsup>" ^ string_for_mtype M ^ "\<^esup>" |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
136 |
fun aux outer_prec m = |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
137 |
let |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
138 |
val prec = precedence_of_mterm m |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
139 |
val need_parens = (prec < outer_prec) |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
140 |
in |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
141 |
(if need_parens then "(" else "") ^ |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
142 |
(case m of |
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
143 |
MRaw (t, M) => Syntax.string_of_term ctxt t ^ mtype_annotation M |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
144 |
| MAbs (s, _, M, a, m) => |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
145 |
"\<lambda>" ^ s ^ mtype_annotation M ^ ".\<^bsup>" ^ |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
146 |
string_for_annotation_atom a ^ "\<^esup> " ^ aux prec m |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
147 |
| MApp (m1, m2) => aux prec m1 ^ " " ^ aux (prec + 1) m2) ^ |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
148 |
(if need_parens then ")" else "") |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
149 |
end |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
150 |
in aux 0 end |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
151 |
|
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
152 |
fun mtype_of_mterm (MRaw (_, M)) = M |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
153 |
| mtype_of_mterm (MAbs (_, _, M, a, m)) = MFun (M, a, mtype_of_mterm m) |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
154 |
| mtype_of_mterm (MApp (m1, _)) = |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
155 |
case mtype_of_mterm m1 of |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
156 |
MFun (_, _, M12) => M12 |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
157 |
| M1 => raise MTYPE ("Nitpick_Mono.mtype_of_mterm", [M1], []) |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
158 |
|
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
159 |
fun strip_mcomb (MApp (m1, m2)) = strip_mcomb m1 ||> (fn ms => append ms [m2]) |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
160 |
| strip_mcomb m = (m, []) |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
161 |
|
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
162 |
fun initial_mdata hol_ctxt binarize no_harmless alpha_T = |
35190
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents:
35079
diff
changeset
|
163 |
({hol_ctxt = hol_ctxt, binarize = binarize, alpha_T = alpha_T, |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
164 |
no_harmless = no_harmless, max_fresh = Unsynchronized.ref 0, |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
165 |
datatype_mcache = Unsynchronized.ref [], |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
166 |
constr_mcache = Unsynchronized.ref []} : mdata) |
33192 | 167 |
|
168 |
fun could_exist_alpha_subtype alpha_T (T as Type (_, Ts)) = |
|
34936
c4f04bee79f3
some work on Nitpick's support for quotient types;
blanchet
parents:
34123
diff
changeset
|
169 |
T = alpha_T orelse (not (is_fp_iterator_type T) andalso |
c4f04bee79f3
some work on Nitpick's support for quotient types;
blanchet
parents:
34123
diff
changeset
|
170 |
exists (could_exist_alpha_subtype alpha_T) Ts) |
33192 | 171 |
| could_exist_alpha_subtype alpha_T T = (T = alpha_T) |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
172 |
fun could_exist_alpha_sub_mtype _ (alpha_T as TFree _) T = |
34121
5e831d805118
get rid of polymorphic equality in Nitpick's code + a few minor cleanups
blanchet
parents:
33982
diff
changeset
|
173 |
could_exist_alpha_subtype alpha_T T |
37256
0dca1ec52999
thread along context instead of theory for typedef lookup
blanchet
parents:
36385
diff
changeset
|
174 |
| could_exist_alpha_sub_mtype ctxt alpha_T T = |
0dca1ec52999
thread along context instead of theory for typedef lookup
blanchet
parents:
36385
diff
changeset
|
175 |
(T = alpha_T orelse is_datatype ctxt [(NONE, true)] T) |
33192 | 176 |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
177 |
fun exists_alpha_sub_mtype MAlpha = true |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
178 |
| exists_alpha_sub_mtype (MFun (M1, _, M2)) = |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
179 |
exists exists_alpha_sub_mtype [M1, M2] |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
180 |
| exists_alpha_sub_mtype (MPair (M1, M2)) = |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
181 |
exists exists_alpha_sub_mtype [M1, M2] |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
182 |
| exists_alpha_sub_mtype (MType (_, Ms)) = exists exists_alpha_sub_mtype Ms |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
183 |
| exists_alpha_sub_mtype (MRec _) = true |
33192 | 184 |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
185 |
fun exists_alpha_sub_mtype_fresh MAlpha = true |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
186 |
| exists_alpha_sub_mtype_fresh (MFun (_, V _, _)) = true |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
187 |
| exists_alpha_sub_mtype_fresh (MFun (_, _, M2)) = |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
188 |
exists_alpha_sub_mtype_fresh M2 |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
189 |
| exists_alpha_sub_mtype_fresh (MPair (M1, M2)) = |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
190 |
exists exists_alpha_sub_mtype_fresh [M1, M2] |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
191 |
| exists_alpha_sub_mtype_fresh (MType (_, Ms)) = |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
192 |
exists exists_alpha_sub_mtype_fresh Ms |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
193 |
| exists_alpha_sub_mtype_fresh (MRec _) = true |
33192 | 194 |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
195 |
fun constr_mtype_for_binders z Ms = |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
196 |
fold_rev (fn M => curry3 MFun M (A Gen)) Ms (MRec z) |
33192 | 197 |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
198 |
fun repair_mtype _ _ MAlpha = MAlpha |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
199 |
| repair_mtype cache seen (MFun (M1, a, M2)) = |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
200 |
MFun (repair_mtype cache seen M1, a, repair_mtype cache seen M2) |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
201 |
| repair_mtype cache seen (MPair Mp) = |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
202 |
MPair (pairself (repair_mtype cache seen) Mp) |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
203 |
| repair_mtype cache seen (MType (s, Ms)) = |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
204 |
MType (s, maps (flatten_mtype o repair_mtype cache seen) Ms) |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
205 |
| repair_mtype cache seen (MRec (z as (s, _))) = |
33192 | 206 |
case AList.lookup (op =) cache z |> the of |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
207 |
MRec _ => MType (s, []) |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
208 |
| M => if member (op =) seen M then MType (s, []) |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
209 |
else repair_mtype cache (M :: seen) M |
33192 | 210 |
|
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
211 |
fun repair_datatype_mcache cache = |
33192 | 212 |
let |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
213 |
fun repair_one (z, M) = |
33192 | 214 |
Unsynchronized.change cache |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
215 |
(AList.update (op =) (z, repair_mtype (!cache) [] M)) |
33192 | 216 |
in List.app repair_one (rev (!cache)) end |
217 |
||
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
218 |
fun repair_constr_mcache dtype_cache constr_mcache = |
33192 | 219 |
let |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
220 |
fun repair_one (x, M) = |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
221 |
Unsynchronized.change constr_mcache |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
222 |
(AList.update (op =) (x, repair_mtype dtype_cache [] M)) |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
223 |
in List.app repair_one (!constr_mcache) end |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
224 |
|
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
225 |
fun is_fin_fun_supported_type @{typ prop} = true |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
226 |
| is_fin_fun_supported_type @{typ bool} = true |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
227 |
| is_fin_fun_supported_type (Type (@{type_name option}, _)) = true |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
228 |
| is_fin_fun_supported_type _ = false |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
229 |
fun fin_fun_body _ _ (t as @{term False}) = SOME t |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
230 |
| fin_fun_body _ _ (t as Const (@{const_name None}, _)) = SOME t |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
231 |
| fin_fun_body dom_T ran_T |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
232 |
((t0 as Const (@{const_name If}, _)) |
38864
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
haftmann
parents:
38795
diff
changeset
|
233 |
$ (t1 as Const (@{const_name HOL.eq}, _) $ Bound 0 $ t1') |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
234 |
$ t2 $ t3) = |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
235 |
(if loose_bvar1 (t1', 0) then |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
236 |
NONE |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
237 |
else case fin_fun_body dom_T ran_T t3 of |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
238 |
NONE => NONE |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
239 |
| SOME t3 => |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
240 |
SOME (t0 $ (Const (@{const_name is_unknown}, dom_T --> bool_T) $ t1') |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
241 |
$ (Const (@{const_name unknown}, ran_T)) $ (t0 $ t1 $ t2 $ t3))) |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
242 |
| fin_fun_body _ _ _ = NONE |
33192 | 243 |
|
35672 | 244 |
fun fresh_mfun_for_fun_type (mdata as {max_fresh, ...} : mdata) all_minus |
245 |
T1 T2 = |
|
33192 | 246 |
let |
35672 | 247 |
val M1 = fresh_mtype_for_type mdata all_minus T1 |
248 |
val M2 = fresh_mtype_for_type mdata all_minus T2 |
|
249 |
val a = if not all_minus andalso exists_alpha_sub_mtype_fresh M1 andalso |
|
250 |
is_fin_fun_supported_type (body_type T2) then |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
251 |
V (Unsynchronized.inc max_fresh) |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
252 |
else |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
253 |
A Gen |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
254 |
in (M1, a, M2) end |
37256
0dca1ec52999
thread along context instead of theory for typedef lookup
blanchet
parents:
36385
diff
changeset
|
255 |
and fresh_mtype_for_type (mdata as {hol_ctxt as {ctxt, ...}, binarize, alpha_T, |
35672 | 256 |
datatype_mcache, constr_mcache, ...}) |
257 |
all_minus = |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
258 |
let |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
259 |
fun do_type T = |
33192 | 260 |
if T = alpha_T then |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
261 |
MAlpha |
33192 | 262 |
else case T of |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
263 |
Type (@{type_name fun}, [T1, T2]) => |
39342 | 264 |
MFun (fresh_mfun_for_fun_type mdata all_minus T1 T2) |
38190
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
blanchet
parents:
38179
diff
changeset
|
265 |
| Type (@{type_name prod}, [T1, T2]) => MPair (pairself do_type (T1, T2)) |
33192 | 266 |
| Type (z as (s, _)) => |
37256
0dca1ec52999
thread along context instead of theory for typedef lookup
blanchet
parents:
36385
diff
changeset
|
267 |
if could_exist_alpha_sub_mtype ctxt alpha_T T then |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
268 |
case AList.lookup (op =) (!datatype_mcache) z of |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
269 |
SOME M => M |
33192 | 270 |
| NONE => |
271 |
let |
|
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
272 |
val _ = Unsynchronized.change datatype_mcache (cons (z, MRec z)) |
35219
15a5f213ef5b
fix bug in Nitpick's monotonicity code w.r.t. binary integers
blanchet
parents:
35190
diff
changeset
|
273 |
val xs = binarized_and_boxed_datatype_constrs hol_ctxt binarize T |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
274 |
val (all_Ms, constr_Ms) = |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
275 |
fold_rev (fn (_, T') => fn (all_Ms, constr_Ms) => |
33192 | 276 |
let |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
277 |
val binder_Ms = map do_type (binder_types T') |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
278 |
val new_Ms = filter exists_alpha_sub_mtype_fresh |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
279 |
binder_Ms |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
280 |
val constr_M = constr_mtype_for_binders z |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
281 |
binder_Ms |
33192 | 282 |
in |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
283 |
(union (op =) new_Ms all_Ms, |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
284 |
constr_M :: constr_Ms) |
33192 | 285 |
end) |
286 |
xs ([], []) |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
287 |
val M = MType (s, all_Ms) |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
288 |
val _ = Unsynchronized.change datatype_mcache |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
289 |
(AList.update (op =) (z, M)) |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
290 |
val _ = Unsynchronized.change constr_mcache |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
291 |
(append (xs ~~ constr_Ms)) |
33192 | 292 |
in |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
293 |
if forall (not o is_MRec o snd) (!datatype_mcache) then |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
294 |
(repair_datatype_mcache datatype_mcache; |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
295 |
repair_constr_mcache (!datatype_mcache) constr_mcache; |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
296 |
AList.lookup (op =) (!datatype_mcache) z |> the) |
33192 | 297 |
else |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
298 |
M |
33192 | 299 |
end |
300 |
else |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
301 |
MType (s, []) |
37260
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
302 |
| _ => MType (simple_string_of_typ T, []) |
33192 | 303 |
in do_type end |
304 |
||
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
305 |
fun prodM_factors (MPair (M1, M2)) = maps prodM_factors [M1, M2] |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
306 |
| prodM_factors M = [M] |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
307 |
fun curried_strip_mtype (MFun (M1, _, M2)) = |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
308 |
curried_strip_mtype M2 |>> append (prodM_factors M1) |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
309 |
| curried_strip_mtype M = ([], M) |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
310 |
fun sel_mtype_from_constr_mtype s M = |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
311 |
let val (arg_Ms, dataM) = curried_strip_mtype M in |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
312 |
MFun (dataM, A Gen, |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
313 |
case sel_no_from_name s of ~1 => bool_M | n => nth arg_Ms n) |
33192 | 314 |
end |
315 |
||
37256
0dca1ec52999
thread along context instead of theory for typedef lookup
blanchet
parents:
36385
diff
changeset
|
316 |
fun mtype_for_constr (mdata as {hol_ctxt = {ctxt, ...}, alpha_T, constr_mcache, |
33192 | 317 |
...}) (x as (_, T)) = |
37256
0dca1ec52999
thread along context instead of theory for typedef lookup
blanchet
parents:
36385
diff
changeset
|
318 |
if could_exist_alpha_sub_mtype ctxt alpha_T T then |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
319 |
case AList.lookup (op =) (!constr_mcache) x of |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
320 |
SOME M => M |
35384 | 321 |
| NONE => if T = alpha_T then |
35672 | 322 |
let val M = fresh_mtype_for_type mdata false T in |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
323 |
(Unsynchronized.change constr_mcache (cons (x, M)); M) |
35384 | 324 |
end |
325 |
else |
|
35672 | 326 |
(fresh_mtype_for_type mdata false (body_type T); |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
327 |
AList.lookup (op =) (!constr_mcache) x |> the) |
33192 | 328 |
else |
35672 | 329 |
fresh_mtype_for_type mdata false T |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
330 |
fun mtype_for_sel (mdata as {hol_ctxt, binarize, ...}) (x as (s, _)) = |
35190
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents:
35079
diff
changeset
|
331 |
x |> binarized_and_boxed_constr_for_sel hol_ctxt binarize |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
332 |
|> mtype_for_constr mdata |> sel_mtype_from_constr_mtype s |
33192 | 333 |
|
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
334 |
fun resolve_annotation_atom lits (V x) = |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
335 |
x |> AList.lookup (op =) lits |> Option.map A |> the_default (V x) |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
336 |
| resolve_annotation_atom _ a = a |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
337 |
fun resolve_mtype lits = |
33192 | 338 |
let |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
339 |
fun aux MAlpha = MAlpha |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
340 |
| aux (MFun (M1, a, M2)) = |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
341 |
MFun (aux M1, resolve_annotation_atom lits a, aux M2) |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
342 |
| aux (MPair Mp) = MPair (pairself aux Mp) |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
343 |
| aux (MType (s, Ms)) = MType (s, map aux Ms) |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
344 |
| aux (MRec z) = MRec z |
33192 | 345 |
in aux end |
346 |
||
347 |
datatype comp_op = Eq | Leq |
|
348 |
||
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
349 |
type comp = annotation_atom * annotation_atom * comp_op * var list |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
350 |
type annotation_expr = literal list |
33192 | 351 |
|
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
352 |
type constraint_set = literal list * comp list * annotation_expr list |
33192 | 353 |
|
354 |
fun string_for_comp_op Eq = "=" |
|
355 |
| string_for_comp_op Leq = "\<le>" |
|
356 |
||
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
357 |
fun string_for_annotation_expr [] = "\<bot>" |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
358 |
| string_for_annotation_expr lits = |
33192 | 359 |
space_implode " \<or> " (map string_for_literal lits) |
360 |
||
361 |
fun do_literal _ NONE = NONE |
|
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
362 |
| do_literal (x, a) (SOME lits) = |
33192 | 363 |
case AList.lookup (op =) lits x of |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
364 |
SOME a' => if a = a' then SOME lits else NONE |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
365 |
| NONE => SOME ((x, a) :: lits) |
33192 | 366 |
|
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
367 |
fun do_annotation_atom_comp Eq [] a1 a2 (accum as (lits, comps)) = |
33192 | 368 |
(case (a1, a2) of |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
369 |
(A sn1, A sn2) => if sn1 = sn2 then SOME accum else NONE |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
370 |
| (V x1, A sn2) => |
33192 | 371 |
Option.map (rpair comps) (do_literal (x1, sn2) (SOME lits)) |
372 |
| (V _, V _) => SOME (lits, insert (op =) (a1, a2, Eq, []) comps) |
|
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
373 |
| _ => do_annotation_atom_comp Eq [] a2 a1 accum) |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
374 |
| do_annotation_atom_comp Leq [] a1 a2 (accum as (lits, comps)) = |
33192 | 375 |
(case (a1, a2) of |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
376 |
(_, A Gen) => SOME accum |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
377 |
| (A Fls, _) => SOME accum |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
378 |
| (A Gen, A Fls) => NONE |
33192 | 379 |
| (V _, V _) => SOME (lits, insert (op =) (a1, a2, Leq, []) comps) |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
380 |
| _ => do_annotation_atom_comp Eq [] a1 a2 accum) |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
381 |
| do_annotation_atom_comp cmp xs a1 a2 (lits, comps) = |
33192 | 382 |
SOME (lits, insert (op =) (a1, a2, cmp, xs) comps) |
383 |
||
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
384 |
fun do_mtype_comp _ _ _ _ NONE = NONE |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
385 |
| do_mtype_comp _ _ MAlpha MAlpha accum = accum |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
386 |
| do_mtype_comp Eq xs (MFun (M11, a1, M12)) (MFun (M21, a2, M22)) |
33192 | 387 |
(SOME accum) = |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
388 |
accum |> do_annotation_atom_comp Eq xs a1 a2 |> do_mtype_comp Eq xs M11 M21 |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
389 |
|> do_mtype_comp Eq xs M12 M22 |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
390 |
| do_mtype_comp Leq xs (MFun (M11, a1, M12)) (MFun (M21, a2, M22)) |
33192 | 391 |
(SOME accum) = |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
392 |
(if exists_alpha_sub_mtype M11 then |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
393 |
accum |> do_annotation_atom_comp Leq xs a1 a2 |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
394 |
|> do_mtype_comp Leq xs M21 M11 |
33192 | 395 |
|> (case a2 of |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
396 |
A Gen => I |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
397 |
| A Fls => do_mtype_comp Leq xs M11 M21 |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
398 |
| V x => do_mtype_comp Leq (x :: xs) M11 M21) |
33192 | 399 |
else |
400 |
SOME accum) |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
401 |
|> do_mtype_comp Leq xs M12 M22 |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
402 |
| do_mtype_comp cmp xs (M1 as MPair (M11, M12)) (M2 as MPair (M21, M22)) |
33192 | 403 |
accum = |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
404 |
(accum |> fold (uncurry (do_mtype_comp cmp xs)) [(M11, M21), (M12, M22)] |
40722
441260986b63
make two copies (!) of Library.UnequalLengths coincide with ListPair.UnequalLengths;
wenzelm
parents:
40146
diff
changeset
|
405 |
handle ListPair.UnequalLengths => |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
406 |
raise MTYPE ("Nitpick_Mono.do_mtype_comp", [M1, M2], [])) |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
407 |
| do_mtype_comp _ _ (MType _) (MType _) accum = |
33192 | 408 |
accum (* no need to compare them thanks to the cache *) |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
409 |
| do_mtype_comp cmp _ M1 M2 _ = |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
410 |
raise MTYPE ("Nitpick_Mono.do_mtype_comp (" ^ string_for_comp_op cmp ^ ")", |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
411 |
[M1, M2], []) |
33192 | 412 |
|
35832 | 413 |
fun add_mtype_comp cmp M1 M2 ((lits, comps, sexps) : constraint_set) = |
38179 | 414 |
(trace_msg (fn () => "*** Add " ^ string_for_mtype M1 ^ " " ^ |
415 |
string_for_comp_op cmp ^ " " ^ string_for_mtype M2); |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
416 |
case do_mtype_comp cmp [] M1 M2 (SOME (lits, comps)) of |
38179 | 417 |
NONE => (trace_msg (K "**** Unsolvable"); raise UNSOLVABLE ()) |
35812
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
418 |
| SOME (lits, comps) => (lits, comps, sexps)) |
33192 | 419 |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
420 |
val add_mtypes_equal = add_mtype_comp Eq |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
421 |
val add_is_sub_mtype = add_mtype_comp Leq |
33192 | 422 |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
423 |
fun do_notin_mtype_fv _ _ _ NONE = NONE |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
424 |
| do_notin_mtype_fv Minus _ MAlpha accum = accum |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
425 |
| do_notin_mtype_fv Plus [] MAlpha _ = NONE |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
426 |
| do_notin_mtype_fv Plus [(x, sn)] MAlpha (SOME (lits, sexps)) = |
33192 | 427 |
SOME lits |> do_literal (x, sn) |> Option.map (rpair sexps) |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
428 |
| do_notin_mtype_fv Plus sexp MAlpha (SOME (lits, sexps)) = |
33192 | 429 |
SOME (lits, insert (op =) sexp sexps) |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
430 |
| do_notin_mtype_fv sn sexp (MFun (M1, A a, M2)) accum = |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
431 |
accum |> (if a = Fls andalso sn = Plus then do_notin_mtype_fv Plus sexp M1 |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
432 |
else I) |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
433 |
|> (if a = Gen orelse sn = Plus then do_notin_mtype_fv Minus sexp M1 |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
434 |
else I) |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
435 |
|> do_notin_mtype_fv sn sexp M2 |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
436 |
| do_notin_mtype_fv Plus sexp (MFun (M1, V x, M2)) accum = |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
437 |
accum |> (case do_literal (x, Gen) (SOME sexp) of |
33192 | 438 |
NONE => I |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
439 |
| SOME sexp' => do_notin_mtype_fv Plus sexp' M1) |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
440 |
|> do_notin_mtype_fv Minus sexp M1 |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
441 |
|> do_notin_mtype_fv Plus sexp M2 |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
442 |
| do_notin_mtype_fv Minus sexp (MFun (M1, V x, M2)) accum = |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
443 |
accum |> (case do_literal (x, Fls) (SOME sexp) of |
33192 | 444 |
NONE => I |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
445 |
| SOME sexp' => do_notin_mtype_fv Plus sexp' M1) |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
446 |
|> do_notin_mtype_fv Minus sexp M2 |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
447 |
| do_notin_mtype_fv sn sexp (MPair (M1, M2)) accum = |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
448 |
accum |> fold (do_notin_mtype_fv sn sexp) [M1, M2] |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
449 |
| do_notin_mtype_fv sn sexp (MType (_, Ms)) accum = |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
450 |
accum |> fold (do_notin_mtype_fv sn sexp) Ms |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
451 |
| do_notin_mtype_fv _ _ M _ = |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
452 |
raise MTYPE ("Nitpick_Mono.do_notin_mtype_fv", [M], []) |
33192 | 453 |
|
35832 | 454 |
fun add_notin_mtype_fv sn M ((lits, comps, sexps) : constraint_set) = |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
455 |
(trace_msg (fn () => "*** Add " ^ string_for_mtype M ^ " is " ^ |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
456 |
(case sn of Minus => "concrete" | Plus => "complete")); |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
457 |
case do_notin_mtype_fv sn [] M (SOME (lits, sexps)) of |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
458 |
NONE => (trace_msg (K "**** Unsolvable"); raise UNSOLVABLE ()) |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
459 |
| SOME (lits, sexps) => (lits, comps, sexps)) |
33192 | 460 |
|
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
461 |
val add_mtype_is_concrete = add_notin_mtype_fv Minus |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
462 |
val add_mtype_is_complete = add_notin_mtype_fv Plus |
33192 | 463 |
|
40986
cfd91aa80701
use boolean pair to encode annotation, which may now take four values
blanchet
parents:
40985
diff
changeset
|
464 |
fun fst_var n = 2 * n |
cfd91aa80701
use boolean pair to encode annotation, which may now take four values
blanchet
parents:
40985
diff
changeset
|
465 |
fun snd_var n = 2 * n + 1 |
35384 | 466 |
|
40986
cfd91aa80701
use boolean pair to encode annotation, which may now take four values
blanchet
parents:
40985
diff
changeset
|
467 |
val bool_table = |
cfd91aa80701
use boolean pair to encode annotation, which may now take four values
blanchet
parents:
40985
diff
changeset
|
468 |
[(Gen, (false, false)), |
cfd91aa80701
use boolean pair to encode annotation, which may now take four values
blanchet
parents:
40985
diff
changeset
|
469 |
(New, (false, true)), |
cfd91aa80701
use boolean pair to encode annotation, which may now take four values
blanchet
parents:
40985
diff
changeset
|
470 |
(Fls, (true, false)), |
cfd91aa80701
use boolean pair to encode annotation, which may now take four values
blanchet
parents:
40985
diff
changeset
|
471 |
(Tru, (true, true))] |
cfd91aa80701
use boolean pair to encode annotation, which may now take four values
blanchet
parents:
40985
diff
changeset
|
472 |
|
cfd91aa80701
use boolean pair to encode annotation, which may now take four values
blanchet
parents:
40985
diff
changeset
|
473 |
val bools_from_annotation = AList.lookup (op =) bool_table #> the |
cfd91aa80701
use boolean pair to encode annotation, which may now take four values
blanchet
parents:
40985
diff
changeset
|
474 |
val annotation_from_bools = AList.find (op =) bool_table #> the_single |
33192 | 475 |
|
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
476 |
fun prop_for_literal (x, a) = |
40986
cfd91aa80701
use boolean pair to encode annotation, which may now take four values
blanchet
parents:
40985
diff
changeset
|
477 |
let val (b1, b2) = bools_from_annotation a in |
cfd91aa80701
use boolean pair to encode annotation, which may now take four values
blanchet
parents:
40985
diff
changeset
|
478 |
PropLogic.SAnd (PropLogic.BoolVar (fst_var x) |> not b1 ? PropLogic.Not, |
cfd91aa80701
use boolean pair to encode annotation, which may now take four values
blanchet
parents:
40985
diff
changeset
|
479 |
PropLogic.BoolVar (snd_var x) |> not b2 ? PropLogic.Not) |
cfd91aa80701
use boolean pair to encode annotation, which may now take four values
blanchet
parents:
40985
diff
changeset
|
480 |
end |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
481 |
fun prop_for_annotation_atom_eq (A a', a) = |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
482 |
if a = a' then PropLogic.True else PropLogic.False |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
483 |
| prop_for_annotation_atom_eq (V x, a) = prop_for_literal (x, a) |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
484 |
fun prop_for_annotation_expr xs = PropLogic.exists (map prop_for_literal xs) |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
485 |
fun prop_for_exists_eq xs a = |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
486 |
PropLogic.exists (map (fn x => prop_for_literal (x, a)) xs) |
33192 | 487 |
fun prop_for_comp (a1, a2, Eq, []) = |
488 |
PropLogic.SAnd (prop_for_comp (a1, a2, Leq, []), |
|
489 |
prop_for_comp (a2, a1, Leq, [])) |
|
490 |
| prop_for_comp (a1, a2, Leq, []) = |
|
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
491 |
PropLogic.SOr (prop_for_annotation_atom_eq (a1, Fls), |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
492 |
prop_for_annotation_atom_eq (a2, Gen)) |
33192 | 493 |
| prop_for_comp (a1, a2, cmp, xs) = |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
494 |
PropLogic.SOr (prop_for_exists_eq xs Gen, prop_for_comp (a1, a2, cmp, [])) |
33192 | 495 |
|
40986
cfd91aa80701
use boolean pair to encode annotation, which may now take four values
blanchet
parents:
40985
diff
changeset
|
496 |
fun fix_bool_options (NONE, NONE) = (false, false) |
cfd91aa80701
use boolean pair to encode annotation, which may now take four values
blanchet
parents:
40985
diff
changeset
|
497 |
| fix_bool_options (NONE, SOME b) = (b, b) |
cfd91aa80701
use boolean pair to encode annotation, which may now take four values
blanchet
parents:
40985
diff
changeset
|
498 |
| fix_bool_options (SOME b, NONE) = (b, b) |
cfd91aa80701
use boolean pair to encode annotation, which may now take four values
blanchet
parents:
40985
diff
changeset
|
499 |
| fix_bool_options (SOME b1, SOME b2) = (b1, b2) |
cfd91aa80701
use boolean pair to encode annotation, which may now take four values
blanchet
parents:
40985
diff
changeset
|
500 |
|
34123
c4988215a691
distinguish better between "complete" (vs. incomplete) types and "concrete" (vs. abstract) types in Nitpick;
blanchet
parents:
34121
diff
changeset
|
501 |
fun literals_from_assignments max_var assigns lits = |
33192 | 502 |
fold (fn x => fn accum => |
503 |
if AList.defined (op =) lits x then |
|
504 |
accum |
|
40986
cfd91aa80701
use boolean pair to encode annotation, which may now take four values
blanchet
parents:
40985
diff
changeset
|
505 |
else case (fst_var x, snd_var x) |> pairself assigns of |
cfd91aa80701
use boolean pair to encode annotation, which may now take four values
blanchet
parents:
40985
diff
changeset
|
506 |
(NONE, NONE) => accum |
cfd91aa80701
use boolean pair to encode annotation, which may now take four values
blanchet
parents:
40985
diff
changeset
|
507 |
| bp => (x, annotation_from_bools (fix_bool_options bp)) :: accum) |
cfd91aa80701
use boolean pair to encode annotation, which may now take four values
blanchet
parents:
40985
diff
changeset
|
508 |
(max_var downto 1) lits |
33192 | 509 |
|
510 |
fun string_for_comp (a1, a2, cmp, xs) = |
|
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
511 |
string_for_annotation_atom a1 ^ " " ^ string_for_comp_op cmp ^ |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
512 |
subscript_string_for_vars " \<and> " xs ^ " " ^ string_for_annotation_atom a2 |
33192 | 513 |
|
514 |
fun print_problem lits comps sexps = |
|
38179 | 515 |
trace_msg (fn () => "*** Problem:\n" ^ |
516 |
cat_lines (map string_for_literal lits @ |
|
517 |
map string_for_comp comps @ |
|
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
518 |
map string_for_annotation_expr sexps)) |
33192 | 519 |
|
520 |
fun print_solution lits = |
|
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
521 |
let val (fs, gs) = List.partition (curry (op =) Fls o snd) lits in |
38179 | 522 |
trace_msg (fn () => "*** Solution:\n" ^ |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
523 |
"F: " ^ commas (map (string_for_var o fst) fs) ^ "\n" ^ |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
524 |
"G: " ^ commas (map (string_for_var o fst) gs)) |
33192 | 525 |
end |
526 |
||
35812
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
527 |
fun solve max_var (lits, comps, sexps) = |
40146 | 528 |
let |
529 |
fun do_assigns assigns = |
|
530 |
SOME (literals_from_assignments max_var assigns lits |
|
531 |
|> tap print_solution) |
|
532 |
val _ = print_problem lits comps sexps |
|
533 |
val prop = PropLogic.all (map prop_for_literal lits @ |
|
534 |
map prop_for_comp comps @ |
|
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
535 |
map prop_for_annotation_expr sexps) |
40986
cfd91aa80701
use boolean pair to encode annotation, which may now take four values
blanchet
parents:
40985
diff
changeset
|
536 |
val default_val = fst (bools_from_annotation Gen) |
40146 | 537 |
in |
538 |
if PropLogic.eval (K default_val) prop then |
|
539 |
do_assigns (K (SOME default_val)) |
|
540 |
else |
|
541 |
let |
|
542 |
(* use the first ML solver (to avoid startup overhead) *) |
|
543 |
val solvers = !SatSolver.solvers |
|
544 |
|> filter (member (op =) ["dptsat", "dpll"] o fst) |
|
545 |
in |
|
546 |
case snd (hd solvers) prop of |
|
547 |
SatSolver.SATISFIABLE assigns => do_assigns assigns |
|
548 |
| _ => NONE |
|
549 |
end |
|
550 |
end |
|
33192 | 551 |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
552 |
type mtype_schema = mtyp * constraint_set |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
553 |
type mtype_context = |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
554 |
{bound_Ts: typ list, |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
555 |
bound_Ms: mtyp list, |
40987
7d52af07bff1
added frame component to Gamma in monotonicity calculus
blanchet
parents:
40986
diff
changeset
|
556 |
bound_frame: (int * annotation_atom) list, |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
557 |
frees: (styp * mtyp) list, |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
558 |
consts: (styp * mtyp) list} |
33192 | 559 |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
560 |
type accumulator = mtype_context * constraint_set |
33192 | 561 |
|
40987
7d52af07bff1
added frame component to Gamma in monotonicity calculus
blanchet
parents:
40986
diff
changeset
|
562 |
val initial_gamma = |
7d52af07bff1
added frame component to Gamma in monotonicity calculus
blanchet
parents:
40986
diff
changeset
|
563 |
{bound_Ts = [], bound_Ms = [], bound_frame = [], frees = [], consts = []} |
33192 | 564 |
|
40987
7d52af07bff1
added frame component to Gamma in monotonicity calculus
blanchet
parents:
40986
diff
changeset
|
565 |
fun push_bound T M {bound_Ts, bound_Ms, bound_frame, frees, consts} = |
7d52af07bff1
added frame component to Gamma in monotonicity calculus
blanchet
parents:
40986
diff
changeset
|
566 |
{bound_Ts = T :: bound_Ts, bound_Ms = M :: bound_Ms, |
7d52af07bff1
added frame component to Gamma in monotonicity calculus
blanchet
parents:
40986
diff
changeset
|
567 |
bound_frame = bound_frame, frees = frees, consts = consts} |
7d52af07bff1
added frame component to Gamma in monotonicity calculus
blanchet
parents:
40986
diff
changeset
|
568 |
fun pop_bound {bound_Ts, bound_Ms, bound_frame, frees, consts} = |
7d52af07bff1
added frame component to Gamma in monotonicity calculus
blanchet
parents:
40986
diff
changeset
|
569 |
{bound_Ts = tl bound_Ts, bound_Ms = tl bound_Ms, bound_frame = bound_frame, |
7d52af07bff1
added frame component to Gamma in monotonicity calculus
blanchet
parents:
40986
diff
changeset
|
570 |
frees = frees, consts = consts} |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
571 |
handle List.Empty => initial_gamma (* FIXME: needed? *) |
33192 | 572 |
|
39359 | 573 |
fun consider_term (mdata as {hol_ctxt = {thy, ctxt, stds, ...}, alpha_T, |
574 |
max_fresh, ...}) = |
|
33192 | 575 |
let |
37476
0681e46b4022
optimized code generated for datatype cases + more;
blanchet
parents:
37267
diff
changeset
|
576 |
fun is_enough_eta_expanded t = |
0681e46b4022
optimized code generated for datatype cases + more;
blanchet
parents:
37267
diff
changeset
|
577 |
case strip_comb t of |
39359 | 578 |
(Const x, ts) => the_default 0 (arity_of_built_in_const thy stds x) |
37476
0681e46b4022
optimized code generated for datatype cases + more;
blanchet
parents:
37267
diff
changeset
|
579 |
<= length ts |
0681e46b4022
optimized code generated for datatype cases + more;
blanchet
parents:
37267
diff
changeset
|
580 |
| _ => true |
35672 | 581 |
val mtype_for = fresh_mtype_for_type mdata false |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
582 |
fun plus_set_mtype_for_dom M = |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
583 |
MFun (M, A (if exists_alpha_sub_mtype M then Fls else Gen), bool_M) |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
584 |
fun do_all T (gamma, cset) = |
33192 | 585 |
let |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
586 |
val abs_M = mtype_for (domain_type (domain_type T)) |
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
587 |
val body_M = mtype_for (body_type T) |
33192 | 588 |
in |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
589 |
(MFun (MFun (abs_M, A Gen, body_M), A Gen, body_M), |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
590 |
(gamma, cset |> add_mtype_is_complete abs_M)) |
33192 | 591 |
end |
592 |
fun do_equals T (gamma, cset) = |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
593 |
let val M = mtype_for (domain_type T) in |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
594 |
(MFun (M, A Gen, MFun (M, V (Unsynchronized.inc max_fresh), |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
595 |
mtype_for (nth_range_type 2 T))), |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
596 |
(gamma, cset |> add_mtype_is_concrete M)) |
33192 | 597 |
end |
598 |
fun do_robust_set_operation T (gamma, cset) = |
|
599 |
let |
|
600 |
val set_T = domain_type T |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
601 |
val M1 = mtype_for set_T |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
602 |
val M2 = mtype_for set_T |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
603 |
val M3 = mtype_for set_T |
33192 | 604 |
in |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
605 |
(MFun (M1, A Gen, MFun (M2, A Gen, M3)), |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
606 |
(gamma, cset |> add_is_sub_mtype M1 M3 |> add_is_sub_mtype M2 M3)) |
33192 | 607 |
end |
608 |
fun do_fragile_set_operation T (gamma, cset) = |
|
609 |
let |
|
610 |
val set_T = domain_type T |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
611 |
val set_M = mtype_for set_T |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
612 |
fun custom_mtype_for (T as Type (@{type_name fun}, [T1, T2])) = |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
613 |
if T = set_T then set_M |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
614 |
else MFun (custom_mtype_for T1, A Gen, custom_mtype_for T2) |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
615 |
| custom_mtype_for T = mtype_for T |
33192 | 616 |
in |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
617 |
(custom_mtype_for T, (gamma, cset |> add_mtype_is_concrete set_M)) |
33192 | 618 |
end |
619 |
fun do_pair_constr T accum = |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
620 |
case mtype_for (nth_range_type 2 T) of |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
621 |
M as MPair (a_M, b_M) => |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
622 |
(MFun (a_M, A Gen, MFun (b_M, A Gen, M)), accum) |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
623 |
| M => raise MTYPE ("Nitpick_Mono.consider_term.do_pair_constr", [M], []) |
33192 | 624 |
fun do_nth_pair_sel n T = |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
625 |
case mtype_for (domain_type T) of |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
626 |
M as MPair (a_M, b_M) => |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
627 |
pair (MFun (M, A Gen, if n = 0 then a_M else b_M)) |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
628 |
| M => raise MTYPE ("Nitpick_Mono.consider_term.do_nth_pair_sel", [M], []) |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
629 |
fun do_bounded_quantifier t0 abs_s abs_T connective_t bound_t body_t accum = |
33192 | 630 |
let |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
631 |
val abs_M = mtype_for abs_T |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
632 |
val (bound_m, accum) = |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
633 |
accum |>> push_bound abs_T abs_M |> do_term bound_t |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
634 |
val expected_bound_M = plus_set_mtype_for_dom abs_M |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
635 |
val (body_m, accum) = |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
636 |
accum ||> add_mtypes_equal expected_bound_M (mtype_of_mterm bound_m) |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
637 |
|> do_term body_t ||> apfst pop_bound |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
638 |
val bound_M = mtype_of_mterm bound_m |
39345 | 639 |
val (M1, a, _) = dest_MFun bound_M |
33192 | 640 |
in |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
641 |
(MApp (MRaw (t0, MFun (bound_M, A Gen, bool_M)), |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
642 |
MAbs (abs_s, abs_T, M1, a, |
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
643 |
MApp (MApp (MRaw (connective_t, |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
644 |
mtype_for (fastype_of connective_t)), |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
645 |
MApp (bound_m, MRaw (Bound 0, M1))), |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
646 |
body_m))), accum) |
33192 | 647 |
end |
40987
7d52af07bff1
added frame component to Gamma in monotonicity calculus
blanchet
parents:
40986
diff
changeset
|
648 |
and do_term t (accum as ({bound_Ts, bound_Ms, bound_frame, frees, consts}, |
7d52af07bff1
added frame component to Gamma in monotonicity calculus
blanchet
parents:
40986
diff
changeset
|
649 |
cset)) = |
38179 | 650 |
(trace_msg (fn () => " \<Gamma> \<turnstile> " ^ |
651 |
Syntax.string_of_term ctxt t ^ " : _?"); |
|
37704
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents:
37678
diff
changeset
|
652 |
case t of |
33192 | 653 |
Const (x as (s, T)) => |
654 |
(case AList.lookup (op =) consts x of |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
655 |
SOME M => (M, accum) |
33192 | 656 |
| NONE => |
657 |
if not (could_exist_alpha_subtype alpha_T T) then |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
658 |
(mtype_for T, accum) |
33192 | 659 |
else case s of |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
660 |
@{const_name all} => do_all T accum |
33192 | 661 |
| @{const_name "=="} => do_equals T accum |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
662 |
| @{const_name All} => do_all T accum |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
663 |
| @{const_name Ex} => |
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
664 |
let val set_T = domain_type T in |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
665 |
do_term (Abs (Name.uu, set_T, |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
666 |
@{const Not} $ (HOLogic.mk_eq |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
667 |
(Abs (Name.uu, domain_type set_T, |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
668 |
@{const False}), |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
669 |
Bound 0)))) accum |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
670 |
|>> mtype_of_mterm |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
671 |
end |
38864
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
haftmann
parents:
38795
diff
changeset
|
672 |
| @{const_name HOL.eq} => do_equals T accum |
37476
0681e46b4022
optimized code generated for datatype cases + more;
blanchet
parents:
37267
diff
changeset
|
673 |
| @{const_name The} => |
38179 | 674 |
(trace_msg (K "*** The"); raise UNSOLVABLE ()) |
37476
0681e46b4022
optimized code generated for datatype cases + more;
blanchet
parents:
37267
diff
changeset
|
675 |
| @{const_name Eps} => |
38179 | 676 |
(trace_msg (K "*** Eps"); raise UNSOLVABLE ()) |
33192 | 677 |
| @{const_name If} => |
678 |
do_robust_set_operation (range_type T) accum |
|
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
679 |
|>> curry3 MFun bool_M (A Gen) |
33192 | 680 |
| @{const_name Pair} => do_pair_constr T accum |
681 |
| @{const_name fst} => do_nth_pair_sel 0 T accum |
|
39316
b6c4385ab400
change defaults of Auto Nitpick so that it consumes less resources (time and Kodkod threads)
blanchet
parents:
38864
diff
changeset
|
682 |
| @{const_name snd} => do_nth_pair_sel 1 T accum |
33192 | 683 |
| @{const_name Id} => |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
684 |
(MFun (mtype_for (domain_type T), A Gen, bool_M), accum) |
33192 | 685 |
| @{const_name converse} => |
686 |
let |
|
687 |
val x = Unsynchronized.inc max_fresh |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
688 |
fun mtype_for_set T = |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
689 |
MFun (mtype_for (domain_type T), V x, bool_M) |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
690 |
val ab_set_M = domain_type T |> mtype_for_set |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
691 |
val ba_set_M = range_type T |> mtype_for_set |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
692 |
in (MFun (ab_set_M, A Gen, ba_set_M), accum) end |
33192 | 693 |
| @{const_name trancl} => do_fragile_set_operation T accum |
694 |
| @{const_name rel_comp} => |
|
695 |
let |
|
696 |
val x = Unsynchronized.inc max_fresh |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
697 |
fun mtype_for_set T = |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
698 |
MFun (mtype_for (domain_type T), V x, bool_M) |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
699 |
val bc_set_M = domain_type T |> mtype_for_set |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
700 |
val ab_set_M = domain_type (range_type T) |> mtype_for_set |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
701 |
val ac_set_M = nth_range_type 2 T |> mtype_for_set |
33192 | 702 |
in |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
703 |
(MFun (bc_set_M, A Gen, MFun (ab_set_M, A Gen, ac_set_M)), |
33192 | 704 |
accum) |
705 |
end |
|
706 |
| @{const_name image} => |
|
707 |
let |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
708 |
val a_M = mtype_for (domain_type (domain_type T)) |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
709 |
val b_M = mtype_for (range_type (domain_type T)) |
33192 | 710 |
in |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
711 |
(MFun (MFun (a_M, A Gen, b_M), A Gen, |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
712 |
MFun (plus_set_mtype_for_dom a_M, A Gen, |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
713 |
plus_set_mtype_for_dom b_M)), accum) |
33192 | 714 |
end |
35672 | 715 |
| @{const_name finite} => |
716 |
let val M1 = mtype_for (domain_type (domain_type T)) in |
|
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
717 |
(MFun (plus_set_mtype_for_dom M1, A Gen, bool_M), accum) |
35672 | 718 |
end |
33192 | 719 |
| @{const_name Sigma} => |
720 |
let |
|
721 |
val x = Unsynchronized.inc max_fresh |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
722 |
fun mtype_for_set T = |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
723 |
MFun (mtype_for (domain_type T), V x, bool_M) |
33192 | 724 |
val a_set_T = domain_type T |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
725 |
val a_M = mtype_for (domain_type a_set_T) |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
726 |
val b_set_M = mtype_for_set (range_type (domain_type |
33192 | 727 |
(range_type T))) |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
728 |
val a_set_M = mtype_for_set a_set_T |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
729 |
val a_to_b_set_M = MFun (a_M, A Gen, b_set_M) |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
730 |
val ab_set_M = mtype_for_set (nth_range_type 2 T) |
33192 | 731 |
in |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
732 |
(MFun (a_set_M, A Gen, MFun (a_to_b_set_M, A Gen, ab_set_M)), |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
733 |
accum) |
33192 | 734 |
end |
35220
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
735 |
| _ => |
39360
cdf2c3341422
eliminate more clutter related to "fast_descrs" optimization
blanchet
parents:
39359
diff
changeset
|
736 |
if s = @{const_name safe_The} then |
35671
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
737 |
let |
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
738 |
val a_set_M = mtype_for (domain_type T) |
ed2c3830d881
improved Nitpick's precision for "card" and "setsum" + fix incorrect outcome code w.r.t. "bisim_depth = -1"
blanchet
parents:
35665
diff
changeset
|
739 |
val a_M = dest_MFun a_set_M |> #1 |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
740 |
in (MFun (a_set_M, A Gen, a_M), accum) end |
35220
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
741 |
else if s = @{const_name ord_class.less_eq} andalso |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
742 |
is_set_type (domain_type T) then |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
743 |
do_fragile_set_operation T accum |
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
744 |
else if is_sel s then |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
745 |
(mtype_for_sel mdata x, accum) |
37256
0dca1ec52999
thread along context instead of theory for typedef lookup
blanchet
parents:
36385
diff
changeset
|
746 |
else if is_constr ctxt stds x then |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
747 |
(mtype_for_constr mdata x, accum) |
39359 | 748 |
else if is_built_in_const thy stds x then |
35672 | 749 |
(fresh_mtype_for_type mdata true T, accum) |
35220
2bcdae5f4fdb
added support for nonstandard "nat"s to Nitpick and fixed bugs in binary "nat"s and "int"s
blanchet
parents:
35219
diff
changeset
|
750 |
else |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
751 |
let val M = mtype_for T in |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
752 |
(M, ({bound_Ts = bound_Ts, bound_Ms = bound_Ms, |
40987
7d52af07bff1
added frame component to Gamma in monotonicity calculus
blanchet
parents:
40986
diff
changeset
|
753 |
bound_frame = bound_frame, frees = frees, |
7d52af07bff1
added frame component to Gamma in monotonicity calculus
blanchet
parents:
40986
diff
changeset
|
754 |
consts = (x, M) :: consts}, cset)) |
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
755 |
end) |>> curry MRaw t |
33192 | 756 |
| Free (x as (_, T)) => |
757 |
(case AList.lookup (op =) frees x of |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
758 |
SOME M => (M, accum) |
33192 | 759 |
| NONE => |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
760 |
let val M = mtype_for T in |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
761 |
(M, ({bound_Ts = bound_Ts, bound_Ms = bound_Ms, |
40987
7d52af07bff1
added frame component to Gamma in monotonicity calculus
blanchet
parents:
40986
diff
changeset
|
762 |
bound_frame = bound_frame, frees = (x, M) :: frees, |
7d52af07bff1
added frame component to Gamma in monotonicity calculus
blanchet
parents:
40986
diff
changeset
|
763 |
consts = consts}, cset)) |
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
764 |
end) |>> curry MRaw t |
38179 | 765 |
| Var _ => (trace_msg (K "*** Var"); raise UNSOLVABLE ()) |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
766 |
| Bound j => (MRaw (t, nth bound_Ms j), accum) |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
767 |
| Abs (s, T, t') => |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
768 |
(case fin_fun_body T (fastype_of1 (T :: bound_Ts, t')) t' of |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
769 |
SOME t' => |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
770 |
let |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
771 |
val M = mtype_for T |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
772 |
val a = V (Unsynchronized.inc max_fresh) |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
773 |
val (m', accum) = do_term t' (accum |>> push_bound T M) |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
774 |
in (MAbs (s, T, M, a, m'), accum |>> pop_bound) end |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
775 |
| NONE => |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
776 |
((case t' of |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
777 |
t1' $ Bound 0 => |
37476
0681e46b4022
optimized code generated for datatype cases + more;
blanchet
parents:
37267
diff
changeset
|
778 |
if not (loose_bvar1 (t1', 0)) andalso |
0681e46b4022
optimized code generated for datatype cases + more;
blanchet
parents:
37267
diff
changeset
|
779 |
is_enough_eta_expanded t1' then |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
780 |
do_term (incr_boundvars ~1 t1') accum |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
781 |
else |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
782 |
raise SAME () |
38864
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
haftmann
parents:
38795
diff
changeset
|
783 |
| (t11 as Const (@{const_name HOL.eq}, _)) $ Bound 0 $ t13 => |
37476
0681e46b4022
optimized code generated for datatype cases + more;
blanchet
parents:
37267
diff
changeset
|
784 |
if not (loose_bvar1 (t13, 0)) then |
0681e46b4022
optimized code generated for datatype cases + more;
blanchet
parents:
37267
diff
changeset
|
785 |
do_term (incr_boundvars ~1 (t11 $ t13)) accum |
0681e46b4022
optimized code generated for datatype cases + more;
blanchet
parents:
37267
diff
changeset
|
786 |
else |
0681e46b4022
optimized code generated for datatype cases + more;
blanchet
parents:
37267
diff
changeset
|
787 |
raise SAME () |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
788 |
| _ => raise SAME ()) |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
789 |
handle SAME () => |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
790 |
let |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
791 |
val M = mtype_for T |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
792 |
val (m', accum) = do_term t' (accum |>> push_bound T M) |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
793 |
in (MAbs (s, T, M, A Gen, m'), accum |>> pop_bound) end)) |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
794 |
| (t0 as Const (@{const_name All}, _)) |
38786
e46e7a9cb622
formerly unnamed infix impliciation now named HOL.implies
haftmann
parents:
38199
diff
changeset
|
795 |
$ Abs (s', T', (t10 as @{const HOL.implies}) $ (t11 $ Bound 0) $ t12) => |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
796 |
do_bounded_quantifier t0 s' T' t10 t11 t12 accum |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
797 |
| (t0 as Const (@{const_name Ex}, _)) |
38795
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents:
38786
diff
changeset
|
798 |
$ Abs (s', T', (t10 as @{const HOL.conj}) $ (t11 $ Bound 0) $ t12) => |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
799 |
do_bounded_quantifier t0 s' T' t10 t11 t12 accum |
33192 | 800 |
| Const (@{const_name Let}, _) $ t1 $ t2 => |
801 |
do_term (betapply (t2, t1)) accum |
|
802 |
| t1 $ t2 => |
|
803 |
let |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
804 |
val (m1, accum) = do_term t1 accum |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
805 |
val (m2, accum) = do_term t2 accum |
33192 | 806 |
in |
35812
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
807 |
let |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
808 |
val M11 = mtype_of_mterm m1 |> dest_MFun |> #1 |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
809 |
val M2 = mtype_of_mterm m2 |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
810 |
in (MApp (m1, m2), accum ||> add_is_sub_mtype M2 M11) end |
33192 | 811 |
end) |
38179 | 812 |
|> tap (fn (m, _) => trace_msg (fn () => " \<Gamma> \<turnstile> " ^ |
813 |
string_for_mterm ctxt m)) |
|
33192 | 814 |
in do_term end |
815 |
||
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
816 |
fun force_minus_funs 0 _ = I |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
817 |
| force_minus_funs n (M as MFun (M1, _, M2)) = |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
818 |
add_mtypes_equal M (MFun (M1, A Gen, M2)) #> force_minus_funs (n - 1) M2 |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
819 |
| force_minus_funs _ M = |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
820 |
raise MTYPE ("Nitpick_Mono.force_minus_funs", [M], []) |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
821 |
fun consider_general_equals mdata def (x as (_, T)) t1 t2 accum = |
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
822 |
let |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
823 |
val (m1, accum) = consider_term mdata t1 accum |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
824 |
val (m2, accum) = consider_term mdata t2 accum |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
825 |
val M1 = mtype_of_mterm m1 |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
826 |
val M2 = mtype_of_mterm m2 |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
827 |
val accum = accum ||> add_mtypes_equal M1 M2 |
35672 | 828 |
val body_M = fresh_mtype_for_type mdata false (nth_range_type 2 T) |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
829 |
val m = MApp (MApp (MRaw (Const x, |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
830 |
MFun (M1, A Gen, MFun (M2, A Gen, body_M))), m1), m2) |
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
831 |
in |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
832 |
(m, if def then |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
833 |
let val (head_m, arg_ms) = strip_mcomb m1 in |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
834 |
accum ||> force_minus_funs (length arg_ms) (mtype_of_mterm head_m) |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
835 |
end |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
836 |
else |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
837 |
accum) |
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
838 |
end |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
839 |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
840 |
fun consider_general_formula (mdata as {hol_ctxt = {ctxt, ...}, ...}) = |
33192 | 841 |
let |
35672 | 842 |
val mtype_for = fresh_mtype_for_type mdata false |
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
843 |
val do_term = consider_term mdata |
35812
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
844 |
fun do_formula sn t accum = |
33192 | 845 |
let |
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
846 |
fun do_quantifier (quant_x as (quant_s, _)) abs_s abs_T body_t = |
33192 | 847 |
let |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
848 |
val abs_M = mtype_for abs_T |
34982
7b8c366e34a2
added support for nonstandard models to Nitpick (based on an idea by Koen Claessen) and did other fixes to Nitpick
blanchet
parents:
34936
diff
changeset
|
849 |
val side_cond = ((sn = Minus) = (quant_s = @{const_name Ex})) |
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
850 |
val (body_m, accum) = |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
851 |
accum ||> side_cond ? add_mtype_is_complete abs_M |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
852 |
|>> push_bound abs_T abs_M |> do_formula sn body_t |
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
853 |
val body_M = mtype_of_mterm body_m |
33192 | 854 |
in |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
855 |
(MApp (MRaw (Const quant_x, |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
856 |
MFun (MFun (abs_M, A Gen, body_M), A Gen, body_M)), |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
857 |
MAbs (abs_s, abs_T, abs_M, A Gen, body_m)), |
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
858 |
accum |>> pop_bound) |
33192 | 859 |
end |
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
860 |
fun do_equals x t1 t2 = |
33192 | 861 |
case sn of |
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
862 |
Plus => do_term t accum |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
863 |
| Minus => consider_general_equals mdata false x t1 t2 accum |
33192 | 864 |
in |
38179 | 865 |
(trace_msg (fn () => " \<Gamma> \<turnstile> " ^ |
866 |
Syntax.string_of_term ctxt t ^ " : o\<^sup>" ^ |
|
867 |
string_for_sign sn ^ "?"); |
|
37704
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents:
37678
diff
changeset
|
868 |
case t of |
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents:
37678
diff
changeset
|
869 |
Const (x as (@{const_name all}, _)) $ Abs (s1, T1, t1) => |
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents:
37678
diff
changeset
|
870 |
do_quantifier x s1 T1 t1 |
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents:
37678
diff
changeset
|
871 |
| Const (x as (@{const_name "=="}, _)) $ t1 $ t2 => do_equals x t1 t2 |
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents:
37678
diff
changeset
|
872 |
| @{const Trueprop} $ t1 => |
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents:
37678
diff
changeset
|
873 |
let val (m1, accum) = do_formula sn t1 accum in |
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents:
37678
diff
changeset
|
874 |
(MApp (MRaw (@{const Trueprop}, mtype_for (bool_T --> prop_T)), |
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents:
37678
diff
changeset
|
875 |
m1), accum) |
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents:
37678
diff
changeset
|
876 |
end |
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents:
37678
diff
changeset
|
877 |
| @{const Not} $ t1 => |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
878 |
let val (m1, accum) = do_formula (negate_sign sn) t1 accum in |
37704
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents:
37678
diff
changeset
|
879 |
(MApp (MRaw (@{const Not}, mtype_for (bool_T --> bool_T)), m1), |
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents:
37678
diff
changeset
|
880 |
accum) |
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents:
37678
diff
changeset
|
881 |
end |
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents:
37678
diff
changeset
|
882 |
| Const (x as (@{const_name All}, _)) $ Abs (s1, T1, t1) => |
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents:
37678
diff
changeset
|
883 |
do_quantifier x s1 T1 t1 |
39345 | 884 |
| Const (x0 as (@{const_name Ex}, T0)) |
37704
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents:
37678
diff
changeset
|
885 |
$ (t1 as Abs (s1, T1, t1')) => |
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents:
37678
diff
changeset
|
886 |
(case sn of |
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents:
37678
diff
changeset
|
887 |
Plus => do_quantifier x0 s1 T1 t1' |
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents:
37678
diff
changeset
|
888 |
| Minus => |
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents:
37678
diff
changeset
|
889 |
(* FIXME: Move elsewhere *) |
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents:
37678
diff
changeset
|
890 |
do_term (@{const Not} |
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents:
37678
diff
changeset
|
891 |
$ (HOLogic.eq_const (domain_type T0) $ t1 |
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents:
37678
diff
changeset
|
892 |
$ Abs (Name.uu, T1, @{const False}))) accum) |
38864
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
haftmann
parents:
38795
diff
changeset
|
893 |
| Const (x as (@{const_name HOL.eq}, _)) $ t1 $ t2 => |
37704
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents:
37678
diff
changeset
|
894 |
do_equals x t1 t2 |
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents:
37678
diff
changeset
|
895 |
| Const (@{const_name Let}, _) $ t1 $ t2 => |
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents:
37678
diff
changeset
|
896 |
do_formula sn (betapply (t2, t1)) accum |
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents:
37678
diff
changeset
|
897 |
| (t0 as Const (s0, _)) $ t1 $ t2 => |
38199
8a9cace789d6
deal correctly with Pure.conjunction in mono check
blanchet
parents:
38190
diff
changeset
|
898 |
if s0 = @{const_name "==>"} orelse |
8a9cace789d6
deal correctly with Pure.conjunction in mono check
blanchet
parents:
38190
diff
changeset
|
899 |
s0 = @{const_name Pure.conjunction} orelse |
38795
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents:
38786
diff
changeset
|
900 |
s0 = @{const_name HOL.conj} orelse |
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents:
38786
diff
changeset
|
901 |
s0 = @{const_name HOL.disj} orelse |
38786
e46e7a9cb622
formerly unnamed infix impliciation now named HOL.implies
haftmann
parents:
38199
diff
changeset
|
902 |
s0 = @{const_name HOL.implies} then |
37704
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents:
37678
diff
changeset
|
903 |
let |
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents:
37678
diff
changeset
|
904 |
val impl = (s0 = @{const_name "==>"} orelse |
38786
e46e7a9cb622
formerly unnamed infix impliciation now named HOL.implies
haftmann
parents:
38199
diff
changeset
|
905 |
s0 = @{const_name HOL.implies}) |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
906 |
val (m1, accum) = |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
907 |
do_formula (sn |> impl ? negate_sign) t1 accum |
37704
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents:
37678
diff
changeset
|
908 |
val (m2, accum) = do_formula sn t2 accum |
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents:
37678
diff
changeset
|
909 |
in |
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents:
37678
diff
changeset
|
910 |
(MApp (MApp (MRaw (t0, mtype_for (fastype_of t0)), m1), m2), |
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents:
37678
diff
changeset
|
911 |
accum) |
39316
b6c4385ab400
change defaults of Auto Nitpick so that it consumes less resources (time and Kodkod threads)
blanchet
parents:
38864
diff
changeset
|
912 |
end |
37704
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents:
37678
diff
changeset
|
913 |
else |
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents:
37678
diff
changeset
|
914 |
do_term t accum |
c6161bee8486
adapt Nitpick to "prod_case" and "*" -> "sum" renaming;
blanchet
parents:
37678
diff
changeset
|
915 |
| _ => do_term t accum) |
33192 | 916 |
end |
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
917 |
|> tap (fn (m, _) => |
38179 | 918 |
trace_msg (fn () => "\<Gamma> \<turnstile> " ^ |
919 |
string_for_mterm ctxt m ^ " : o\<^sup>" ^ |
|
920 |
string_for_sign sn)) |
|
33192 | 921 |
in do_formula end |
922 |
||
923 |
(* The harmless axiom optimization below is somewhat too aggressive in the face |
|
924 |
of (rather peculiar) user-defined axioms. *) |
|
925 |
val harmless_consts = |
|
926 |
[@{const_name ord_class.less}, @{const_name ord_class.less_eq}] |
|
927 |
val bounteous_consts = [@{const_name bisim}] |
|
928 |
||
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
929 |
fun is_harmless_axiom ({no_harmless = true, ...} : mdata) _ = false |
39359 | 930 |
| is_harmless_axiom {hol_ctxt = {thy, stds, ...}, ...} t = |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
931 |
Term.add_consts t [] |
39359 | 932 |
|> filter_out (is_built_in_const thy stds) |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
933 |
|> (forall (member (op =) harmless_consts o original_name o fst) orf |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
934 |
exists (member (op =) bounteous_consts o fst)) |
33192 | 935 |
|
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
936 |
fun consider_nondefinitional_axiom mdata t = |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
937 |
if is_harmless_axiom mdata t then pair (MRaw (t, dummy_M)) |
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
938 |
else consider_general_formula mdata Plus t |
33192 | 939 |
|
37256
0dca1ec52999
thread along context instead of theory for typedef lookup
blanchet
parents:
36385
diff
changeset
|
940 |
fun consider_definitional_axiom (mdata as {hol_ctxt = {ctxt, ...}, ...}) t = |
0dca1ec52999
thread along context instead of theory for typedef lookup
blanchet
parents:
36385
diff
changeset
|
941 |
if not (is_constr_pattern_formula ctxt t) then |
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
942 |
consider_nondefinitional_axiom mdata t |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
943 |
else if is_harmless_axiom mdata t then |
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
944 |
pair (MRaw (t, dummy_M)) |
33192 | 945 |
else |
946 |
let |
|
35672 | 947 |
val mtype_for = fresh_mtype_for_type mdata false |
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
948 |
val do_term = consider_term mdata |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
949 |
fun do_all quant_t abs_s abs_T body_t accum = |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
950 |
let |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
951 |
val abs_M = mtype_for abs_T |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
952 |
val (body_m, accum) = |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
953 |
accum |>> push_bound abs_T abs_M |> do_formula body_t |
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
954 |
val body_M = mtype_of_mterm body_m |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
955 |
in |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
956 |
(MApp (MRaw (quant_t, MFun (MFun (abs_M, A Gen, body_M), A Gen, |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
957 |
body_M)), |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
958 |
MAbs (abs_s, abs_T, abs_M, A Gen, body_m)), |
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
959 |
accum |>> pop_bound) |
33192 | 960 |
end |
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
961 |
and do_conjunction t0 t1 t2 accum = |
33192 | 962 |
let |
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
963 |
val (m1, accum) = do_formula t1 accum |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
964 |
val (m2, accum) = do_formula t2 accum |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
965 |
in |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
966 |
(MApp (MApp (MRaw (t0, mtype_for (fastype_of t0)), m1), m2), accum) |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
967 |
end |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
968 |
and do_implies t0 t1 t2 accum = |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
969 |
let |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
970 |
val (m1, accum) = do_term t1 accum |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
971 |
val (m2, accum) = do_formula t2 accum |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
972 |
in |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
973 |
(MApp (MApp (MRaw (t0, mtype_for (fastype_of t0)), m1), m2), accum) |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
974 |
end |
35812
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
975 |
and do_formula t accum = |
33192 | 976 |
case t of |
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
977 |
(t0 as Const (@{const_name all}, _)) $ Abs (s1, T1, t1) => |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
978 |
do_all t0 s1 T1 t1 accum |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
979 |
| @{const Trueprop} $ t1 => |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
980 |
let val (m1, accum) = do_formula t1 accum in |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
981 |
(MApp (MRaw (@{const Trueprop}, mtype_for (bool_T --> prop_T)), |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
982 |
m1), accum) |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
983 |
end |
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
984 |
| Const (x as (@{const_name "=="}, _)) $ t1 $ t2 => |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
985 |
consider_general_equals mdata true x t1 t2 accum |
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
986 |
| (t0 as @{const "==>"}) $ t1 $ t2 => do_implies t0 t1 t2 accum |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
987 |
| (t0 as @{const Pure.conjunction}) $ t1 $ t2 => |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
988 |
do_conjunction t0 t1 t2 accum |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
989 |
| (t0 as Const (@{const_name All}, _)) $ Abs (s0, T1, t1) => |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
990 |
do_all t0 s0 T1 t1 accum |
38864
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
haftmann
parents:
38795
diff
changeset
|
991 |
| Const (x as (@{const_name HOL.eq}, _)) $ t1 $ t2 => |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
992 |
consider_general_equals mdata true x t1 t2 accum |
38795
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents:
38786
diff
changeset
|
993 |
| (t0 as @{const HOL.conj}) $ t1 $ t2 => do_conjunction t0 t1 t2 accum |
38786
e46e7a9cb622
formerly unnamed infix impliciation now named HOL.implies
haftmann
parents:
38199
diff
changeset
|
994 |
| (t0 as @{const HOL.implies}) $ t1 $ t2 => do_implies t0 t1 t2 accum |
33232
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33192
diff
changeset
|
995 |
| _ => raise TERM ("Nitpick_Mono.consider_definitional_axiom.\ |
33192 | 996 |
\do_formula", [t]) |
997 |
in do_formula t end |
|
998 |
||
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
999 |
fun string_for_mtype_of_term ctxt lits t M = |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
1000 |
Syntax.string_of_term ctxt t ^ " : " ^ string_for_mtype (resolve_mtype lits M) |
33192 | 1001 |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
1002 |
fun print_mtype_context ctxt lits ({frees, consts, ...} : mtype_context) = |
38179 | 1003 |
trace_msg (fn () => |
37476
0681e46b4022
optimized code generated for datatype cases + more;
blanchet
parents:
37267
diff
changeset
|
1004 |
map (fn (x, M) => string_for_mtype_of_term ctxt lits (Free x) M) frees @ |
0681e46b4022
optimized code generated for datatype cases + more;
blanchet
parents:
37267
diff
changeset
|
1005 |
map (fn (x, M) => string_for_mtype_of_term ctxt lits (Const x) M) consts |
0681e46b4022
optimized code generated for datatype cases + more;
blanchet
parents:
37267
diff
changeset
|
1006 |
|> cat_lines) |
33192 | 1007 |
|
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
1008 |
fun amass f t (ms, accum) = |
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
1009 |
let val (m, accum) = f t accum in (m :: ms, accum) end |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
1010 |
|
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
1011 |
fun infer which no_harmless (hol_ctxt as {ctxt, ...}) binarize alpha_T |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
1012 |
(nondef_ts, def_ts) = |
33192 | 1013 |
let |
38179 | 1014 |
val _ = trace_msg (fn () => "****** " ^ which ^ " analysis: " ^ |
1015 |
string_for_mtype MAlpha ^ " is " ^ |
|
1016 |
Syntax.string_of_typ ctxt alpha_T) |
|
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
1017 |
val mdata as {max_fresh, constr_mcache, ...} = |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
1018 |
initial_mdata hol_ctxt binarize no_harmless alpha_T |
35812
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1019 |
val accum = (initial_gamma, ([], [], [])) |
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
1020 |
val (nondef_ms, accum) = |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
1021 |
([], accum) |> amass (consider_general_formula mdata Plus) (hd nondef_ts) |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
1022 |
|> fold (amass (consider_nondefinitional_axiom mdata)) |
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
1023 |
(tl nondef_ts) |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
1024 |
val (def_ms, (gamma, cset)) = |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
1025 |
([], accum) |> fold (amass (consider_definitional_axiom mdata)) def_ts |
33192 | 1026 |
in |
1027 |
case solve (!max_fresh) cset of |
|
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
1028 |
SOME lits => (print_mtype_context ctxt lits gamma; |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
1029 |
SOME (lits, (nondef_ms, def_ms), !constr_mcache)) |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
1030 |
| _ => NONE |
33192 | 1031 |
end |
35812
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1032 |
handle UNSOLVABLE () => NONE |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1033 |
| MTYPE (loc, Ms, Ts) => |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
1034 |
raise BAD (loc, commas (map string_for_mtype Ms @ |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
1035 |
map (Syntax.string_of_typ ctxt) Ts)) |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
1036 |
| MTERM (loc, ms) => |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
1037 |
raise BAD (loc, commas (map (string_for_mterm ctxt) ms)) |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
1038 |
|
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
1039 |
val formulas_monotonic = is_some oooo infer "Monotonicity" false |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
1040 |
|
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
1041 |
fun fin_fun_constr T1 T2 = |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
1042 |
(@{const_name FinFun}, (T1 --> T2) --> Type (@{type_name fin_fun}, [T1, T2])) |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
1043 |
|
39359 | 1044 |
fun finitize_funs (hol_ctxt as {thy, ctxt, stds, constr_cache, ...}) |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
1045 |
binarize finitizes alpha_T tsp = |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
1046 |
case infer "Finiteness" true hol_ctxt binarize alpha_T tsp of |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
1047 |
SOME (lits, msp, constr_mtypes) => |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
1048 |
if forall (curry (op =) Gen o snd) lits then |
35812
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1049 |
tsp |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1050 |
else |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1051 |
let |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1052 |
fun should_finitize T a = |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1053 |
case triple_lookup (type_match thy) finitizes T of |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1054 |
SOME (SOME false) => false |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
1055 |
| _ => resolve_annotation_atom lits a = A Fls |
35812
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1056 |
fun type_from_mtype T M = |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1057 |
case (M, T) of |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1058 |
(MAlpha, _) => T |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1059 |
| (MFun (M1, a, M2), Type (@{type_name fun}, Ts)) => |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1060 |
Type (if should_finitize T a then @{type_name fin_fun} |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1061 |
else @{type_name fun}, map2 type_from_mtype Ts [M1, M2]) |
38190
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
blanchet
parents:
38179
diff
changeset
|
1062 |
| (MPair (M1, M2), Type (@{type_name prod}, Ts)) => |
b02e204b613a
get rid of all "optimizations" regarding "unit" and other cardinality-1 types
blanchet
parents:
38179
diff
changeset
|
1063 |
Type (@{type_name prod}, map2 type_from_mtype Ts [M1, M2]) |
35812
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1064 |
| (MType _, _) => T |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1065 |
| _ => raise MTYPE ("Nitpick_Mono.finitize_funs.type_from_mtype", |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1066 |
[M], [T]) |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1067 |
fun finitize_constr (x as (s, T)) = |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1068 |
(s, case AList.lookup (op =) constr_mtypes x of |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1069 |
SOME M => type_from_mtype T M |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1070 |
| NONE => T) |
37267
5c47d633c84d
fix code that used to raise an exception if bound variables were given a finite function type, because the old vs. new bound variable types were confused
blanchet
parents:
37260
diff
changeset
|
1071 |
fun term_from_mterm new_Ts old_Ts m = |
35812
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1072 |
case m of |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1073 |
MRaw (t, M) => |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1074 |
let |
37267
5c47d633c84d
fix code that used to raise an exception if bound variables were given a finite function type, because the old vs. new bound variable types were confused
blanchet
parents:
37260
diff
changeset
|
1075 |
val T = fastype_of1 (old_Ts, t) |
35812
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1076 |
val T' = type_from_mtype T M |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1077 |
in |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1078 |
case t of |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1079 |
Const (x as (s, _)) => |
37476
0681e46b4022
optimized code generated for datatype cases + more;
blanchet
parents:
37267
diff
changeset
|
1080 |
if s = @{const_name finite} then |
35812
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1081 |
case domain_type T' of |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1082 |
set_T' as Type (@{type_name fin_fun}, _) => |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1083 |
Abs (Name.uu, set_T', @{const True}) |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1084 |
| _ => Const (s, T') |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1085 |
else if s = @{const_name "=="} orelse |
38864
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
haftmann
parents:
38795
diff
changeset
|
1086 |
s = @{const_name HOL.eq} then |
37476
0681e46b4022
optimized code generated for datatype cases + more;
blanchet
parents:
37267
diff
changeset
|
1087 |
let |
0681e46b4022
optimized code generated for datatype cases + more;
blanchet
parents:
37267
diff
changeset
|
1088 |
val T = |
0681e46b4022
optimized code generated for datatype cases + more;
blanchet
parents:
37267
diff
changeset
|
1089 |
case T' of |
0681e46b4022
optimized code generated for datatype cases + more;
blanchet
parents:
37267
diff
changeset
|
1090 |
Type (_, [T1, Type (_, [T2, T3])]) => |
0681e46b4022
optimized code generated for datatype cases + more;
blanchet
parents:
37267
diff
changeset
|
1091 |
T1 --> T2 --> T3 |
0681e46b4022
optimized code generated for datatype cases + more;
blanchet
parents:
37267
diff
changeset
|
1092 |
| _ => raise TYPE ("Nitpick_Mono.finitize_funs.\ |
0681e46b4022
optimized code generated for datatype cases + more;
blanchet
parents:
37267
diff
changeset
|
1093 |
\term_from_mterm", [T, T'], []) |
0681e46b4022
optimized code generated for datatype cases + more;
blanchet
parents:
37267
diff
changeset
|
1094 |
in coerce_term hol_ctxt new_Ts T' T (Const (s, T)) end |
39359 | 1095 |
else if is_built_in_const thy stds x then |
37267
5c47d633c84d
fix code that used to raise an exception if bound variables were given a finite function type, because the old vs. new bound variable types were confused
blanchet
parents:
37260
diff
changeset
|
1096 |
coerce_term hol_ctxt new_Ts T' T t |
37256
0dca1ec52999
thread along context instead of theory for typedef lookup
blanchet
parents:
36385
diff
changeset
|
1097 |
else if is_constr ctxt stds x then |
35812
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1098 |
Const (finitize_constr x) |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1099 |
else if is_sel s then |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1100 |
let |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1101 |
val n = sel_no_from_name s |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1102 |
val x' = |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1103 |
x |> binarized_and_boxed_constr_for_sel hol_ctxt binarize |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1104 |
|> finitize_constr |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1105 |
val x'' = |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1106 |
binarized_and_boxed_nth_sel_for_constr hol_ctxt binarize |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1107 |
x' n |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1108 |
in Const x'' end |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1109 |
else |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1110 |
Const (s, T') |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1111 |
| Free (s, T) => Free (s, type_from_mtype T M) |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1112 |
| Bound _ => t |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1113 |
| _ => raise MTERM ("Nitpick_Mono.finitize_funs.term_from_mterm", |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1114 |
[m]) |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1115 |
end |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1116 |
| MApp (m1, m2) => |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1117 |
let |
37267
5c47d633c84d
fix code that used to raise an exception if bound variables were given a finite function type, because the old vs. new bound variable types were confused
blanchet
parents:
37260
diff
changeset
|
1118 |
val (t1, t2) = pairself (term_from_mterm new_Ts old_Ts) (m1, m2) |
5c47d633c84d
fix code that used to raise an exception if bound variables were given a finite function type, because the old vs. new bound variable types were confused
blanchet
parents:
37260
diff
changeset
|
1119 |
val (T1, T2) = pairself (curry fastype_of1 new_Ts) (t1, t2) |
35812
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1120 |
val (t1', T2') = |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1121 |
case T1 of |
39316
b6c4385ab400
change defaults of Auto Nitpick so that it consumes less resources (time and Kodkod threads)
blanchet
parents:
38864
diff
changeset
|
1122 |
Type (s, [T11, T12]) => |
35812
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1123 |
(if s = @{type_name fin_fun} then |
37256
0dca1ec52999
thread along context instead of theory for typedef lookup
blanchet
parents:
36385
diff
changeset
|
1124 |
select_nth_constr_arg ctxt stds (fin_fun_constr T11 T12) t1 |
35812
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1125 |
0 (T11 --> T12) |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1126 |
else |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1127 |
t1, T11) |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1128 |
| _ => raise TYPE ("Nitpick_Mono.finitize_funs.term_from_mterm", |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1129 |
[T1], []) |
37267
5c47d633c84d
fix code that used to raise an exception if bound variables were given a finite function type, because the old vs. new bound variable types were confused
blanchet
parents:
37260
diff
changeset
|
1130 |
in betapply (t1', coerce_term hol_ctxt new_Ts T2' T2 t2) end |
5c47d633c84d
fix code that used to raise an exception if bound variables were given a finite function type, because the old vs. new bound variable types were confused
blanchet
parents:
37260
diff
changeset
|
1131 |
| MAbs (s, old_T, M, a, m') => |
35812
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1132 |
let |
37267
5c47d633c84d
fix code that used to raise an exception if bound variables were given a finite function type, because the old vs. new bound variable types were confused
blanchet
parents:
37260
diff
changeset
|
1133 |
val new_T = type_from_mtype old_T M |
5c47d633c84d
fix code that used to raise an exception if bound variables were given a finite function type, because the old vs. new bound variable types were confused
blanchet
parents:
37260
diff
changeset
|
1134 |
val t' = term_from_mterm (new_T :: new_Ts) (old_T :: old_Ts) m' |
5c47d633c84d
fix code that used to raise an exception if bound variables were given a finite function type, because the old vs. new bound variable types were confused
blanchet
parents:
37260
diff
changeset
|
1135 |
val T' = fastype_of1 (new_T :: new_Ts, t') |
35812
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1136 |
in |
37267
5c47d633c84d
fix code that used to raise an exception if bound variables were given a finite function type, because the old vs. new bound variable types were confused
blanchet
parents:
37260
diff
changeset
|
1137 |
Abs (s, new_T, t') |
5c47d633c84d
fix code that used to raise an exception if bound variables were given a finite function type, because the old vs. new bound variable types were confused
blanchet
parents:
37260
diff
changeset
|
1138 |
|> should_finitize (new_T --> T') a |
5c47d633c84d
fix code that used to raise an exception if bound variables were given a finite function type, because the old vs. new bound variable types were confused
blanchet
parents:
37260
diff
changeset
|
1139 |
? construct_value ctxt stds (fin_fun_constr new_T T') o single |
35812
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1140 |
end |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1141 |
in |
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1142 |
Unsynchronized.change constr_cache (map (apsnd (map finitize_constr))); |
37267
5c47d633c84d
fix code that used to raise an exception if bound variables were given a finite function type, because the old vs. new bound variable types were confused
blanchet
parents:
37260
diff
changeset
|
1143 |
pairself (map (term_from_mterm [] [])) msp |
35812
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1144 |
end |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
1145 |
| NONE => tsp |
33192 | 1146 |
|
1147 |
end; |