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