author | boehmes |
Sun, 04 May 2014 16:17:53 +0200 | |
changeset 56845 | 691da43fbdd4 |
parent 56815 | 848d507584db |
child 56851 | 35ff4ede3409 |
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 : |
41001
11715564e2ad
removed old baggage from monotonicity calculus -- the "calculus" option didn't really work anyway because of on-the-fly simplifications
blanchet
parents:
41000
diff
changeset
|
14 |
hol_context -> bool -> typ -> term list * term list -> bool |
33192 | 15 |
end; |
16 |
||
33232
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33192
diff
changeset
|
17 |
structure Nitpick_Mono : NITPICK_MONO = |
33192 | 18 |
struct |
19 |
||
33232
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33192
diff
changeset
|
20 |
open Nitpick_Util |
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
blanchet
parents:
33192
diff
changeset
|
21 |
open Nitpick_HOL |
33192 | 22 |
|
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
23 |
datatype sign = Plus | Minus |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
24 |
|
33192 | 25 |
type var = int |
26 |
||
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
27 |
datatype annotation = Gen | New | Fls | Tru |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
28 |
datatype annotation_atom = A of annotation | V of var |
33192 | 29 |
|
40996
63112be4a469
added "Neq" operator to monotonicity inference module
blanchet
parents:
40995
diff
changeset
|
30 |
type assign_literal = var * (sign * annotation) |
33192 | 31 |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
32 |
datatype mtyp = |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
33 |
MAlpha | |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
34 |
MFun of mtyp * annotation_atom * mtyp | |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
35 |
MPair of mtyp * mtyp | |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
36 |
MType of string * mtyp list | |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
37 |
MRec of string * typ list |
33192 | 38 |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
39 |
type mdata = |
35070
96136eb6218f
split "nitpick_hol.ML" into two files to make it more manageable;
blanchet
parents:
34998
diff
changeset
|
40 |
{hol_ctxt: hol_context, |
35190
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents:
35079
diff
changeset
|
41 |
binarize: bool, |
33192 | 42 |
alpha_T: typ, |
43 |
max_fresh: int Unsynchronized.ref, |
|
55890 | 44 |
data_type_mcache: ((string * typ list) * mtyp) list Unsynchronized.ref, |
55889 | 45 |
constr_mcache: ((string * typ) * mtyp) list Unsynchronized.ref} |
33192 | 46 |
|
35812
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
47 |
exception UNSOLVABLE of unit |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
48 |
exception MTYPE of string * mtyp list * typ list |
33192 | 49 |
|
38179 | 50 |
val trace = Unsynchronized.ref false |
55889 | 51 |
|
38179 | 52 |
fun trace_msg msg = if !trace then tracing (msg ()) else () |
33192 | 53 |
|
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
54 |
fun string_for_sign Plus = "+" |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
55 |
| string_for_sign Minus = "-" |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
56 |
|
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
57 |
fun negate_sign Plus = Minus |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
58 |
| negate_sign Minus = Plus |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
59 |
|
33192 | 60 |
val string_for_var = signed_string_of_int |
55889 | 61 |
|
33192 | 62 |
fun string_for_vars sep [] = "0\<^bsub>" ^ sep ^ "\<^esub>" |
63 |
| string_for_vars sep xs = space_implode sep (map string_for_var xs) |
|
55889 | 64 |
|
33192 | 65 |
fun subscript_string_for_vars sep xs = |
66 |
if null xs then "" else "\<^bsub>" ^ string_for_vars sep xs ^ "\<^esub>" |
|
67 |
||
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
68 |
fun string_for_annotation Gen = "G" |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
69 |
| string_for_annotation New = "N" |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
70 |
| string_for_annotation Fls = "F" |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
71 |
| string_for_annotation Tru = "T" |
33192 | 72 |
|
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
73 |
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
|
74 |
| string_for_annotation_atom (V x) = string_for_var x |
33192 | 75 |
|
40996
63112be4a469
added "Neq" operator to monotonicity inference module
blanchet
parents:
40995
diff
changeset
|
76 |
fun string_for_assign_literal (x, (sn, a)) = |
63112be4a469
added "Neq" operator to monotonicity inference module
blanchet
parents:
40995
diff
changeset
|
77 |
string_for_var x ^ (case sn of Plus => " = " | Minus => " \<noteq> ") ^ |
63112be4a469
added "Neq" operator to monotonicity inference module
blanchet
parents:
40995
diff
changeset
|
78 |
string_for_annotation a |
33192 | 79 |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
80 |
val bool_M = MType (@{type_name bool}, []) |
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
81 |
val dummy_M = MType (nitpick_prefix ^ "dummy", []) |
33192 | 82 |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
83 |
fun is_MRec (MRec _) = true |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
84 |
| is_MRec _ = false |
55889 | 85 |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
86 |
fun dest_MFun (MFun z) = z |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
87 |
| dest_MFun M = raise MTYPE ("Nitpick_Mono.dest_MFun", [M], []) |
33192 | 88 |
|
89 |
val no_prec = 100 |
|
90 |
||
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
91 |
fun precedence_of_mtype (MFun _) = 1 |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
92 |
| precedence_of_mtype (MPair _) = 2 |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
93 |
| precedence_of_mtype _ = no_prec |
33192 | 94 |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
95 |
val string_for_mtype = |
33192 | 96 |
let |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
97 |
fun aux outer_prec M = |
33192 | 98 |
let |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
99 |
val prec = precedence_of_mtype M |
33192 | 100 |
val need_parens = (prec < outer_prec) |
101 |
in |
|
102 |
(if need_parens then "(" else "") ^ |
|
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
103 |
(if M = dummy_M then |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
104 |
"_" |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
105 |
else case M of |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
106 |
MAlpha => "\<alpha>" |
40988 | 107 |
| MFun (M1, aa, M2) => |
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
108 |
aux (prec + 1) M1 ^ " \<Rightarrow>\<^bsup>" ^ |
40988 | 109 |
string_for_annotation_atom aa ^ "\<^esup> " ^ aux prec M2 |
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
110 |
| 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
|
111 |
| MType (s, []) => |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
112 |
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
|
113 |
else s |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
114 |
| MType (s, Ms) => "(" ^ commas (map (aux 0) Ms) ^ ") " ^ s |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
115 |
| MRec (s, _) => "[" ^ s ^ "]") ^ |
33192 | 116 |
(if need_parens then ")" else "") |
117 |
end |
|
118 |
in aux 0 end |
|
119 |
||
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
120 |
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
|
121 |
| flatten_mtype (MType (_, Ms)) = maps flatten_mtype Ms |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
122 |
| flatten_mtype M = [M] |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
123 |
|
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
124 |
fun initial_mdata hol_ctxt binarize alpha_T = |
35190
ce653cc27a94
make sure that Nitpick uses binary notation consistently if "binary_ints" is enabled
blanchet
parents:
35079
diff
changeset
|
125 |
({hol_ctxt = hol_ctxt, binarize = binarize, alpha_T = alpha_T, |
55890 | 126 |
max_fresh = Unsynchronized.ref 0, data_type_mcache = Unsynchronized.ref [], |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
127 |
constr_mcache = Unsynchronized.ref []} : mdata) |
33192 | 128 |
|
129 |
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
|
130 |
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
|
131 |
exists (could_exist_alpha_subtype alpha_T) Ts) |
33192 | 132 |
| could_exist_alpha_subtype alpha_T T = (T = alpha_T) |
55889 | 133 |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
134 |
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
|
135 |
could_exist_alpha_subtype alpha_T T |
37256
0dca1ec52999
thread along context instead of theory for typedef lookup
blanchet
parents:
36385
diff
changeset
|
136 |
| could_exist_alpha_sub_mtype ctxt alpha_T T = |
55890 | 137 |
(T = alpha_T orelse is_data_type ctxt T) |
33192 | 138 |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
139 |
fun exists_alpha_sub_mtype MAlpha = true |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
140 |
| exists_alpha_sub_mtype (MFun (M1, _, M2)) = |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
141 |
exists exists_alpha_sub_mtype [M1, M2] |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
142 |
| exists_alpha_sub_mtype (MPair (M1, M2)) = |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
143 |
exists exists_alpha_sub_mtype [M1, M2] |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
144 |
| 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
|
145 |
| exists_alpha_sub_mtype (MRec _) = true |
33192 | 146 |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
147 |
fun exists_alpha_sub_mtype_fresh MAlpha = true |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
148 |
| exists_alpha_sub_mtype_fresh (MFun (_, V _, _)) = true |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
149 |
| exists_alpha_sub_mtype_fresh (MFun (_, _, M2)) = |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
150 |
exists_alpha_sub_mtype_fresh M2 |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
151 |
| exists_alpha_sub_mtype_fresh (MPair (M1, M2)) = |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
152 |
exists exists_alpha_sub_mtype_fresh [M1, M2] |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
153 |
| exists_alpha_sub_mtype_fresh (MType (_, Ms)) = |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
154 |
exists exists_alpha_sub_mtype_fresh Ms |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
155 |
| exists_alpha_sub_mtype_fresh (MRec _) = true |
33192 | 156 |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
157 |
fun constr_mtype_for_binders z Ms = |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
158 |
fold_rev (fn M => curry3 MFun M (A Gen)) Ms (MRec z) |
33192 | 159 |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
160 |
fun repair_mtype _ _ MAlpha = MAlpha |
40988 | 161 |
| repair_mtype cache seen (MFun (M1, aa, M2)) = |
162 |
MFun (repair_mtype cache seen M1, aa, repair_mtype cache seen M2) |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
163 |
| repair_mtype cache seen (MPair Mp) = |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
164 |
MPair (pairself (repair_mtype cache seen) Mp) |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
165 |
| repair_mtype cache seen (MType (s, Ms)) = |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
166 |
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
|
167 |
| repair_mtype cache seen (MRec (z as (s, _))) = |
33192 | 168 |
case AList.lookup (op =) cache z |> the of |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
169 |
MRec _ => MType (s, []) |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
170 |
| M => if member (op =) seen M then MType (s, []) |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
171 |
else repair_mtype cache (M :: seen) M |
33192 | 172 |
|
55890 | 173 |
fun repair_data_type_mcache cache = |
33192 | 174 |
let |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
175 |
fun repair_one (z, M) = |
33192 | 176 |
Unsynchronized.change cache |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
177 |
(AList.update (op =) (z, repair_mtype (!cache) [] M)) |
33192 | 178 |
in List.app repair_one (rev (!cache)) end |
179 |
||
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
180 |
fun repair_constr_mcache dtype_cache constr_mcache = |
33192 | 181 |
let |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
182 |
fun repair_one (x, M) = |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
183 |
Unsynchronized.change constr_mcache |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
184 |
(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
|
185 |
in List.app repair_one (!constr_mcache) end |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
186 |
|
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
187 |
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
|
188 |
| is_fin_fun_supported_type @{typ bool} = true |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
189 |
| 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
|
190 |
| is_fin_fun_supported_type _ = false |
41009
91495051c2ec
improve precision of finite functions in monotonicity calculus
blanchet
parents:
41007
diff
changeset
|
191 |
|
91495051c2ec
improve precision of finite functions in monotonicity calculus
blanchet
parents:
41007
diff
changeset
|
192 |
(* TODO: clean this up *) |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
193 |
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
|
194 |
| 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
|
195 |
| fin_fun_body dom_T ran_T |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
196 |
((t0 as Const (@{const_name If}, _)) |
38864
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
haftmann
parents:
38795
diff
changeset
|
197 |
$ (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
|
198 |
$ t2 $ t3) = |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
199 |
(if loose_bvar1 (t1', 0) then |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
200 |
NONE |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
201 |
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
|
202 |
NONE => NONE |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
203 |
| SOME t3 => |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
204 |
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
|
205 |
$ (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
|
206 |
| fin_fun_body _ _ _ = NONE |
33192 | 207 |
|
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
208 |
(* FIXME: make sure well-annotated *) |
40998
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
209 |
|
35672 | 210 |
fun fresh_mfun_for_fun_type (mdata as {max_fresh, ...} : mdata) all_minus |
211 |
T1 T2 = |
|
33192 | 212 |
let |
35672 | 213 |
val M1 = fresh_mtype_for_type mdata all_minus T1 |
214 |
val M2 = fresh_mtype_for_type mdata all_minus T2 |
|
40988 | 215 |
val aa = if not all_minus andalso exists_alpha_sub_mtype_fresh M1 andalso |
216 |
is_fin_fun_supported_type (body_type T2) then |
|
217 |
V (Unsynchronized.inc max_fresh) |
|
218 |
else |
|
219 |
A Gen |
|
220 |
in (M1, aa, M2) end |
|
37256
0dca1ec52999
thread along context instead of theory for typedef lookup
blanchet
parents:
36385
diff
changeset
|
221 |
and fresh_mtype_for_type (mdata as {hol_ctxt as {ctxt, ...}, binarize, alpha_T, |
55890 | 222 |
data_type_mcache, constr_mcache, ...}) |
35672 | 223 |
all_minus = |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
224 |
let |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
225 |
fun do_type T = |
33192 | 226 |
if T = alpha_T then |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
227 |
MAlpha |
33192 | 228 |
else case T of |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
229 |
Type (@{type_name fun}, [T1, T2]) => |
39342 | 230 |
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
|
231 |
| Type (@{type_name prod}, [T1, T2]) => MPair (pairself do_type (T1, T2)) |
46081
8f6465f7021b
ported mono calculus to handle "set" type constructors
blanchet
parents:
41471
diff
changeset
|
232 |
| Type (@{type_name set}, [T']) => do_type (T' --> bool_T) |
33192 | 233 |
| Type (z as (s, _)) => |
37256
0dca1ec52999
thread along context instead of theory for typedef lookup
blanchet
parents:
36385
diff
changeset
|
234 |
if could_exist_alpha_sub_mtype ctxt alpha_T T then |
55890 | 235 |
case AList.lookup (op =) (!data_type_mcache) z of |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
236 |
SOME M => M |
33192 | 237 |
| NONE => |
238 |
let |
|
55890 | 239 |
val _ = Unsynchronized.change data_type_mcache (cons (z, MRec z)) |
240 |
val xs = binarized_and_boxed_data_type_constrs hol_ctxt binarize T |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
241 |
val (all_Ms, constr_Ms) = |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
242 |
fold_rev (fn (_, T') => fn (all_Ms, constr_Ms) => |
33192 | 243 |
let |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
244 |
val binder_Ms = map do_type (binder_types T') |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
245 |
val new_Ms = filter exists_alpha_sub_mtype_fresh |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
246 |
binder_Ms |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
247 |
val constr_M = constr_mtype_for_binders z |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
248 |
binder_Ms |
33192 | 249 |
in |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
250 |
(union (op =) new_Ms all_Ms, |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
251 |
constr_M :: constr_Ms) |
33192 | 252 |
end) |
253 |
xs ([], []) |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
254 |
val M = MType (s, all_Ms) |
55890 | 255 |
val _ = Unsynchronized.change data_type_mcache |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
256 |
(AList.update (op =) (z, M)) |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
257 |
val _ = Unsynchronized.change constr_mcache |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
258 |
(append (xs ~~ constr_Ms)) |
33192 | 259 |
in |
55890 | 260 |
if forall (not o is_MRec o snd) (!data_type_mcache) then |
261 |
(repair_data_type_mcache data_type_mcache; |
|
262 |
repair_constr_mcache (!data_type_mcache) constr_mcache; |
|
263 |
AList.lookup (op =) (!data_type_mcache) z |> the) |
|
33192 | 264 |
else |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
265 |
M |
33192 | 266 |
end |
267 |
else |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
268 |
MType (s, []) |
37260
dde817e6dfb1
added "atoms" option to Nitpick (request from Karlsruhe) + wrap Refute. functions to "nitpick_util.ML"
blanchet
parents:
37256
diff
changeset
|
269 |
| _ => MType (simple_string_of_typ T, []) |
33192 | 270 |
in do_type end |
271 |
||
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
272 |
val ground_and_sole_base_constrs = [] |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
273 |
(* FIXME: [@{const_name Nil}, @{const_name None}], cf. lists_empty *) |
41009
91495051c2ec
improve precision of finite functions in monotonicity calculus
blanchet
parents:
41007
diff
changeset
|
274 |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
275 |
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
|
276 |
| prodM_factors M = [M] |
55889 | 277 |
|
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
278 |
fun curried_strip_mtype (MFun (M1, _, M2)) = |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
279 |
curried_strip_mtype M2 |>> append (prodM_factors M1) |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
280 |
| curried_strip_mtype M = ([], M) |
55889 | 281 |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
282 |
fun sel_mtype_from_constr_mtype s M = |
41009
91495051c2ec
improve precision of finite functions in monotonicity calculus
blanchet
parents:
41007
diff
changeset
|
283 |
let |
91495051c2ec
improve precision of finite functions in monotonicity calculus
blanchet
parents:
41007
diff
changeset
|
284 |
val (arg_Ms, dataM) = curried_strip_mtype M |
91495051c2ec
improve precision of finite functions in monotonicity calculus
blanchet
parents:
41007
diff
changeset
|
285 |
val a = if member (op =) ground_and_sole_base_constrs |
91495051c2ec
improve precision of finite functions in monotonicity calculus
blanchet
parents:
41007
diff
changeset
|
286 |
(constr_name_for_sel_like s) then |
91495051c2ec
improve precision of finite functions in monotonicity calculus
blanchet
parents:
41007
diff
changeset
|
287 |
Fls |
91495051c2ec
improve precision of finite functions in monotonicity calculus
blanchet
parents:
41007
diff
changeset
|
288 |
else |
91495051c2ec
improve precision of finite functions in monotonicity calculus
blanchet
parents:
41007
diff
changeset
|
289 |
Gen |
91495051c2ec
improve precision of finite functions in monotonicity calculus
blanchet
parents:
41007
diff
changeset
|
290 |
in |
91495051c2ec
improve precision of finite functions in monotonicity calculus
blanchet
parents:
41007
diff
changeset
|
291 |
MFun (dataM, A a, |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
292 |
case sel_no_from_name s of ~1 => bool_M | n => nth arg_Ms n) |
33192 | 293 |
end |
294 |
||
37256
0dca1ec52999
thread along context instead of theory for typedef lookup
blanchet
parents:
36385
diff
changeset
|
295 |
fun mtype_for_constr (mdata as {hol_ctxt = {ctxt, ...}, alpha_T, constr_mcache, |
33192 | 296 |
...}) (x as (_, T)) = |
37256
0dca1ec52999
thread along context instead of theory for typedef lookup
blanchet
parents:
36385
diff
changeset
|
297 |
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
|
298 |
case AList.lookup (op =) (!constr_mcache) x of |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
299 |
SOME M => M |
35384 | 300 |
| NONE => if T = alpha_T then |
35672 | 301 |
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
|
302 |
(Unsynchronized.change constr_mcache (cons (x, M)); M) |
35384 | 303 |
end |
304 |
else |
|
35672 | 305 |
(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
|
306 |
AList.lookup (op =) (!constr_mcache) x |> the) |
33192 | 307 |
else |
35672 | 308 |
fresh_mtype_for_type mdata false T |
55889 | 309 |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
310 |
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
|
311 |
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
|
312 |
|> mtype_for_constr mdata |> sel_mtype_from_constr_mtype s |
33192 | 313 |
|
40989 | 314 |
fun resolve_annotation_atom asgs (V x) = |
315 |
x |> AList.lookup (op =) asgs |> Option.map A |> the_default (V x) |
|
40988 | 316 |
| resolve_annotation_atom _ aa = aa |
55889 | 317 |
|
40989 | 318 |
fun resolve_mtype asgs = |
33192 | 319 |
let |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
320 |
fun aux MAlpha = MAlpha |
40988 | 321 |
| aux (MFun (M1, aa, M2)) = |
40989 | 322 |
MFun (aux M1, resolve_annotation_atom asgs aa, aux M2) |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
323 |
| aux (MPair Mp) = MPair (pairself aux Mp) |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
324 |
| aux (MType (s, Ms)) = MType (s, map aux Ms) |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
325 |
| aux (MRec z) = MRec z |
33192 | 326 |
in aux end |
327 |
||
40996
63112be4a469
added "Neq" operator to monotonicity inference module
blanchet
parents:
40995
diff
changeset
|
328 |
datatype comp_op = Eq | Neq | Leq |
33192 | 329 |
|
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
330 |
type comp = annotation_atom * annotation_atom * comp_op * var list |
40995
3cae30b60577
started implementing connectives in new monotonicity calculus
blanchet
parents:
40994
diff
changeset
|
331 |
type assign_clause = assign_literal list |
33192 | 332 |
|
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
333 |
type constraint_set = comp list * assign_clause list |
33192 | 334 |
|
335 |
fun string_for_comp_op Eq = "=" |
|
40996
63112be4a469
added "Neq" operator to monotonicity inference module
blanchet
parents:
40995
diff
changeset
|
336 |
| string_for_comp_op Neq = "\<noteq>" |
33192 | 337 |
| string_for_comp_op Leq = "\<le>" |
338 |
||
40990
a36d4d869439
adapt monotonicity code to four annotation types
blanchet
parents:
40989
diff
changeset
|
339 |
fun string_for_comp (aa1, aa2, cmp, xs) = |
a36d4d869439
adapt monotonicity code to four annotation types
blanchet
parents:
40989
diff
changeset
|
340 |
string_for_annotation_atom aa1 ^ " " ^ string_for_comp_op cmp ^ |
a36d4d869439
adapt monotonicity code to four annotation types
blanchet
parents:
40989
diff
changeset
|
341 |
subscript_string_for_vars " \<and> " xs ^ " " ^ string_for_annotation_atom aa2 |
33192 | 342 |
|
40999
69d0d445c46a
fixed bug in clause handling in monotonicity code, whereby the unsound rule False | x <--> False was used to simplify constraints
blanchet
parents:
40998
diff
changeset
|
343 |
fun string_for_assign_clause NONE = "\<top>" |
69d0d445c46a
fixed bug in clause handling in monotonicity code, whereby the unsound rule False | x <--> False was used to simplify constraints
blanchet
parents:
40998
diff
changeset
|
344 |
| string_for_assign_clause (SOME []) = "\<bot>" |
69d0d445c46a
fixed bug in clause handling in monotonicity code, whereby the unsound rule False | x <--> False was used to simplify constraints
blanchet
parents:
40998
diff
changeset
|
345 |
| string_for_assign_clause (SOME asgs) = |
40995
3cae30b60577
started implementing connectives in new monotonicity calculus
blanchet
parents:
40994
diff
changeset
|
346 |
space_implode " \<or> " (map string_for_assign_literal asgs) |
40990
a36d4d869439
adapt monotonicity code to four annotation types
blanchet
parents:
40989
diff
changeset
|
347 |
|
40996
63112be4a469
added "Neq" operator to monotonicity inference module
blanchet
parents:
40995
diff
changeset
|
348 |
fun add_assign_literal (x, (sn, a)) clauses = |
63112be4a469
added "Neq" operator to monotonicity inference module
blanchet
parents:
40995
diff
changeset
|
349 |
if exists (fn [(x', (sn', a'))] => |
63112be4a469
added "Neq" operator to monotonicity inference module
blanchet
parents:
40995
diff
changeset
|
350 |
x = x' andalso ((sn = sn' andalso a <> a') orelse |
63112be4a469
added "Neq" operator to monotonicity inference module
blanchet
parents:
40995
diff
changeset
|
351 |
(sn <> sn' andalso a = a')) |
63112be4a469
added "Neq" operator to monotonicity inference module
blanchet
parents:
40995
diff
changeset
|
352 |
| _ => false) clauses then |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
353 |
NONE |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
354 |
else |
41017 | 355 |
SOME ([(x, (sn, a))] :: clauses) |
33192 | 356 |
|
40991
902ad76994d5
proper handling of assignment disjunctions vs. conjunctions
blanchet
parents:
40990
diff
changeset
|
357 |
fun add_assign_disjunct _ NONE = NONE |
902ad76994d5
proper handling of assignment disjunctions vs. conjunctions
blanchet
parents:
40990
diff
changeset
|
358 |
| add_assign_disjunct asg (SOME asgs) = SOME (insert (op =) asg asgs) |
902ad76994d5
proper handling of assignment disjunctions vs. conjunctions
blanchet
parents:
40990
diff
changeset
|
359 |
|
40999
69d0d445c46a
fixed bug in clause handling in monotonicity code, whereby the unsound rule False | x <--> False was used to simplify constraints
blanchet
parents:
40998
diff
changeset
|
360 |
fun add_assign_clause NONE = I |
69d0d445c46a
fixed bug in clause handling in monotonicity code, whereby the unsound rule False | x <--> False was used to simplify constraints
blanchet
parents:
40998
diff
changeset
|
361 |
| add_assign_clause (SOME clause) = insert (op =) clause |
40997
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
362 |
|
40996
63112be4a469
added "Neq" operator to monotonicity inference module
blanchet
parents:
40995
diff
changeset
|
363 |
fun annotation_comp Eq a1 a2 = (a1 = a2) |
63112be4a469
added "Neq" operator to monotonicity inference module
blanchet
parents:
40995
diff
changeset
|
364 |
| annotation_comp Neq a1 a2 = (a1 <> a2) |
63112be4a469
added "Neq" operator to monotonicity inference module
blanchet
parents:
40995
diff
changeset
|
365 |
| annotation_comp Leq a1 a2 = (a1 = a2 orelse a2 = Gen) |
63112be4a469
added "Neq" operator to monotonicity inference module
blanchet
parents:
40995
diff
changeset
|
366 |
|
40997
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
367 |
fun sign_for_comp_op Eq = Plus |
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
368 |
| sign_for_comp_op Neq = Minus |
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
369 |
| sign_for_comp_op Leq = raise BAD ("sign_for_comp_op", "unexpected \"Leq\"") |
40996
63112be4a469
added "Neq" operator to monotonicity inference module
blanchet
parents:
40995
diff
changeset
|
370 |
|
63112be4a469
added "Neq" operator to monotonicity inference module
blanchet
parents:
40995
diff
changeset
|
371 |
fun do_annotation_atom_comp Leq [] aa1 aa2 (cset as (comps, clauses)) = |
40988 | 372 |
(case (aa1, aa2) of |
40996
63112be4a469
added "Neq" operator to monotonicity inference module
blanchet
parents:
40995
diff
changeset
|
373 |
(A a1, A a2) => if annotation_comp Leq a1 a2 then SOME cset else NONE |
63112be4a469
added "Neq" operator to monotonicity inference module
blanchet
parents:
40995
diff
changeset
|
374 |
| _ => SOME (insert (op =) (aa1, aa2, Leq, []) comps, clauses)) |
63112be4a469
added "Neq" operator to monotonicity inference module
blanchet
parents:
40995
diff
changeset
|
375 |
| do_annotation_atom_comp cmp [] aa1 aa2 (cset as (comps, clauses)) = |
63112be4a469
added "Neq" operator to monotonicity inference module
blanchet
parents:
40995
diff
changeset
|
376 |
(case (aa1, aa2) of |
63112be4a469
added "Neq" operator to monotonicity inference module
blanchet
parents:
40995
diff
changeset
|
377 |
(A a1, A a2) => if annotation_comp cmp a1 a2 then SOME cset else NONE |
40988 | 378 |
| (V x1, A a2) => |
40997
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
379 |
clauses |> add_assign_literal (x1, (sign_for_comp_op cmp, a2)) |
40996
63112be4a469
added "Neq" operator to monotonicity inference module
blanchet
parents:
40995
diff
changeset
|
380 |
|> Option.map (pair comps) |
63112be4a469
added "Neq" operator to monotonicity inference module
blanchet
parents:
40995
diff
changeset
|
381 |
| (A _, V _) => do_annotation_atom_comp cmp [] aa2 aa1 cset |
63112be4a469
added "Neq" operator to monotonicity inference module
blanchet
parents:
40995
diff
changeset
|
382 |
| (V _, V _) => SOME (insert (op =) (aa1, aa2, cmp, []) comps, clauses)) |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
383 |
| do_annotation_atom_comp cmp xs aa1 aa2 (comps, clauses) = |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
384 |
SOME (insert (op =) (aa1, aa2, cmp, xs) comps, clauses) |
33192 | 385 |
|
40997
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
386 |
fun add_annotation_atom_comp cmp xs aa1 aa2 (comps, clauses) = |
40998
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
387 |
(trace_msg (fn () => "*** Add " ^ string_for_comp (aa1, aa2, cmp, xs)); |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
388 |
case do_annotation_atom_comp cmp xs aa1 aa2 (comps, clauses) of |
40993
52ee2a187cdb
support 3 monotonicity calculi in one and fix soundness bug
blanchet
parents:
40991
diff
changeset
|
389 |
NONE => (trace_msg (K "**** Unsolvable"); raise UNSOLVABLE ()) |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
390 |
| SOME cset => cset) |
40993
52ee2a187cdb
support 3 monotonicity calculi in one and fix soundness bug
blanchet
parents:
40991
diff
changeset
|
391 |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
392 |
fun do_mtype_comp _ _ _ _ NONE = NONE |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
393 |
| do_mtype_comp _ _ MAlpha MAlpha cset = cset |
40988 | 394 |
| do_mtype_comp Eq xs (MFun (M11, aa1, M12)) (MFun (M21, aa2, M22)) |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
395 |
(SOME cset) = |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
396 |
cset |> do_annotation_atom_comp Eq xs aa1 aa2 |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
397 |
|> do_mtype_comp Eq xs M11 M21 |> do_mtype_comp Eq xs M12 M22 |
40988 | 398 |
| do_mtype_comp Leq xs (MFun (M11, aa1, M12)) (MFun (M21, aa2, M22)) |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
399 |
(SOME cset) = |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
400 |
(if exists_alpha_sub_mtype M11 then |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
401 |
cset |> do_annotation_atom_comp Leq xs aa1 aa2 |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
402 |
|> do_mtype_comp Leq xs M21 M11 |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
403 |
|> (case aa2 of |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
404 |
A Gen => I |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
405 |
| A _ => do_mtype_comp Leq xs M11 M21 |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
406 |
| V x => do_mtype_comp Leq (x :: xs) M11 M21) |
33192 | 407 |
else |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
408 |
SOME cset) |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
409 |
|> do_mtype_comp Leq xs M12 M22 |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
410 |
| do_mtype_comp cmp xs (M1 as MPair (M11, M12)) (M2 as MPair (M21, M22)) |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
411 |
cset = |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
412 |
(cset |> fold (uncurry (do_mtype_comp cmp xs)) [(M11, M21), (M12, M22)] |
41017 | 413 |
handle ListPair.UnequalLengths => |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
414 |
raise MTYPE ("Nitpick_Mono.do_mtype_comp", [M1, M2], [])) |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
415 |
| do_mtype_comp _ _ (MType _) (MType _) cset = |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
416 |
cset (* 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
|
417 |
| do_mtype_comp cmp _ M1 M2 _ = |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
418 |
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
|
419 |
[M1, M2], []) |
33192 | 420 |
|
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
421 |
fun add_mtype_comp cmp M1 M2 cset = |
40991
902ad76994d5
proper handling of assignment disjunctions vs. conjunctions
blanchet
parents:
40990
diff
changeset
|
422 |
(trace_msg (fn () => "*** Add " ^ string_for_mtype M1 ^ " " ^ |
902ad76994d5
proper handling of assignment disjunctions vs. conjunctions
blanchet
parents:
40990
diff
changeset
|
423 |
string_for_comp_op cmp ^ " " ^ string_for_mtype M2); |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
424 |
case SOME cset |> do_mtype_comp cmp [] M1 M2 of |
40991
902ad76994d5
proper handling of assignment disjunctions vs. conjunctions
blanchet
parents:
40990
diff
changeset
|
425 |
NONE => (trace_msg (K "**** Unsolvable"); raise UNSOLVABLE ()) |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
426 |
| SOME cset => cset) |
33192 | 427 |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
428 |
val add_mtypes_equal = add_mtype_comp Eq |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
429 |
val add_is_sub_mtype = add_mtype_comp Leq |
33192 | 430 |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
431 |
fun do_notin_mtype_fv _ _ _ NONE = NONE |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
432 |
| do_notin_mtype_fv Minus _ MAlpha cset = cset |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
433 |
| do_notin_mtype_fv Plus [] MAlpha _ = NONE |
40996
63112be4a469
added "Neq" operator to monotonicity inference module
blanchet
parents:
40995
diff
changeset
|
434 |
| do_notin_mtype_fv Plus [asg] MAlpha (SOME clauses) = |
63112be4a469
added "Neq" operator to monotonicity inference module
blanchet
parents:
40995
diff
changeset
|
435 |
clauses |> add_assign_literal asg |
41000
4bbff1684465
implemented All rules from new monotonicity calculus
blanchet
parents:
40999
diff
changeset
|
436 |
| do_notin_mtype_fv Plus unless MAlpha (SOME clauses) = |
4bbff1684465
implemented All rules from new monotonicity calculus
blanchet
parents:
40999
diff
changeset
|
437 |
SOME (insert (op =) unless clauses) |
4bbff1684465
implemented All rules from new monotonicity calculus
blanchet
parents:
40999
diff
changeset
|
438 |
| do_notin_mtype_fv sn unless (MFun (M1, A a, M2)) cset = |
4bbff1684465
implemented All rules from new monotonicity calculus
blanchet
parents:
40999
diff
changeset
|
439 |
cset |> (if a <> Gen andalso sn = Plus then do_notin_mtype_fv Plus unless M1 |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
440 |
else I) |
41000
4bbff1684465
implemented All rules from new monotonicity calculus
blanchet
parents:
40999
diff
changeset
|
441 |
|> (if a = Gen orelse sn = Plus then do_notin_mtype_fv Minus unless M1 |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
442 |
else I) |
41000
4bbff1684465
implemented All rules from new monotonicity calculus
blanchet
parents:
40999
diff
changeset
|
443 |
|> do_notin_mtype_fv sn unless M2 |
4bbff1684465
implemented All rules from new monotonicity calculus
blanchet
parents:
40999
diff
changeset
|
444 |
| do_notin_mtype_fv Plus unless (MFun (M1, V x, M2)) cset = |
4bbff1684465
implemented All rules from new monotonicity calculus
blanchet
parents:
40999
diff
changeset
|
445 |
cset |> (case add_assign_disjunct (x, (Plus, Gen)) (SOME unless) of |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
446 |
NONE => I |
41000
4bbff1684465
implemented All rules from new monotonicity calculus
blanchet
parents:
40999
diff
changeset
|
447 |
| SOME unless' => do_notin_mtype_fv Plus unless' M1) |
4bbff1684465
implemented All rules from new monotonicity calculus
blanchet
parents:
40999
diff
changeset
|
448 |
|> do_notin_mtype_fv Minus unless M1 |
4bbff1684465
implemented All rules from new monotonicity calculus
blanchet
parents:
40999
diff
changeset
|
449 |
|> do_notin_mtype_fv Plus unless M2 |
4bbff1684465
implemented All rules from new monotonicity calculus
blanchet
parents:
40999
diff
changeset
|
450 |
| do_notin_mtype_fv Minus unless (MFun (M1, V x, M2)) cset = |
40996
63112be4a469
added "Neq" operator to monotonicity inference module
blanchet
parents:
40995
diff
changeset
|
451 |
cset |> (case fold (fn a => add_assign_disjunct (x, (Plus, a))) [Fls, Tru] |
41000
4bbff1684465
implemented All rules from new monotonicity calculus
blanchet
parents:
40999
diff
changeset
|
452 |
(SOME unless) of |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
453 |
NONE => I |
41000
4bbff1684465
implemented All rules from new monotonicity calculus
blanchet
parents:
40999
diff
changeset
|
454 |
| SOME unless' => do_notin_mtype_fv Plus unless' M1) |
4bbff1684465
implemented All rules from new monotonicity calculus
blanchet
parents:
40999
diff
changeset
|
455 |
|> do_notin_mtype_fv Minus unless M2 |
4bbff1684465
implemented All rules from new monotonicity calculus
blanchet
parents:
40999
diff
changeset
|
456 |
| do_notin_mtype_fv sn unless (MPair (M1, M2)) cset = |
4bbff1684465
implemented All rules from new monotonicity calculus
blanchet
parents:
40999
diff
changeset
|
457 |
cset |> fold (do_notin_mtype_fv sn unless) [M1, M2] |
4bbff1684465
implemented All rules from new monotonicity calculus
blanchet
parents:
40999
diff
changeset
|
458 |
| do_notin_mtype_fv sn unless (MType (_, Ms)) cset = |
4bbff1684465
implemented All rules from new monotonicity calculus
blanchet
parents:
40999
diff
changeset
|
459 |
cset |> fold (do_notin_mtype_fv sn unless) Ms |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
460 |
| do_notin_mtype_fv _ _ M _ = |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
461 |
raise MTYPE ("Nitpick_Mono.do_notin_mtype_fv", [M], []) |
33192 | 462 |
|
41000
4bbff1684465
implemented All rules from new monotonicity calculus
blanchet
parents:
40999
diff
changeset
|
463 |
fun add_notin_mtype_fv sn unless M (comps, clauses) = |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
464 |
(trace_msg (fn () => "*** Add " ^ string_for_mtype M ^ " is " ^ |
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
465 |
(case sn of Minus => "concrete" | Plus => "complete")); |
41000
4bbff1684465
implemented All rules from new monotonicity calculus
blanchet
parents:
40999
diff
changeset
|
466 |
case SOME clauses |> do_notin_mtype_fv sn unless M of |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
467 |
NONE => (trace_msg (K "**** Unsolvable"); raise UNSOLVABLE ()) |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
468 |
| SOME clauses => (comps, clauses)) |
33192 | 469 |
|
41109 | 470 |
fun add_mtype_is_concrete x y z = add_notin_mtype_fv Minus x y z |
471 |
fun add_mtype_is_complete x y z = add_notin_mtype_fv Plus x y z |
|
33192 | 472 |
|
40986
cfd91aa80701
use boolean pair to encode annotation, which may now take four values
blanchet
parents:
40985
diff
changeset
|
473 |
val bool_table = |
cfd91aa80701
use boolean pair to encode annotation, which may now take four values
blanchet
parents:
40985
diff
changeset
|
474 |
[(Gen, (false, false)), |
cfd91aa80701
use boolean pair to encode annotation, which may now take four values
blanchet
parents:
40985
diff
changeset
|
475 |
(New, (false, true)), |
cfd91aa80701
use boolean pair to encode annotation, which may now take four values
blanchet
parents:
40985
diff
changeset
|
476 |
(Fls, (true, false)), |
cfd91aa80701
use boolean pair to encode annotation, which may now take four values
blanchet
parents:
40985
diff
changeset
|
477 |
(Tru, (true, true))] |
cfd91aa80701
use boolean pair to encode annotation, which may now take four values
blanchet
parents:
40985
diff
changeset
|
478 |
|
40993
52ee2a187cdb
support 3 monotonicity calculi in one and fix soundness bug
blanchet
parents:
40991
diff
changeset
|
479 |
fun fst_var n = 2 * n |
52ee2a187cdb
support 3 monotonicity calculi in one and fix soundness bug
blanchet
parents:
40991
diff
changeset
|
480 |
fun snd_var n = 2 * n + 1 |
52ee2a187cdb
support 3 monotonicity calculi in one and fix soundness bug
blanchet
parents:
40991
diff
changeset
|
481 |
|
40986
cfd91aa80701
use boolean pair to encode annotation, which may now take four values
blanchet
parents:
40985
diff
changeset
|
482 |
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
|
483 |
val annotation_from_bools = AList.find (op =) bool_table #> the_single |
33192 | 484 |
|
41471 | 485 |
fun prop_for_bool b = if b then Prop_Logic.True else Prop_Logic.False |
55889 | 486 |
|
40990
a36d4d869439
adapt monotonicity code to four annotation types
blanchet
parents:
40989
diff
changeset
|
487 |
fun prop_for_bool_var_equality (v1, v2) = |
46081
8f6465f7021b
ported mono calculus to handle "set" type constructors
blanchet
parents:
41471
diff
changeset
|
488 |
Prop_Logic.SAnd (Prop_Logic.SOr (Prop_Logic.BoolVar v1, |
8f6465f7021b
ported mono calculus to handle "set" type constructors
blanchet
parents:
41471
diff
changeset
|
489 |
Prop_Logic.SNot (Prop_Logic.BoolVar v2)), |
8f6465f7021b
ported mono calculus to handle "set" type constructors
blanchet
parents:
41471
diff
changeset
|
490 |
Prop_Logic.SOr (Prop_Logic.SNot (Prop_Logic.BoolVar v1), |
8f6465f7021b
ported mono calculus to handle "set" type constructors
blanchet
parents:
41471
diff
changeset
|
491 |
Prop_Logic.BoolVar v2)) |
55889 | 492 |
|
40996
63112be4a469
added "Neq" operator to monotonicity inference module
blanchet
parents:
40995
diff
changeset
|
493 |
fun prop_for_assign (x, a) = |
40986
cfd91aa80701
use boolean pair to encode annotation, which may now take four values
blanchet
parents:
40985
diff
changeset
|
494 |
let val (b1, b2) = bools_from_annotation a in |
41471 | 495 |
Prop_Logic.SAnd (Prop_Logic.BoolVar (fst_var x) |> not b1 ? Prop_Logic.SNot, |
46081
8f6465f7021b
ported mono calculus to handle "set" type constructors
blanchet
parents:
41471
diff
changeset
|
496 |
Prop_Logic.BoolVar (snd_var x) |> not b2 ? Prop_Logic.SNot) |
40986
cfd91aa80701
use boolean pair to encode annotation, which may now take four values
blanchet
parents:
40985
diff
changeset
|
497 |
end |
55889 | 498 |
|
40996
63112be4a469
added "Neq" operator to monotonicity inference module
blanchet
parents:
40995
diff
changeset
|
499 |
fun prop_for_assign_literal (x, (Plus, a)) = prop_for_assign (x, a) |
46081
8f6465f7021b
ported mono calculus to handle "set" type constructors
blanchet
parents:
41471
diff
changeset
|
500 |
| prop_for_assign_literal (x, (Minus, a)) = |
8f6465f7021b
ported mono calculus to handle "set" type constructors
blanchet
parents:
41471
diff
changeset
|
501 |
Prop_Logic.SNot (prop_for_assign (x, a)) |
55889 | 502 |
|
40990
a36d4d869439
adapt monotonicity code to four annotation types
blanchet
parents:
40989
diff
changeset
|
503 |
fun prop_for_atom_assign (A a', a) = prop_for_bool (a = a') |
40996
63112be4a469
added "Neq" operator to monotonicity inference module
blanchet
parents:
40995
diff
changeset
|
504 |
| prop_for_atom_assign (V x, a) = prop_for_assign_literal (x, (Plus, a)) |
55889 | 505 |
|
40990
a36d4d869439
adapt monotonicity code to four annotation types
blanchet
parents:
40989
diff
changeset
|
506 |
fun prop_for_atom_equality (aa1, A a2) = prop_for_atom_assign (aa1, a2) |
a36d4d869439
adapt monotonicity code to four annotation types
blanchet
parents:
40989
diff
changeset
|
507 |
| prop_for_atom_equality (A a1, aa2) = prop_for_atom_assign (aa2, a1) |
a36d4d869439
adapt monotonicity code to four annotation types
blanchet
parents:
40989
diff
changeset
|
508 |
| prop_for_atom_equality (V x1, V x2) = |
41471 | 509 |
Prop_Logic.SAnd (prop_for_bool_var_equality (pairself fst_var (x1, x2)), |
40999
69d0d445c46a
fixed bug in clause handling in monotonicity code, whereby the unsound rule False | x <--> False was used to simplify constraints
blanchet
parents:
40998
diff
changeset
|
510 |
prop_for_bool_var_equality (pairself snd_var (x1, x2))) |
55889 | 511 |
|
41471 | 512 |
val prop_for_assign_clause = Prop_Logic.exists o map prop_for_assign_literal |
55889 | 513 |
|
40995
3cae30b60577
started implementing connectives in new monotonicity calculus
blanchet
parents:
40994
diff
changeset
|
514 |
fun prop_for_exists_var_assign_literal xs a = |
41471 | 515 |
Prop_Logic.exists (map (fn x => prop_for_assign_literal (x, (Plus, a))) xs) |
55889 | 516 |
|
40988 | 517 |
fun prop_for_comp (aa1, aa2, Eq, []) = |
41471 | 518 |
Prop_Logic.SAnd (prop_for_comp (aa1, aa2, Leq, []), |
40990
a36d4d869439
adapt monotonicity code to four annotation types
blanchet
parents:
40989
diff
changeset
|
519 |
prop_for_comp (aa2, aa1, Leq, [])) |
40996
63112be4a469
added "Neq" operator to monotonicity inference module
blanchet
parents:
40995
diff
changeset
|
520 |
| prop_for_comp (aa1, aa2, Neq, []) = |
41471 | 521 |
Prop_Logic.SNot (prop_for_comp (aa1, aa2, Eq, [])) |
40988 | 522 |
| prop_for_comp (aa1, aa2, Leq, []) = |
46081
8f6465f7021b
ported mono calculus to handle "set" type constructors
blanchet
parents:
41471
diff
changeset
|
523 |
Prop_Logic.SOr (prop_for_atom_equality (aa1, aa2), |
8f6465f7021b
ported mono calculus to handle "set" type constructors
blanchet
parents:
41471
diff
changeset
|
524 |
prop_for_atom_assign (aa2, Gen)) |
40988 | 525 |
| prop_for_comp (aa1, aa2, cmp, xs) = |
41471 | 526 |
Prop_Logic.SOr (prop_for_exists_var_assign_literal xs Gen, |
40990
a36d4d869439
adapt monotonicity code to four annotation types
blanchet
parents:
40989
diff
changeset
|
527 |
prop_for_comp (aa1, aa2, cmp, [])) |
33192 | 528 |
|
40990
a36d4d869439
adapt monotonicity code to four annotation types
blanchet
parents:
40989
diff
changeset
|
529 |
fun extract_assigns max_var assigns asgs = |
33192 | 530 |
fold (fn x => fn accum => |
40989 | 531 |
if AList.defined (op =) asgs x then |
33192 | 532 |
accum |
40986
cfd91aa80701
use boolean pair to encode annotation, which may now take four values
blanchet
parents:
40985
diff
changeset
|
533 |
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
|
534 |
(NONE, NONE) => accum |
40993
52ee2a187cdb
support 3 monotonicity calculi in one and fix soundness bug
blanchet
parents:
40991
diff
changeset
|
535 |
| bp => (x, annotation_from_bools (pairself (the_default false) bp)) |
52ee2a187cdb
support 3 monotonicity calculi in one and fix soundness bug
blanchet
parents:
40991
diff
changeset
|
536 |
:: accum) |
40989 | 537 |
(max_var downto 1) asgs |
33192 | 538 |
|
41001
11715564e2ad
removed old baggage from monotonicity calculus -- the "calculus" option didn't really work anyway because of on-the-fly simplifications
blanchet
parents:
41000
diff
changeset
|
539 |
fun print_problem comps clauses = |
11715564e2ad
removed old baggage from monotonicity calculus -- the "calculus" option didn't really work anyway because of on-the-fly simplifications
blanchet
parents:
41000
diff
changeset
|
540 |
trace_msg (fn () => "*** Problem:\n" ^ |
11715564e2ad
removed old baggage from monotonicity calculus -- the "calculus" option didn't really work anyway because of on-the-fly simplifications
blanchet
parents:
41000
diff
changeset
|
541 |
cat_lines (map string_for_comp comps @ |
11715564e2ad
removed old baggage from monotonicity calculus -- the "calculus" option didn't really work anyway because of on-the-fly simplifications
blanchet
parents:
41000
diff
changeset
|
542 |
map (string_for_assign_clause o SOME) clauses)) |
33192 | 543 |
|
40989 | 544 |
fun print_solution asgs = |
545 |
trace_msg (fn () => "*** Solution:\n" ^ |
|
546 |
(asgs |
|
547 |
|> map swap |
|
548 |
|> AList.group (op =) |
|
549 |
|> map (fn (a, xs) => string_for_annotation a ^ ": " ^ |
|
41005
60d931de0709
fixed quantifier handling of new monotonicity calculus
blanchet
parents:
41004
diff
changeset
|
550 |
string_for_vars ", " (sort int_ord xs)) |
55628 | 551 |
|> cat_lines)) |
33192 | 552 |
|
54816
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
blanchet
parents:
47433
diff
changeset
|
553 |
(* The ML solver timeout should correspond more or less to the overhead of invoking an external |
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
blanchet
parents:
47433
diff
changeset
|
554 |
prover. *) |
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
blanchet
parents:
47433
diff
changeset
|
555 |
val ml_solver_timeout = seconds 0.02 |
41007
75275c796b46
use ML SAT solvers up to a certain time limit, then switch to faster solvers with a timeout -- this becomes necessary with the new, more powerful monotonicity calculus
blanchet
parents:
41005
diff
changeset
|
556 |
|
75275c796b46
use ML SAT solvers up to a certain time limit, then switch to faster solvers with a timeout -- this becomes necessary with the new, more powerful monotonicity calculus
blanchet
parents:
41005
diff
changeset
|
557 |
fun solve tac_timeout max_var (comps, clauses) = |
40146 | 558 |
let |
41013
117345744f44
fixed bug in monotonicity solution display, whereby the polarity of literals was ignored
blanchet
parents:
41012
diff
changeset
|
559 |
val asgs = |
117345744f44
fixed bug in monotonicity solution display, whereby the polarity of literals was ignored
blanchet
parents:
41012
diff
changeset
|
560 |
map_filter (fn [(x, (Plus, a))] => SOME (x, a) | _ => NONE) clauses |
40146 | 561 |
fun do_assigns assigns = |
40990
a36d4d869439
adapt monotonicity code to four annotation types
blanchet
parents:
40989
diff
changeset
|
562 |
SOME (extract_assigns max_var assigns asgs |> tap print_solution) |
41001
11715564e2ad
removed old baggage from monotonicity calculus -- the "calculus" option didn't really work anyway because of on-the-fly simplifications
blanchet
parents:
41000
diff
changeset
|
563 |
val _ = print_problem comps clauses |
40993
52ee2a187cdb
support 3 monotonicity calculi in one and fix soundness bug
blanchet
parents:
40991
diff
changeset
|
564 |
val prop = |
46081
8f6465f7021b
ported mono calculus to handle "set" type constructors
blanchet
parents:
41471
diff
changeset
|
565 |
Prop_Logic.all (map prop_for_comp comps @ |
8f6465f7021b
ported mono calculus to handle "set" type constructors
blanchet
parents:
41471
diff
changeset
|
566 |
map prop_for_assign_clause clauses) |
40146 | 567 |
in |
41471 | 568 |
if Prop_Logic.eval (K false) prop then |
40990
a36d4d869439
adapt monotonicity code to four annotation types
blanchet
parents:
40989
diff
changeset
|
569 |
do_assigns (K (SOME false)) |
41471 | 570 |
else if Prop_Logic.eval (K true) prop then |
40990
a36d4d869439
adapt monotonicity code to four annotation types
blanchet
parents:
40989
diff
changeset
|
571 |
do_assigns (K (SOME true)) |
40146 | 572 |
else |
573 |
let |
|
574 |
(* use the first ML solver (to avoid startup overhead) *) |
|
41007
75275c796b46
use ML SAT solvers up to a certain time limit, then switch to faster solvers with a timeout -- this becomes necessary with the new, more powerful monotonicity calculus
blanchet
parents:
41005
diff
changeset
|
575 |
val (ml_solvers, nonml_solvers) = |
56147 | 576 |
SatSolver.get_solvers () |
56845 | 577 |
|> List.partition (member (op =) ["dptsat", "dpll_p"] o fst) |
41007
75275c796b46
use ML SAT solvers up to a certain time limit, then switch to faster solvers with a timeout -- this becomes necessary with the new, more powerful monotonicity calculus
blanchet
parents:
41005
diff
changeset
|
578 |
val res = |
75275c796b46
use ML SAT solvers up to a certain time limit, then switch to faster solvers with a timeout -- this becomes necessary with the new, more powerful monotonicity calculus
blanchet
parents:
41005
diff
changeset
|
579 |
if null nonml_solvers then |
54816
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
blanchet
parents:
47433
diff
changeset
|
580 |
TimeLimit.timeLimit tac_timeout (snd (hd ml_solvers)) prop |
41007
75275c796b46
use ML SAT solvers up to a certain time limit, then switch to faster solvers with a timeout -- this becomes necessary with the new, more powerful monotonicity calculus
blanchet
parents:
41005
diff
changeset
|
581 |
else |
54816
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
blanchet
parents:
47433
diff
changeset
|
582 |
TimeLimit.timeLimit ml_solver_timeout (snd (hd ml_solvers)) prop |
41007
75275c796b46
use ML SAT solvers up to a certain time limit, then switch to faster solvers with a timeout -- this becomes necessary with the new, more powerful monotonicity calculus
blanchet
parents:
41005
diff
changeset
|
583 |
handle TimeLimit.TimeOut => |
54816
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
blanchet
parents:
47433
diff
changeset
|
584 |
TimeLimit.timeLimit tac_timeout |
10d48c2a3e32
made timeouts in Sledgehammer not be 'option's -- simplified lots of code
blanchet
parents:
47433
diff
changeset
|
585 |
(SatSolver.invoke_solver "auto") prop |
40146 | 586 |
in |
41007
75275c796b46
use ML SAT solvers up to a certain time limit, then switch to faster solvers with a timeout -- this becomes necessary with the new, more powerful monotonicity calculus
blanchet
parents:
41005
diff
changeset
|
587 |
case res of |
40146 | 588 |
SatSolver.SATISFIABLE assigns => do_assigns assigns |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
589 |
| _ => (trace_msg (K "*** Unsolvable"); NONE) |
40146 | 590 |
end |
41007
75275c796b46
use ML SAT solvers up to a certain time limit, then switch to faster solvers with a timeout -- this becomes necessary with the new, more powerful monotonicity calculus
blanchet
parents:
41005
diff
changeset
|
591 |
handle TimeLimit.TimeOut => (trace_msg (K "*** Timed out"); NONE) |
40146 | 592 |
end |
33192 | 593 |
|
40998
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
594 |
type mcontext = |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
595 |
{bound_Ts: typ list, |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
596 |
bound_Ms: mtyp list, |
40998
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
597 |
frame: (int * annotation_atom) list, |
55889 | 598 |
frees: ((string * typ) * mtyp) list, |
599 |
consts: ((string * typ) * mtyp) list} |
|
33192 | 600 |
|
40998
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
601 |
fun string_for_bound ctxt Ms (j, aa) = |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
602 |
Syntax.string_of_term ctxt (Bound (length Ms - j - 1)) ^ " :\<^bsup>" ^ |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
603 |
string_for_annotation_atom aa ^ "\<^esup> " ^ |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
604 |
string_for_mtype (nth Ms (length Ms - j - 1)) |
55889 | 605 |
|
40998
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
606 |
fun string_for_free relevant_frees ((s, _), M) = |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
607 |
if member (op =) relevant_frees s then SOME (s ^ " : " ^ string_for_mtype M) |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
608 |
else NONE |
55889 | 609 |
|
41109 | 610 |
fun string_for_mcontext ctxt t ({bound_Ms, frame, frees, ...} : mcontext) = |
40998
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
611 |
(map_filter (string_for_free (Term.add_free_names t [])) frees @ |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
612 |
map (string_for_bound ctxt bound_Ms) frame) |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
613 |
|> commas |> enclose "[" "]" |
33192 | 614 |
|
40987
7d52af07bff1
added frame component to Gamma in monotonicity calculus
blanchet
parents:
40986
diff
changeset
|
615 |
val initial_gamma = |
40998
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
616 |
{bound_Ts = [], bound_Ms = [], frame = [], frees = [], consts = []} |
33192 | 617 |
|
40998
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
618 |
fun push_bound aa T M {bound_Ts, bound_Ms, frame, frees, consts} = |
40987
7d52af07bff1
added frame component to Gamma in monotonicity calculus
blanchet
parents:
40986
diff
changeset
|
619 |
{bound_Ts = T :: bound_Ts, bound_Ms = M :: bound_Ms, |
40998
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
620 |
frame = frame @ [(length bound_Ts, aa)], frees = frees, consts = consts} |
55889 | 621 |
|
40998
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
622 |
fun pop_bound {bound_Ts, bound_Ms, frame, frees, consts} = |
40993
52ee2a187cdb
support 3 monotonicity calculi in one and fix soundness bug
blanchet
parents:
40991
diff
changeset
|
623 |
{bound_Ts = tl bound_Ts, bound_Ms = tl bound_Ms, |
40998
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
624 |
frame = frame |> filter_out (fn (j, _) => j = length bound_Ts - 1), |
40987
7d52af07bff1
added frame component to Gamma in monotonicity calculus
blanchet
parents:
40986
diff
changeset
|
625 |
frees = frees, consts = consts} |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
626 |
handle List.Empty => initial_gamma (* FIXME: needed? *) |
33192 | 627 |
|
40998
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
628 |
fun set_frame frame ({bound_Ts, bound_Ms, frees, consts, ...} : mcontext) = |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
629 |
{bound_Ts = bound_Ts, bound_Ms = bound_Ms, frame = frame, frees = frees, |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
630 |
consts = consts} |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
631 |
|
40998
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
632 |
fun add_comp_frame aa cmp = fold (add_annotation_atom_comp cmp [] aa o snd) |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
633 |
|
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
634 |
fun add_bound_frame j frame = |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
635 |
let |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
636 |
val (new_frame, gen_frame) = List.partition (curry (op =) j o fst) frame |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
637 |
in |
40998
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
638 |
add_comp_frame (A New) Leq new_frame |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
639 |
#> add_comp_frame (A Gen) Eq gen_frame |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
640 |
end |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
641 |
|
40997
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
642 |
fun fresh_frame ({max_fresh, ...} : mdata) fls tru = |
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
643 |
map (apsnd (fn aa => |
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
644 |
case (aa, fls, tru) of |
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
645 |
(A Fls, SOME a, _) => A a |
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
646 |
| (A Tru, _, SOME a) => A a |
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
647 |
| (A Gen, _, _) => A Gen |
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
648 |
| _ => V (Unsynchronized.inc max_fresh))) |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
649 |
|
40997
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
650 |
fun conj_clauses res_aa aa1 aa2 = |
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
651 |
[[(aa1, (Neq, Tru)), (aa2, (Neq, Tru)), (res_aa, (Eq, Tru))], |
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
652 |
[(aa1, (Neq, Fls)), (res_aa, (Eq, Fls))], |
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
653 |
[(aa2, (Neq, Fls)), (res_aa, (Eq, Fls))], |
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
654 |
[(aa1, (Neq, Gen)), (aa2, (Eq, Fls)), (res_aa, (Eq, Gen))], |
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
655 |
[(aa1, (Neq, New)), (aa2, (Eq, Fls)), (res_aa, (Eq, Gen))], |
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
656 |
[(aa1, (Eq, Fls)), (aa2, (Neq, Gen)), (res_aa, (Eq, Gen))], |
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
657 |
[(aa1, (Eq, Fls)), (aa2, (Neq, New)), (res_aa, (Eq, Gen))]] |
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
658 |
|
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
659 |
fun disj_clauses res_aa aa1 aa2 = |
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
660 |
[[(aa1, (Neq, Tru)), (res_aa, (Eq, Tru))], |
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
661 |
[(aa2, (Neq, Tru)), (res_aa, (Eq, Tru))], |
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
662 |
[(aa1, (Neq, Fls)), (aa2, (Neq, Fls)), (res_aa, (Eq, Fls))], |
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
663 |
[(aa1, (Neq, Gen)), (aa2, (Eq, Tru)), (res_aa, (Eq, Gen))], |
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
664 |
[(aa1, (Neq, New)), (aa2, (Eq, Tru)), (res_aa, (Eq, Gen))], |
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
665 |
[(aa1, (Eq, Tru)), (aa2, (Neq, Gen)), (res_aa, (Eq, Gen))], |
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
666 |
[(aa1, (Eq, Tru)), (aa2, (Neq, New)), (res_aa, (Eq, Gen))]] |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
667 |
|
40997
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
668 |
fun imp_clauses res_aa aa1 aa2 = |
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
669 |
[[(aa1, (Neq, Fls)), (res_aa, (Eq, Tru))], |
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
670 |
[(aa2, (Neq, Tru)), (res_aa, (Eq, Tru))], |
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
671 |
[(aa1, (Neq, Tru)), (aa2, (Neq, Fls)), (res_aa, (Eq, Fls))], |
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
672 |
[(aa1, (Neq, Gen)), (aa2, (Eq, Tru)), (res_aa, (Eq, Gen))], |
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
673 |
[(aa1, (Neq, New)), (aa2, (Eq, Tru)), (res_aa, (Eq, Gen))], |
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
674 |
[(aa1, (Eq, Fls)), (aa2, (Neq, Gen)), (res_aa, (Eq, Gen))], |
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
675 |
[(aa1, (Eq, Fls)), (aa2, (Neq, New)), (res_aa, (Eq, Gen))]] |
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
676 |
|
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
677 |
val meta_conj_spec = ("\<and>", conj_clauses) |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
678 |
val meta_imp_spec = ("\<implies>", imp_clauses) |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
679 |
val conj_spec = ("\<and>", conj_clauses) |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
680 |
val disj_spec = ("\<or>", disj_clauses) |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
681 |
val imp_spec = ("\<implies>", imp_clauses) |
41003
7e2a7bd55a00
proper handling of frames for connectives in monotonicity calculus
blanchet
parents:
41002
diff
changeset
|
682 |
|
40999
69d0d445c46a
fixed bug in clause handling in monotonicity code, whereby the unsound rule False | x <--> False was used to simplify constraints
blanchet
parents:
40998
diff
changeset
|
683 |
fun add_annotation_clause_from_quasi_clause _ NONE = NONE |
69d0d445c46a
fixed bug in clause handling in monotonicity code, whereby the unsound rule False | x <--> False was used to simplify constraints
blanchet
parents:
40998
diff
changeset
|
684 |
| add_annotation_clause_from_quasi_clause [] accum = accum |
69d0d445c46a
fixed bug in clause handling in monotonicity code, whereby the unsound rule False | x <--> False was used to simplify constraints
blanchet
parents:
40998
diff
changeset
|
685 |
| add_annotation_clause_from_quasi_clause ((aa, (cmp, a)) :: rest) accum = |
69d0d445c46a
fixed bug in clause handling in monotonicity code, whereby the unsound rule False | x <--> False was used to simplify constraints
blanchet
parents:
40998
diff
changeset
|
686 |
case aa of |
69d0d445c46a
fixed bug in clause handling in monotonicity code, whereby the unsound rule False | x <--> False was used to simplify constraints
blanchet
parents:
40998
diff
changeset
|
687 |
A a' => if annotation_comp cmp a' a then NONE |
69d0d445c46a
fixed bug in clause handling in monotonicity code, whereby the unsound rule False | x <--> False was used to simplify constraints
blanchet
parents:
40998
diff
changeset
|
688 |
else add_annotation_clause_from_quasi_clause rest accum |
69d0d445c46a
fixed bug in clause handling in monotonicity code, whereby the unsound rule False | x <--> False was used to simplify constraints
blanchet
parents:
40998
diff
changeset
|
689 |
| V x => add_annotation_clause_from_quasi_clause rest accum |
69d0d445c46a
fixed bug in clause handling in monotonicity code, whereby the unsound rule False | x <--> False was used to simplify constraints
blanchet
parents:
40998
diff
changeset
|
690 |
|> Option.map (cons (x, (sign_for_comp_op cmp, a))) |
40995
3cae30b60577
started implementing connectives in new monotonicity calculus
blanchet
parents:
40994
diff
changeset
|
691 |
|
41000
4bbff1684465
implemented All rules from new monotonicity calculus
blanchet
parents:
40999
diff
changeset
|
692 |
fun assign_clause_from_quasi_clause unless = |
4bbff1684465
implemented All rules from new monotonicity calculus
blanchet
parents:
40999
diff
changeset
|
693 |
add_annotation_clause_from_quasi_clause unless (SOME []) |
40995
3cae30b60577
started implementing connectives in new monotonicity calculus
blanchet
parents:
40994
diff
changeset
|
694 |
|
40997
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
695 |
fun add_connective_var conn mk_quasi_clauses res_aa aa1 aa2 = |
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
696 |
(trace_msg (fn () => "*** Add " ^ string_for_annotation_atom res_aa ^ " = " ^ |
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
697 |
string_for_annotation_atom aa1 ^ " " ^ conn ^ " " ^ |
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
698 |
string_for_annotation_atom aa2); |
40999
69d0d445c46a
fixed bug in clause handling in monotonicity code, whereby the unsound rule False | x <--> False was used to simplify constraints
blanchet
parents:
40998
diff
changeset
|
699 |
fold (add_assign_clause o assign_clause_from_quasi_clause) |
69d0d445c46a
fixed bug in clause handling in monotonicity code, whereby the unsound rule False | x <--> False was used to simplify constraints
blanchet
parents:
40998
diff
changeset
|
700 |
(mk_quasi_clauses res_aa aa1 aa2)) |
55889 | 701 |
|
40997
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
702 |
fun add_connective_frames conn mk_quasi_clauses res_frame frame1 frame2 = |
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
703 |
fold I (map3 (fn (_, res_aa) => fn (_, aa1) => fn (_, aa2) => |
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
704 |
add_connective_var conn mk_quasi_clauses res_aa aa1 aa2) |
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
705 |
res_frame frame1 frame2) |
40995
3cae30b60577
started implementing connectives in new monotonicity calculus
blanchet
parents:
40994
diff
changeset
|
706 |
|
41109 | 707 |
fun kill_unused_in_frame is_in (accum as ({frame, ...} : mcontext, _)) = |
40998
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
708 |
let val (used_frame, unused_frame) = List.partition is_in frame in |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
709 |
accum |>> set_frame used_frame |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
710 |
||> add_comp_frame (A Gen) Eq unused_frame |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
711 |
end |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
712 |
|
41109 | 713 |
fun split_frame is_in_fun (gamma as {frame, ...} : mcontext, cset) = |
40998
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
714 |
let |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
715 |
fun bubble fun_frame arg_frame [] cset = |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
716 |
((rev fun_frame, rev arg_frame), cset) |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
717 |
| bubble fun_frame arg_frame ((bound as (_, aa)) :: rest) cset = |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
718 |
if is_in_fun bound then |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
719 |
bubble (bound :: fun_frame) arg_frame rest |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
720 |
(cset |> add_comp_frame aa Leq arg_frame) |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
721 |
else |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
722 |
bubble fun_frame (bound :: arg_frame) rest cset |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
723 |
in cset |> bubble [] [] frame ||> pair gamma end |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
724 |
|
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
725 |
fun add_annotation_atom_comp_alt _ (A Gen) _ _ = I |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
726 |
| add_annotation_atom_comp_alt _ (A _) _ _ = |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
727 |
(trace_msg (K "*** Expected G"); raise UNSOLVABLE ()) |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
728 |
| add_annotation_atom_comp_alt cmp (V x) aa1 aa2 = |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
729 |
add_annotation_atom_comp cmp [x] aa1 aa2 |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
730 |
|
40999
69d0d445c46a
fixed bug in clause handling in monotonicity code, whereby the unsound rule False | x <--> False was used to simplify constraints
blanchet
parents:
40998
diff
changeset
|
731 |
fun add_arg_order1 ((_, aa), (_, prev_aa)) = I |
40998
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
732 |
add_annotation_atom_comp_alt Neq prev_aa (A Gen) aa |
55889 | 733 |
|
40999
69d0d445c46a
fixed bug in clause handling in monotonicity code, whereby the unsound rule False | x <--> False was used to simplify constraints
blanchet
parents:
40998
diff
changeset
|
734 |
fun add_app1 fun_aa ((_, res_aa), (_, arg_aa)) = I |
69d0d445c46a
fixed bug in clause handling in monotonicity code, whereby the unsound rule False | x <--> False was used to simplify constraints
blanchet
parents:
40998
diff
changeset
|
735 |
let |
69d0d445c46a
fixed bug in clause handling in monotonicity code, whereby the unsound rule False | x <--> False was used to simplify constraints
blanchet
parents:
40998
diff
changeset
|
736 |
val clause = [(arg_aa, (Eq, New)), (res_aa, (Eq, Gen))] |
69d0d445c46a
fixed bug in clause handling in monotonicity code, whereby the unsound rule False | x <--> False was used to simplify constraints
blanchet
parents:
40998
diff
changeset
|
737 |
|> assign_clause_from_quasi_clause |
69d0d445c46a
fixed bug in clause handling in monotonicity code, whereby the unsound rule False | x <--> False was used to simplify constraints
blanchet
parents:
40998
diff
changeset
|
738 |
in |
69d0d445c46a
fixed bug in clause handling in monotonicity code, whereby the unsound rule False | x <--> False was used to simplify constraints
blanchet
parents:
40998
diff
changeset
|
739 |
trace_msg (fn () => "*** Add " ^ string_for_assign_clause clause); |
69d0d445c46a
fixed bug in clause handling in monotonicity code, whereby the unsound rule False | x <--> False was used to simplify constraints
blanchet
parents:
40998
diff
changeset
|
740 |
apsnd (add_assign_clause clause) |
69d0d445c46a
fixed bug in clause handling in monotonicity code, whereby the unsound rule False | x <--> False was used to simplify constraints
blanchet
parents:
40998
diff
changeset
|
741 |
#> add_annotation_atom_comp_alt Leq arg_aa fun_aa res_aa |
69d0d445c46a
fixed bug in clause handling in monotonicity code, whereby the unsound rule False | x <--> False was used to simplify constraints
blanchet
parents:
40998
diff
changeset
|
742 |
end |
55889 | 743 |
|
40998
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
744 |
fun add_app _ [] [] = I |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
745 |
| add_app fun_aa res_frame arg_frame = |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
746 |
add_comp_frame (A New) Leq arg_frame |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
747 |
#> fold add_arg_order1 (tl arg_frame ~~ (fst (split_last arg_frame))) |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
748 |
#> fold (add_app1 fun_aa) (res_frame ~~ arg_frame) |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
749 |
|
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
750 |
fun consider_connective mdata (conn, mk_quasi_clauses) do_t1 do_t2 |
41003
7e2a7bd55a00
proper handling of frames for connectives in monotonicity calculus
blanchet
parents:
41002
diff
changeset
|
751 |
(accum as ({frame, ...}, _)) = |
7e2a7bd55a00
proper handling of frames for connectives in monotonicity calculus
blanchet
parents:
41002
diff
changeset
|
752 |
let |
7e2a7bd55a00
proper handling of frames for connectives in monotonicity calculus
blanchet
parents:
41002
diff
changeset
|
753 |
val frame1 = fresh_frame mdata (SOME Tru) NONE frame |
7e2a7bd55a00
proper handling of frames for connectives in monotonicity calculus
blanchet
parents:
41002
diff
changeset
|
754 |
val frame2 = fresh_frame mdata (SOME Fls) NONE frame |
7e2a7bd55a00
proper handling of frames for connectives in monotonicity calculus
blanchet
parents:
41002
diff
changeset
|
755 |
in |
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
756 |
accum |>> set_frame frame1 |> do_t1 |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
757 |
|>> set_frame frame2 |> do_t2 |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
758 |
|>> set_frame frame |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
759 |
||> apsnd (add_connective_frames conn mk_quasi_clauses frame frame1 |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
760 |
frame2) |
41003
7e2a7bd55a00
proper handling of frames for connectives in monotonicity calculus
blanchet
parents:
41002
diff
changeset
|
761 |
end |
7e2a7bd55a00
proper handling of frames for connectives in monotonicity calculus
blanchet
parents:
41002
diff
changeset
|
762 |
|
55888 | 763 |
fun consider_term (mdata as {hol_ctxt = {ctxt, ...}, alpha_T, max_fresh, ...}) = |
33192 | 764 |
let |
37476
0681e46b4022
optimized code generated for datatype cases + more;
blanchet
parents:
37267
diff
changeset
|
765 |
fun is_enough_eta_expanded t = |
0681e46b4022
optimized code generated for datatype cases + more;
blanchet
parents:
37267
diff
changeset
|
766 |
case strip_comb t of |
55888 | 767 |
(Const x, ts) => the_default 0 (arity_of_built_in_const x) <= length ts |
37476
0681e46b4022
optimized code generated for datatype cases + more;
blanchet
parents:
37267
diff
changeset
|
768 |
| _ => true |
35672 | 769 |
val mtype_for = fresh_mtype_for_type mdata false |
46083 | 770 |
fun mtype_for_set x T = MFun (mtype_for (pseudo_domain_type T), V x, bool_M) |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
771 |
fun do_all T (gamma, cset) = |
33192 | 772 |
let |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
773 |
val abs_M = mtype_for (domain_type (domain_type T)) |
41012
e5a23ffb5524
improve precision of forall in constancy part of the monotonicity calculus
blanchet
parents:
41011
diff
changeset
|
774 |
val x = Unsynchronized.inc max_fresh |
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
775 |
val body_M = mtype_for (body_type T) |
33192 | 776 |
in |
41012
e5a23ffb5524
improve precision of forall in constancy part of the monotonicity calculus
blanchet
parents:
41011
diff
changeset
|
777 |
(MFun (MFun (abs_M, V x, body_M), A Gen, body_M), |
e5a23ffb5524
improve precision of forall in constancy part of the monotonicity calculus
blanchet
parents:
41011
diff
changeset
|
778 |
(gamma, cset |> add_mtype_is_complete [(x, (Plus, Tru))] abs_M)) |
33192 | 779 |
end |
780 |
fun do_equals T (gamma, cset) = |
|
40997
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
781 |
let |
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
782 |
val M = mtype_for (domain_type T) |
41014
e538a4f9dd86
add more flexibility to the monotonicity calculus: instead of hardcoding F-arrows, also allow G-arrows, simulating applications of the Sub rule
blanchet
parents:
41013
diff
changeset
|
783 |
val x = Unsynchronized.inc max_fresh |
40997
67e11a73532a
implemented connectives in new monotonicity calculus
blanchet
parents:
40996
diff
changeset
|
784 |
in |
41014
e538a4f9dd86
add more flexibility to the monotonicity calculus: instead of hardcoding F-arrows, also allow G-arrows, simulating applications of the Sub rule
blanchet
parents:
41013
diff
changeset
|
785 |
(MFun (M, A Gen, MFun (M, V x, mtype_for (nth_range_type 2 T))), |
41000
4bbff1684465
implemented All rules from new monotonicity calculus
blanchet
parents:
40999
diff
changeset
|
786 |
(gamma, cset |> add_mtype_is_concrete [] M |
41014
e538a4f9dd86
add more flexibility to the monotonicity calculus: instead of hardcoding F-arrows, also allow G-arrows, simulating applications of the Sub rule
blanchet
parents:
41013
diff
changeset
|
787 |
|> add_annotation_atom_comp Leq [] (A Fls) (V x))) |
33192 | 788 |
end |
789 |
fun do_robust_set_operation T (gamma, cset) = |
|
790 |
let |
|
791 |
val set_T = domain_type T |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
792 |
val M1 = mtype_for set_T |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
793 |
val M2 = mtype_for set_T |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
794 |
val M3 = mtype_for set_T |
33192 | 795 |
in |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
796 |
(MFun (M1, A Gen, MFun (M2, A Gen, M3)), |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
797 |
(gamma, cset |> add_is_sub_mtype M1 M3 |> add_is_sub_mtype M2 M3)) |
33192 | 798 |
end |
799 |
fun do_fragile_set_operation T (gamma, cset) = |
|
800 |
let |
|
801 |
val set_T = domain_type T |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
802 |
val set_M = mtype_for set_T |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
803 |
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
|
804 |
if T = set_T then set_M |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
805 |
else MFun (custom_mtype_for T1, A Gen, custom_mtype_for T2) |
46081
8f6465f7021b
ported mono calculus to handle "set" type constructors
blanchet
parents:
41471
diff
changeset
|
806 |
| custom_mtype_for (Type (@{type_name set}, [T'])) = |
8f6465f7021b
ported mono calculus to handle "set" type constructors
blanchet
parents:
41471
diff
changeset
|
807 |
custom_mtype_for (T' --> bool_T) |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
808 |
| custom_mtype_for T = mtype_for T |
33192 | 809 |
in |
41000
4bbff1684465
implemented All rules from new monotonicity calculus
blanchet
parents:
40999
diff
changeset
|
810 |
(custom_mtype_for T, (gamma, cset |> add_mtype_is_concrete [] set_M)) |
33192 | 811 |
end |
812 |
fun do_pair_constr T accum = |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
813 |
case mtype_for (nth_range_type 2 T) of |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
814 |
M as MPair (a_M, b_M) => |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
815 |
(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
|
816 |
| M => raise MTYPE ("Nitpick_Mono.consider_term.do_pair_constr", [M], []) |
33192 | 817 |
fun do_nth_pair_sel n T = |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
818 |
case mtype_for (domain_type T) of |
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
819 |
M as MPair (a_M, b_M) => |
40985
8b870370c26f
started generalizing monotonicity code to accommodate new calculus
blanchet
parents:
40722
diff
changeset
|
820 |
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
|
821 |
| M => raise MTYPE ("Nitpick_Mono.consider_term.do_nth_pair_sel", [M], []) |
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
822 |
and do_connect spec t1 t2 accum = |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
823 |
(bool_M, consider_connective mdata spec (snd o do_term t1) |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
824 |
(snd o do_term t2) accum) |
40998
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
825 |
and do_term t |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
826 |
(accum as (gamma as {bound_Ts, bound_Ms, frame, frees, consts}, |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
827 |
cset)) = |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
828 |
(trace_msg (fn () => " " ^ string_for_mcontext ctxt t gamma ^ |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
829 |
" \<turnstile> " ^ Syntax.string_of_term ctxt t ^ |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
830 |
" : _?"); |
40993
52ee2a187cdb
support 3 monotonicity calculi in one and fix soundness bug
blanchet
parents:
40991
diff
changeset
|
831 |
case t of |
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
832 |
@{const False} => (bool_M, accum ||> add_comp_frame (A Fls) Leq frame) |
41018 | 833 |
| Const (@{const_name None}, T) => |
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
834 |
(mtype_for T, accum ||> add_comp_frame (A Fls) Leq frame) |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
835 |
| @{const True} => (bool_M, accum ||> add_comp_frame (A Tru) Leq frame) |
41016
343cdf923c02
introduced hack to exploit the symmetry of equality in monotonicity calculus
blanchet
parents:
41014
diff
changeset
|
836 |
| (t0 as Const (@{const_name HOL.eq}, _)) $ Bound 0 $ t2 => |
343cdf923c02
introduced hack to exploit the symmetry of equality in monotonicity calculus
blanchet
parents:
41014
diff
changeset
|
837 |
(* hack to exploit symmetry of equality when typing "insert" *) |
343cdf923c02
introduced hack to exploit the symmetry of equality in monotonicity calculus
blanchet
parents:
41014
diff
changeset
|
838 |
(if t2 = Bound 0 then do_term @{term True} |
343cdf923c02
introduced hack to exploit the symmetry of equality in monotonicity calculus
blanchet
parents:
41014
diff
changeset
|
839 |
else do_term (t0 $ t2 $ Bound 0)) accum |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
840 |
| Const (x as (s, T)) => |
40993
52ee2a187cdb
support 3 monotonicity calculi in one and fix soundness bug
blanchet
parents:
40991
diff
changeset
|
841 |
(case AList.lookup (op =) consts x of |
52ee2a187cdb
support 3 monotonicity calculi in one and fix soundness bug
blanchet
parents:
40991
diff
changeset
|
842 |
SOME M => (M, accum) |
52ee2a187cdb
support 3 monotonicity calculi in one and fix soundness bug
blanchet
parents:
40991
diff
changeset
|
843 |
| NONE => |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
844 |
if not (could_exist_alpha_subtype alpha_T T) then |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
845 |
(mtype_for T, accum) |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
846 |
else case s of |
56245 | 847 |
@{const_name Pure.all} => do_all T accum |
848 |
| @{const_name Pure.eq} => do_equals T accum |
|
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
849 |
| @{const_name All} => do_all T accum |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
850 |
| @{const_name Ex} => |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
851 |
let val set_T = domain_type T in |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
852 |
do_term (Abs (Name.uu, set_T, |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
853 |
@{const Not} $ (HOLogic.mk_eq |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
854 |
(Abs (Name.uu, domain_type set_T, |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
855 |
@{const False}), |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
856 |
Bound 0)))) accum |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
857 |
end |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
858 |
| @{const_name HOL.eq} => do_equals T accum |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
859 |
| @{const_name The} => |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
860 |
(trace_msg (K "*** The"); raise UNSOLVABLE ()) |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
861 |
| @{const_name Eps} => |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
862 |
(trace_msg (K "*** Eps"); raise UNSOLVABLE ()) |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
863 |
| @{const_name If} => |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
864 |
do_robust_set_operation (range_type T) accum |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
865 |
|>> curry3 MFun bool_M (A Gen) |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
866 |
| @{const_name Pair} => do_pair_constr T accum |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
867 |
| @{const_name fst} => do_nth_pair_sel 0 T accum |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
868 |
| @{const_name snd} => do_nth_pair_sel 1 T accum |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
869 |
| @{const_name Id} => |
46085
447cda88adfe
fixed a few more bugs in \Nitpick's new "set" support
blanchet
parents:
46083
diff
changeset
|
870 |
(MFun (mtype_for (elem_type T), A Gen, bool_M), accum) |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
871 |
| @{const_name converse} => |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
872 |
let |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
873 |
val x = Unsynchronized.inc max_fresh |
46081
8f6465f7021b
ported mono calculus to handle "set" type constructors
blanchet
parents:
41471
diff
changeset
|
874 |
val ab_set_M = domain_type T |> mtype_for_set x |
8f6465f7021b
ported mono calculus to handle "set" type constructors
blanchet
parents:
41471
diff
changeset
|
875 |
val ba_set_M = range_type T |> mtype_for_set x |
41011
5c2f16eae066
added some missing well-annotatedness constraints to monotonicity calculus
blanchet
parents:
41009
diff
changeset
|
876 |
in |
5c2f16eae066
added some missing well-annotatedness constraints to monotonicity calculus
blanchet
parents:
41009
diff
changeset
|
877 |
(MFun (ab_set_M, A Gen, ba_set_M), |
5c2f16eae066
added some missing well-annotatedness constraints to monotonicity calculus
blanchet
parents:
41009
diff
changeset
|
878 |
accum ||> add_annotation_atom_comp Neq [] (V x) (A New)) |
5c2f16eae066
added some missing well-annotatedness constraints to monotonicity calculus
blanchet
parents:
41009
diff
changeset
|
879 |
end |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
880 |
| @{const_name trancl} => do_fragile_set_operation T accum |
47433
07f4bf913230
renamed "rel_comp" to "relcomp" (to be consistent with, e.g., "relpow")
griff
parents:
46085
diff
changeset
|
881 |
| @{const_name relcomp} => |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
882 |
let |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
883 |
val x = Unsynchronized.inc max_fresh |
46081
8f6465f7021b
ported mono calculus to handle "set" type constructors
blanchet
parents:
41471
diff
changeset
|
884 |
val bc_set_M = domain_type T |> mtype_for_set x |
8f6465f7021b
ported mono calculus to handle "set" type constructors
blanchet
parents:
41471
diff
changeset
|
885 |
val ab_set_M = domain_type (range_type T) |> mtype_for_set x |
8f6465f7021b
ported mono calculus to handle "set" type constructors
blanchet
parents:
41471
diff
changeset
|
886 |
val ac_set_M = nth_range_type 2 T |> mtype_for_set x |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
887 |
in |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
888 |
(MFun (bc_set_M, A Gen, MFun (ab_set_M, A Gen, ac_set_M)), |
41011
5c2f16eae066
added some missing well-annotatedness constraints to monotonicity calculus
blanchet
parents:
41009
diff
changeset
|
889 |
accum ||> add_annotation_atom_comp Neq [] (V x) (A New)) |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
890 |
end |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
891 |
| @{const_name finite} => |
41011
5c2f16eae066
added some missing well-annotatedness constraints to monotonicity calculus
blanchet
parents:
41009
diff
changeset
|
892 |
let |
46081
8f6465f7021b
ported mono calculus to handle "set" type constructors
blanchet
parents:
41471
diff
changeset
|
893 |
val M1 = mtype_for (elem_type (domain_type T)) |
41011
5c2f16eae066
added some missing well-annotatedness constraints to monotonicity calculus
blanchet
parents:
41009
diff
changeset
|
894 |
val a = if exists_alpha_sub_mtype M1 then Fls else Gen |
5c2f16eae066
added some missing well-annotatedness constraints to monotonicity calculus
blanchet
parents:
41009
diff
changeset
|
895 |
in (MFun (MFun (M1, A a, bool_M), A Gen, bool_M), accum) end |
41050
effbaa323cf0
updated monotonicity calculus w.r.t. set products
blanchet
parents:
41049
diff
changeset
|
896 |
| @{const_name prod} => |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
897 |
let |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
898 |
val x = Unsynchronized.inc max_fresh |
46081
8f6465f7021b
ported mono calculus to handle "set" type constructors
blanchet
parents:
41471
diff
changeset
|
899 |
val a_set_M = domain_type T |> mtype_for_set x |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
900 |
val b_set_M = |
46081
8f6465f7021b
ported mono calculus to handle "set" type constructors
blanchet
parents:
41471
diff
changeset
|
901 |
range_type (domain_type (range_type T)) |> mtype_for_set x |
8f6465f7021b
ported mono calculus to handle "set" type constructors
blanchet
parents:
41471
diff
changeset
|
902 |
val ab_set_M = nth_range_type 2 T |> mtype_for_set x |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
903 |
in |
41050
effbaa323cf0
updated monotonicity calculus w.r.t. set products
blanchet
parents:
41049
diff
changeset
|
904 |
(MFun (a_set_M, A Gen, MFun (b_set_M, A Gen, ab_set_M)), |
41011
5c2f16eae066
added some missing well-annotatedness constraints to monotonicity calculus
blanchet
parents:
41009
diff
changeset
|
905 |
accum ||> add_annotation_atom_comp Neq [] (V x) (A New)) |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
906 |
end |
40993
52ee2a187cdb
support 3 monotonicity calculi in one and fix soundness bug
blanchet
parents:
40991
diff
changeset
|
907 |
| _ => |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
908 |
if s = @{const_name safe_The} then |
33192 | 909 |
let |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
910 |
val a_set_M = mtype_for (domain_type T) |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
911 |
val a_M = dest_MFun a_set_M |> #1 |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
912 |
in (MFun (a_set_M, A Gen, a_M), accum) end |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
913 |
else if s = @{const_name ord_class.less_eq} andalso |
46081
8f6465f7021b
ported mono calculus to handle "set" type constructors
blanchet
parents:
41471
diff
changeset
|
914 |
is_set_like_type (domain_type T) then |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
915 |
do_fragile_set_operation T accum |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
916 |
else if is_sel s then |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
917 |
(mtype_for_sel mdata x, accum) |
55888 | 918 |
else if is_constr ctxt x then |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
919 |
(mtype_for_constr mdata x, accum) |
55888 | 920 |
else if is_built_in_const x then |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
921 |
(fresh_mtype_for_type mdata true T, accum) |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
922 |
else |
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
923 |
let val M = mtype_for T in |
40998
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
924 |
(M, ({bound_Ts = bound_Ts, bound_Ms = bound_Ms, frame = frame, |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
925 |
frees = frees, consts = (x, M) :: consts}, cset)) |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
926 |
end) |
40998
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
927 |
||> apsnd (add_comp_frame (A Gen) Eq frame) |
33192 | 928 |
| Free (x as (_, T)) => |
929 |
(case AList.lookup (op =) frees x of |
|
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
930 |
SOME M => (M, accum) |
33192 | 931 |
| NONE => |
35385
29f81babefd7
improved precision of infinite "shallow" datatypes in Nitpick;
blanchet
parents:
35384
diff
changeset
|
932 |
let val M = mtype_for T in |
40998
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
933 |
(M, ({bound_Ts = bound_Ts, bound_Ms = bound_Ms, frame = frame, |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
934 |
frees = (x, M) :: frees, consts = consts}, cset)) |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
935 |
end) |
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
936 |
||> apsnd (add_comp_frame (A Gen) Eq frame) |
38179 | 937 |
| Var _ => (trace_msg (K "*** Var"); raise UNSOLVABLE ()) |
40994
3bdb8df0daf0
more work on frames in the new monotonicity calculus
blanchet
parents:
40993
diff
changeset
|
938 |
| Bound j => |
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
939 |
(nth bound_Ms j, |
40998
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
940 |
accum ||> add_bound_frame (length bound_Ts - j - 1) frame) |
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
941 |
| Abs (_, T, t') => |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
942 |
(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
|
943 |
SOME t' => |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
944 |
let |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
945 |
val M = mtype_for T |
41014
e538a4f9dd86
add more flexibility to the monotonicity calculus: instead of hardcoding F-arrows, also allow G-arrows, simulating applications of the Sub rule
blanchet
parents:
41013
diff
changeset
|
946 |
val x = Unsynchronized.inc max_fresh |
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
947 |
val (M', accum) = do_term t' (accum |>> push_bound (V x) T M) |
41014
e538a4f9dd86
add more flexibility to the monotonicity calculus: instead of hardcoding F-arrows, also allow G-arrows, simulating applications of the Sub rule
blanchet
parents:
41013
diff
changeset
|
948 |
in |
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
949 |
(MFun (M, V x, M'), |
41014
e538a4f9dd86
add more flexibility to the monotonicity calculus: instead of hardcoding F-arrows, also allow G-arrows, simulating applications of the Sub rule
blanchet
parents:
41013
diff
changeset
|
950 |
accum |>> pop_bound |
e538a4f9dd86
add more flexibility to the monotonicity calculus: instead of hardcoding F-arrows, also allow G-arrows, simulating applications of the Sub rule
blanchet
parents:
41013
diff
changeset
|
951 |
||> add_annotation_atom_comp Leq [] (A Fls) (V x)) |
e538a4f9dd86
add more flexibility to the monotonicity calculus: instead of hardcoding F-arrows, also allow G-arrows, simulating applications of the Sub rule
blanchet
parents:
41013
diff
changeset
|
952 |
end |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
953 |
| NONE => |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
954 |
((case t' of |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
955 |
t1' $ Bound 0 => |
37476
0681e46b4022
optimized code generated for datatype cases + more;
blanchet
parents:
37267
diff
changeset
|
956 |
if not (loose_bvar1 (t1', 0)) andalso |
0681e46b4022
optimized code generated for datatype cases + more;
blanchet
parents:
37267
diff
changeset
|
957 |
is_enough_eta_expanded t1' then |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
958 |
do_term (incr_boundvars ~1 t1') accum |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
959 |
else |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
960 |
raise SAME () |
38864
4abe644fcea5
formerly unnamed infix equality now named HOL.eq
haftmann
parents:
38795
diff
changeset
|
961 |
| (t11 as Const (@{const_name HOL.eq}, _)) $ Bound 0 $ t13 => |
37476
0681e46b4022
optimized code generated for datatype cases + more;
blanchet
parents:
37267
diff
changeset
|
962 |
if not (loose_bvar1 (t13, 0)) then |
0681e46b4022
optimized code generated for datatype cases + more;
blanchet
parents:
37267
diff
changeset
|
963 |
do_term (incr_boundvars ~1 (t11 $ t13)) accum |
0681e46b4022
optimized code generated for datatype cases + more;
blanchet
parents:
37267
diff
changeset
|
964 |
else |
0681e46b4022
optimized code generated for datatype cases + more;
blanchet
parents:
37267
diff
changeset
|
965 |
raise SAME () |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
966 |
| _ => raise SAME ()) |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
967 |
handle SAME () => |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
968 |
let |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
969 |
val M = mtype_for T |
41014
e538a4f9dd86
add more flexibility to the monotonicity calculus: instead of hardcoding F-arrows, also allow G-arrows, simulating applications of the Sub rule
blanchet
parents:
41013
diff
changeset
|
970 |
val x = Unsynchronized.inc max_fresh |
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
971 |
val (M', accum) = |
41014
e538a4f9dd86
add more flexibility to the monotonicity calculus: instead of hardcoding F-arrows, also allow G-arrows, simulating applications of the Sub rule
blanchet
parents:
41013
diff
changeset
|
972 |
do_term t' (accum |>> push_bound (V x) T M) |
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
973 |
in (MFun (M, V x, M'), accum |>> pop_bound) end)) |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
974 |
| @{const Not} $ t1 => do_connect imp_spec t1 @{const False} accum |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
975 |
| @{const conj} $ t1 $ t2 => do_connect conj_spec t1 t2 accum |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
976 |
| @{const disj} $ t1 $ t2 => do_connect disj_spec t1 t2 accum |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
977 |
| @{const implies} $ t1 $ t2 => do_connect imp_spec t1 t2 accum |
33192 | 978 |
| Const (@{const_name Let}, _) $ t1 $ t2 => |
979 |
do_term (betapply (t2, t1)) accum |
|
980 |
| t1 $ t2 => |
|
981 |
let |
|
40998
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
982 |
fun is_in t (j, _) = loose_bvar1 (t, length bound_Ts - j - 1) |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
983 |
val accum as ({frame, ...}, _) = |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
984 |
accum |> kill_unused_in_frame (is_in t) |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
985 |
val ((frame1a, frame1b), accum) = accum |> split_frame (is_in t1) |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
986 |
val frame2a = frame1a |> map (apsnd (K (A Gen))) |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
987 |
val frame2b = |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
988 |
frame1b |> map (apsnd (fn _ => V (Unsynchronized.inc max_fresh))) |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
989 |
val frame2 = frame2a @ frame2b |
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
990 |
val (M1, accum) = accum |>> set_frame frame1a |> do_term t1 |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
991 |
val (M2, accum) = accum |>> set_frame frame2 |> do_term t2 |
33192 | 992 |
in |
35812
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
993 |
let |
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
994 |
val (M11, aa, M12) = M1 |> dest_MFun |
40998
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
995 |
in |
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
996 |
(M12, accum |>> set_frame frame |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
997 |
||> add_is_sub_mtype M2 M11 |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
998 |
||> add_app aa frame1b frame2b) |
40998
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
999 |
end |
33192 | 1000 |
end) |
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1001 |
|> tap (fn (M, (gamma, _)) => |
40998
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
1002 |
trace_msg (fn () => " " ^ string_for_mcontext ctxt t gamma ^ |
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
1003 |
" \<turnstile> " ^ |
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1004 |
Syntax.string_of_term ctxt t ^ " : " ^ |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1005 |
string_for_mtype M)) |
33192 | 1006 |
in do_term end |
1007 |
||
41001
11715564e2ad
removed old baggage from monotonicity calculus -- the "calculus" option didn't really work anyway because of on-the-fly simplifications
blanchet
parents:
41000
diff
changeset
|
1008 |
fun force_gen_funs 0 _ = I |
11715564e2ad
removed old baggage from monotonicity calculus -- the "calculus" option didn't really work anyway because of on-the-fly simplifications
blanchet
parents:
41000
diff
changeset
|
1009 |
| force_gen_funs n (M as MFun (M1, _, M2)) = |
11715564e2ad
removed old baggage from monotonicity calculus -- the "calculus" option didn't really work anyway because of on-the-fly simplifications
blanchet
parents:
41000
diff
changeset
|
1010 |
add_mtypes_equal M (MFun (M1, A Gen, M2)) #> force_gen_funs (n - 1) M2 |
11715564e2ad
removed old baggage from monotonicity calculus -- the "calculus" option didn't really work anyway because of on-the-fly simplifications
blanchet
parents:
41000
diff
changeset
|
1011 |
| force_gen_funs _ M = raise MTYPE ("Nitpick_Mono.force_gen_funs", [M], []) |
55889 | 1012 |
|
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1013 |
fun consider_general_equals mdata def t1 t2 accum = |
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
1014 |
let |
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1015 |
val (M1, accum) = consider_term mdata t1 accum |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1016 |
val (M2, accum) = consider_term mdata t2 accum |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
1017 |
val accum = accum ||> add_mtypes_equal M1 M2 |
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
1018 |
in |
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1019 |
if def then |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1020 |
let |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1021 |
val (head1, args1) = strip_comb t1 |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1022 |
val (head_M1, accum) = consider_term mdata head1 accum |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1023 |
in accum ||> force_gen_funs (length args1) head_M1 end |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1024 |
else |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1025 |
accum |
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
1026 |
end |
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
1027 |
|
41000
4bbff1684465
implemented All rules from new monotonicity calculus
blanchet
parents:
40999
diff
changeset
|
1028 |
fun consider_general_formula (mdata as {hol_ctxt = {ctxt, ...}, max_fresh, |
4bbff1684465
implemented All rules from new monotonicity calculus
blanchet
parents:
40999
diff
changeset
|
1029 |
...}) = |
33192 | 1030 |
let |
35672 | 1031 |
val mtype_for = fresh_mtype_for_type mdata false |
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1032 |
val do_term = snd oo consider_term mdata |
40998
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
1033 |
fun do_formula sn t (accum as (gamma, _)) = |
41000
4bbff1684465
implemented All rules from new monotonicity calculus
blanchet
parents:
40999
diff
changeset
|
1034 |
let |
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1035 |
fun do_quantifier quant_s abs_T body_t = |
41000
4bbff1684465
implemented All rules from new monotonicity calculus
blanchet
parents:
40999
diff
changeset
|
1036 |
let |
4bbff1684465
implemented All rules from new monotonicity calculus
blanchet
parents:
40999
diff
changeset
|
1037 |
val abs_M = mtype_for abs_T |
4bbff1684465
implemented All rules from new monotonicity calculus
blanchet
parents:
40999
diff
changeset
|
1038 |
val x = Unsynchronized.inc max_fresh |
4bbff1684465
implemented All rules from new monotonicity calculus
blanchet
parents:
40999
diff
changeset
|
1039 |
val side_cond = ((sn = Minus) = (quant_s = @{const_name Ex})) |
41001
11715564e2ad
removed old baggage from monotonicity calculus -- the "calculus" option didn't really work anyway because of on-the-fly simplifications
blanchet
parents:
41000
diff
changeset
|
1040 |
fun ann () = if quant_s = @{const_name Ex} then Fls else Tru |
41000
4bbff1684465
implemented All rules from new monotonicity calculus
blanchet
parents:
40999
diff
changeset
|
1041 |
in |
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1042 |
accum ||> side_cond |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1043 |
? add_mtype_is_complete [(x, (Plus, ann ()))] abs_M |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1044 |
|>> push_bound (V x) abs_T abs_M |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1045 |
|> do_formula sn body_t |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1046 |
|>> pop_bound |
41000
4bbff1684465
implemented All rules from new monotonicity calculus
blanchet
parents:
40999
diff
changeset
|
1047 |
end |
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1048 |
fun do_connect spec neg1 t1 t2 = |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1049 |
consider_connective mdata spec |
41003
7e2a7bd55a00
proper handling of frames for connectives in monotonicity calculus
blanchet
parents:
41002
diff
changeset
|
1050 |
(do_formula (sn |> neg1 ? negate_sign) t1) (do_formula sn t2) |
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1051 |
fun do_equals t1 t2 = |
41000
4bbff1684465
implemented All rules from new monotonicity calculus
blanchet
parents:
40999
diff
changeset
|
1052 |
case sn of |
4bbff1684465
implemented All rules from new monotonicity calculus
blanchet
parents:
40999
diff
changeset
|
1053 |
Plus => do_term t accum |
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1054 |
| Minus => consider_general_equals mdata false t1 t2 accum |
41000
4bbff1684465
implemented All rules from new monotonicity calculus
blanchet
parents:
40999
diff
changeset
|
1055 |
in |
41004 | 1056 |
trace_msg (fn () => " " ^ string_for_mcontext ctxt t gamma ^ |
1057 |
" \<turnstile> " ^ Syntax.string_of_term ctxt t ^ |
|
1058 |
" : o\<^sup>" ^ string_for_sign sn ^ "?"); |
|
1059 |
case t of |
|
56245 | 1060 |
Const (s as @{const_name Pure.all}, _) $ Abs (_, T1, t1) => |
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1061 |
do_quantifier s T1 t1 |
56245 | 1062 |
| Const (@{const_name Pure.eq}, _) $ t1 $ t2 => do_equals t1 t2 |
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1063 |
| @{const Trueprop} $ t1 => do_formula sn t1 accum |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1064 |
| Const (s as @{const_name All}, _) $ Abs (_, T1, t1) => |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1065 |
do_quantifier s T1 t1 |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1066 |
| Const (s as @{const_name Ex}, T0) $ (t1 as Abs (_, T1, t1')) => |
41004 | 1067 |
(case sn of |
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1068 |
Plus => do_quantifier s T1 t1' |
41004 | 1069 |
| Minus => |
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1070 |
(* FIXME: Needed? *) |
41004 | 1071 |
do_term (@{const Not} |
1072 |
$ (HOLogic.eq_const (domain_type T0) $ t1 |
|
1073 |
$ Abs (Name.uu, T1, @{const False}))) accum) |
|
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1074 |
| Const (@{const_name HOL.eq}, _) $ t1 $ t2 => do_equals t1 t2 |
41004 | 1075 |
| Const (@{const_name Let}, _) $ t1 $ t2 => |
1076 |
do_formula sn (betapply (t2, t1)) accum |
|
1077 |
| @{const Pure.conjunction} $ t1 $ t2 => |
|
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1078 |
do_connect meta_conj_spec false t1 t2 accum |
56245 | 1079 |
| @{const Pure.imp} $ t1 $ t2 => do_connect meta_imp_spec true t1 t2 accum |
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1080 |
| @{const Not} $ t1 => do_connect imp_spec true t1 @{const False} accum |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1081 |
| @{const conj} $ t1 $ t2 => do_connect conj_spec false t1 t2 accum |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1082 |
| @{const disj} $ t1 $ t2 => do_connect disj_spec false t1 t2 accum |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1083 |
| @{const implies} $ t1 $ t2 => do_connect imp_spec true t1 t2 accum |
41004 | 1084 |
| _ => do_term t accum |
41000
4bbff1684465
implemented All rules from new monotonicity calculus
blanchet
parents:
40999
diff
changeset
|
1085 |
end |
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1086 |
|> tap (fn (gamma, _) => |
41000
4bbff1684465
implemented All rules from new monotonicity calculus
blanchet
parents:
40999
diff
changeset
|
1087 |
trace_msg (fn () => string_for_mcontext ctxt t gamma ^ |
4bbff1684465
implemented All rules from new monotonicity calculus
blanchet
parents:
40999
diff
changeset
|
1088 |
" \<turnstile> " ^ |
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1089 |
Syntax.string_of_term ctxt t ^ |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1090 |
" : o\<^sup>" ^ string_for_sign sn)) |
33192 | 1091 |
in do_formula end |
1092 |
||
1093 |
(* The harmless axiom optimization below is somewhat too aggressive in the face |
|
1094 |
of (rather peculiar) user-defined axioms. *) |
|
1095 |
val harmless_consts = |
|
1096 |
[@{const_name ord_class.less}, @{const_name ord_class.less_eq}] |
|
1097 |
val bounteous_consts = [@{const_name bisim}] |
|
1098 |
||
55888 | 1099 |
fun is_harmless_axiom t = |
1100 |
Term.add_consts t [] |
|
1101 |
|> filter_out is_built_in_const |
|
1102 |
|> (forall (member (op =) harmless_consts o original_name o fst) orf |
|
1103 |
exists (member (op =) bounteous_consts o fst)) |
|
33192 | 1104 |
|
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
1105 |
fun consider_nondefinitional_axiom mdata t = |
55888 | 1106 |
if is_harmless_axiom t then I |
35386
45a4e19d3ebd
more work on the new monotonicity stuff in Nitpick
blanchet
parents:
35385
diff
changeset
|
1107 |
else consider_general_formula mdata Plus t |
33192 | 1108 |
|
41109 | 1109 |
fun consider_definitional_axiom (mdata as {hol_ctxt = {ctxt, ...}, ...} : mdata) t = |
37256
0dca1ec52999
thread along context instead of theory for typedef lookup
blanchet
parents:
36385
diff
changeset
|
1110 |
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
|
1111 |
consider_nondefinitional_axiom mdata t |
55888 | 1112 |
else if is_harmless_axiom t then |
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1113 |
I |
33192 | 1114 |
else |
1115 |
let |
|
35672 | 1116 |
val mtype_for = fresh_mtype_for_type mdata false |
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1117 |
val do_term = snd oo consider_term mdata |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1118 |
fun do_all abs_T body_t accum = |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1119 |
accum |>> push_bound (A Gen) abs_T (mtype_for abs_T) |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1120 |
|> do_formula body_t |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1121 |
|>> pop_bound |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1122 |
and do_implies t1 t2 = do_term t1 #> do_formula t2 |
35812
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1123 |
and do_formula t accum = |
41002 | 1124 |
case t of |
56245 | 1125 |
Const (@{const_name Pure.all}, _) $ Abs (_, T1, t1) => do_all T1 t1 accum |
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1126 |
| @{const Trueprop} $ t1 => do_formula t1 accum |
56245 | 1127 |
| Const (@{const_name Pure.eq}, _) $ t1 $ t2 => |
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1128 |
consider_general_equals mdata true t1 t2 accum |
56245 | 1129 |
| @{const Pure.imp} $ t1 $ t2 => do_implies t1 t2 accum |
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1130 |
| @{const Pure.conjunction} $ t1 $ t2 => |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1131 |
fold (do_formula) [t1, t2] accum |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1132 |
| Const (@{const_name All}, _) $ Abs (_, T1, t1) => do_all T1 t1 accum |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1133 |
| Const (@{const_name HOL.eq}, _) $ t1 $ t2 => |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1134 |
consider_general_equals mdata true t1 t2 accum |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1135 |
| @{const conj} $ t1 $ t2 => fold (do_formula) [t1, t2] accum |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1136 |
| @{const implies} $ t1 $ t2 => do_implies t1 t2 accum |
41002 | 1137 |
| _ => raise TERM ("Nitpick_Mono.consider_definitional_axiom.\ |
1138 |
\do_formula", [t]) |
|
33192 | 1139 |
in do_formula t end |
1140 |
||
40989 | 1141 |
fun string_for_mtype_of_term ctxt asgs t M = |
1142 |
Syntax.string_of_term ctxt t ^ " : " ^ string_for_mtype (resolve_mtype asgs M) |
|
33192 | 1143 |
|
40998
bcd23ddeecef
started implementing the new monotonicity rules for application
blanchet
parents:
40997
diff
changeset
|
1144 |
fun print_mcontext ctxt asgs ({frees, consts, ...} : mcontext) = |
38179 | 1145 |
trace_msg (fn () => |
40989 | 1146 |
map (fn (x, M) => string_for_mtype_of_term ctxt asgs (Free x) M) frees @ |
1147 |
map (fn (x, M) => string_for_mtype_of_term ctxt asgs (Const x) M) consts |
|
37476
0681e46b4022
optimized code generated for datatype cases + more;
blanchet
parents:
37267
diff
changeset
|
1148 |
|> cat_lines) |
33192 | 1149 |
|
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1150 |
fun formulas_monotonic (hol_ctxt as {ctxt, tac_timeout, ...}) binarize alpha_T |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1151 |
(nondef_ts, def_ts) = |
33192 | 1152 |
let |
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1153 |
val _ = trace_msg (fn () => "****** Monotonicity analysis: " ^ |
38179 | 1154 |
string_for_mtype MAlpha ^ " is " ^ |
1155 |
Syntax.string_of_typ ctxt alpha_T) |
|
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1156 |
val mdata as {max_fresh, ...} = initial_mdata hol_ctxt binarize alpha_T |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1157 |
val (gamma, cset) = |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1158 |
(initial_gamma, ([], [])) |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1159 |
|> consider_general_formula mdata Plus (hd nondef_ts) |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1160 |
|> fold (consider_nondefinitional_axiom mdata) (tl nondef_ts) |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1161 |
|> fold (consider_definitional_axiom mdata) def_ts |
33192 | 1162 |
in |
41007
75275c796b46
use ML SAT solvers up to a certain time limit, then switch to faster solvers with a timeout -- this becomes necessary with the new, more powerful monotonicity calculus
blanchet
parents:
41005
diff
changeset
|
1163 |
case solve tac_timeout (!max_fresh) cset of |
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1164 |
SOME asgs => (print_mcontext ctxt asgs gamma; true) |
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1165 |
| _ => false |
33192 | 1166 |
end |
41054
e58d1cdda832
simplify monotonicity code after killing "fin_fun" optimization
blanchet
parents:
41052
diff
changeset
|
1167 |
handle UNSOLVABLE () => false |
35812
394fe2b064cd
solve error in "Nitpick_Mono" + short path when no finite functions are inferred
blanchet
parents:
35672
diff
changeset
|
1168 |
| MTYPE (loc, Ms, Ts) => |
35665
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
1169 |
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
|
1170 |
map (Syntax.string_of_typ ctxt) Ts)) |
ff2bf50505ab
added "finitize" option to Nitpick + remove dependency on "Coinductive_List"
blanchet
parents:
35386
diff
changeset
|
1171 |
|
33192 | 1172 |
end; |