| author | haftmann | 
| Tue, 24 Nov 2009 17:28:25 +0100 | |
| changeset 33955 | fff6f11b1f09 | 
| parent 33705 | 947184dc75c9 | 
| child 33982 | 1ae222745c4a | 
| permissions | -rw-r--r-- | 
| 33192 | 1 | (* Title: HOL/Nitpick/Tools/nitpick_util.ML | 
| 2 | Author: Jasmin Blanchette, TU Muenchen | |
| 3 | Copyright 2008, 2009 | |
| 4 | ||
| 5 | General-purpose functions used by the Nitpick modules. | |
| 6 | *) | |
| 7 | ||
| 33705 
947184dc75c9
removed a few global names in Nitpick (styp, nat_less, pairf)
 blanchet parents: 
33232diff
changeset | 8 | signature NITPICK_UTIL = | 
| 33192 | 9 | sig | 
| 10 | type styp = string * typ | |
| 11 | datatype polarity = Pos | Neg | Neut | |
| 12 | ||
| 13 | exception ARG of string * string | |
| 14 | exception BAD of string * string | |
| 15 | exception LIMIT of string * string | |
| 16 | exception NOT_SUPPORTED of string | |
| 17 | exception SAME of unit | |
| 18 | ||
| 19 | val nitpick_prefix : string | |
| 20 |   val curry3 : ('a * 'b * 'c -> 'd) -> 'a -> 'b -> 'c -> 'd
 | |
| 33705 
947184dc75c9
removed a few global names in Nitpick (styp, nat_less, pairf)
 blanchet parents: 
33232diff
changeset | 21 |   val pairf : ('a -> 'b) -> ('a -> 'c) -> 'a -> 'b * 'c
 | 
| 33192 | 22 | val int_for_bool : bool -> int | 
| 33705 
947184dc75c9
removed a few global names in Nitpick (styp, nat_less, pairf)
 blanchet parents: 
33232diff
changeset | 23 | val nat_minus : int -> int -> int | 
| 33192 | 24 | val reasonable_power : int -> int -> int | 
| 25 | val exact_log : int -> int -> int | |
| 26 | val exact_root : int -> int -> int | |
| 27 | val offset_list : int list -> int list | |
| 28 | val index_seq : int -> int -> int list | |
| 29 | val filter_indices : int list -> 'a list -> 'a list | |
| 30 | val filter_out_indices : int list -> 'a list -> 'a list | |
| 31 |   val fold1 : ('a -> 'a -> 'a) -> 'a list -> 'a
 | |
| 32 | val replicate_list : int -> 'a list -> 'a list | |
| 33 | val n_fold_cartesian_product : 'a list list -> 'a list list | |
| 34 |   val all_distinct_unordered_pairs_of : ''a list -> (''a * ''a) list
 | |
| 35 | val nth_combination : (int * int) list -> int -> int list | |
| 36 | val all_combinations : (int * int) list -> int list list | |
| 37 | val all_permutations : 'a list -> 'a list list | |
| 38 | val batch_list : int -> 'a list -> 'a list list | |
| 39 | val chunk_list_unevenly : int list -> 'a list -> 'a list list | |
| 40 |   val map3 : ('a -> 'b -> 'c -> 'd) -> 'a list -> 'b list -> 'c list -> 'd list
 | |
| 41 | val double_lookup : | |
| 42 |     ('a * 'a -> bool) -> ('a option * 'b) list -> 'a -> 'b option
 | |
| 43 | val triple_lookup : | |
| 44 |     (''a * ''a -> bool) -> (''a option * 'b) list -> ''a -> 'b option
 | |
| 45 | val is_substring_of : string -> string -> bool | |
| 46 | val serial_commas : string -> string list -> string list | |
| 47 | val plural_s : int -> string | |
| 48 | val plural_s_for_list : 'a list -> string | |
| 49 | val flip_polarity : polarity -> polarity | |
| 50 | val prop_T : typ | |
| 51 | val bool_T : typ | |
| 52 | val nat_T : typ | |
| 53 | val int_T : typ | |
| 54 | val subscript : string -> string | |
| 55 | val nat_subscript : int -> string | |
| 56 |   val time_limit : Time.time option -> ('a -> 'b) -> 'a -> 'b
 | |
| 57 | val DETERM_TIMEOUT : Time.time option -> tactic -> tactic | |
| 58 |   val setmp_show_all_types : ('a -> 'b) -> 'a -> 'b
 | |
| 59 | val indent_size : int | |
| 60 | val pstrs : string -> Pretty.T list | |
| 61 | val plain_string_from_yxml : string -> string | |
| 62 | val maybe_quote : string -> string | |
| 63 | end | |
| 64 | ||
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 65 | structure Nitpick_Util : NITPICK_UTIL = | 
| 33192 | 66 | struct | 
| 67 | ||
| 68 | type styp = string * typ | |
| 69 | ||
| 70 | datatype polarity = Pos | Neg | Neut | |
| 71 | ||
| 72 | exception ARG of string * string | |
| 73 | exception BAD of string * string | |
| 74 | exception LIMIT of string * string | |
| 75 | exception NOT_SUPPORTED of string | |
| 76 | exception SAME of unit | |
| 77 | ||
| 78 | val nitpick_prefix = "Nitpick." | |
| 79 | ||
| 80 | (* ('a * 'b * 'c -> 'd) -> 'a -> 'b -> 'c -> 'd *)
 | |
| 81 | fun curry3 f = fn x => fn y => fn z => f (x, y, z) | |
| 82 | ||
| 33705 
947184dc75c9
removed a few global names in Nitpick (styp, nat_less, pairf)
 blanchet parents: 
33232diff
changeset | 83 | (* ('a -> 'b) -> ('a -> 'c) -> 'a -> 'b * 'c *)
 | 
| 
947184dc75c9
removed a few global names in Nitpick (styp, nat_less, pairf)
 blanchet parents: 
33232diff
changeset | 84 | fun pairf f g x = (f x, g x) | 
| 33192 | 85 | |
| 86 | (* bool -> int *) | |
| 87 | fun int_for_bool b = if b then 1 else 0 | |
| 33705 
947184dc75c9
removed a few global names in Nitpick (styp, nat_less, pairf)
 blanchet parents: 
33232diff
changeset | 88 | (* int -> int -> int *) | 
| 
947184dc75c9
removed a few global names in Nitpick (styp, nat_less, pairf)
 blanchet parents: 
33232diff
changeset | 89 | fun nat_minus i j = if i > j then i - j else 0 | 
| 33192 | 90 | |
| 91 | val max_exponent = 16384 | |
| 92 | ||
| 93 | (* int -> int -> int *) | |
| 94 | fun reasonable_power a 0 = 1 | |
| 95 | | reasonable_power a 1 = a | |
| 96 | | reasonable_power 0 _ = 0 | |
| 97 | | reasonable_power 1 _ = 1 | |
| 98 | | reasonable_power a b = | |
| 99 | if b < 0 orelse b > max_exponent then | |
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 100 |       raise LIMIT ("Nitpick_Util.reasonable_power",
 | 
| 33192 | 101 |                    "too large exponent (" ^ signed_string_of_int b ^ ")")
 | 
| 102 | else | |
| 103 | let | |
| 104 | val c = reasonable_power a (b div 2) in | |
| 105 | c * c * reasonable_power a (b mod 2) | |
| 106 | end | |
| 107 | ||
| 108 | (* int -> int -> int *) | |
| 109 | fun exact_log m n = | |
| 110 | let | |
| 111 | val r = Math.ln (Real.fromInt n) / Math.ln (Real.fromInt m) |> Real.round | |
| 112 | in | |
| 113 | if reasonable_power m r = n then | |
| 114 | r | |
| 115 | else | |
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 116 |       raise ARG ("Nitpick_Util.exact_log",
 | 
| 33192 | 117 | commas (map signed_string_of_int [m, n])) | 
| 118 | end | |
| 119 | ||
| 120 | (* int -> int -> int *) | |
| 121 | fun exact_root m n = | |
| 122 | let val r = Math.pow (Real.fromInt n, 1.0 / (Real.fromInt m)) |> Real.round in | |
| 123 | if reasonable_power r m = n then | |
| 124 | r | |
| 125 | else | |
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 126 |       raise ARG ("Nitpick_Util.exact_root",
 | 
| 33192 | 127 | commas (map signed_string_of_int [m, n])) | 
| 128 | end | |
| 129 | ||
| 130 | (* ('a -> 'a -> 'a) -> 'a list -> 'a *)
 | |
| 131 | fun fold1 f = foldl1 (uncurry f) | |
| 132 | ||
| 133 | (* int -> 'a list -> 'a list *) | |
| 134 | fun replicate_list 0 _ = [] | |
| 135 | | replicate_list n xs = xs @ replicate_list (n - 1) xs | |
| 136 | ||
| 137 | (* int list -> int list *) | |
| 138 | fun offset_list ns = rev (tl (fold (fn x => fn xs => (x + hd xs) :: xs) ns [0])) | |
| 139 | (* int -> int -> int list *) | |
| 140 | fun index_seq j0 n = if j0 < 0 then j0 downto j0 - n + 1 else j0 upto j0 + n - 1 | |
| 141 | ||
| 142 | (* int list -> 'a list -> 'a list *) | |
| 143 | fun filter_indices js xs = | |
| 144 | let | |
| 145 | (* int -> int list -> 'a list -> 'a list *) | |
| 146 | fun aux _ [] _ = [] | |
| 147 | | aux i (j :: js) (x :: xs) = | |
| 148 | if i = j then x :: aux (i + 1) js xs else aux (i + 1) (j :: js) xs | |
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 149 |       | aux _ _ _ = raise ARG ("Nitpick_Util.filter_indices",
 | 
| 33192 | 150 | "indices unordered or out of range") | 
| 151 | in aux 0 js xs end | |
| 152 | fun filter_out_indices js xs = | |
| 153 | let | |
| 154 | (* int -> int list -> 'a list -> 'a list *) | |
| 155 | fun aux _ [] xs = xs | |
| 156 | | aux i (j :: js) (x :: xs) = | |
| 157 | if i = j then aux (i + 1) js xs else x :: aux (i + 1) (j :: js) xs | |
| 33232 
f93390060bbe
internal renaming in Nitpick and fixed Kodkodi invokation on Linux;
 blanchet parents: 
33192diff
changeset | 158 |       | aux _ _ _ = raise ARG ("Nitpick_Util.filter_out_indices",
 | 
| 33192 | 159 | "indices unordered or out of range") | 
| 160 | in aux 0 js xs end | |
| 161 | ||
| 162 | (* 'a list -> 'a list list -> 'a list list *) | |
| 163 | fun cartesian_product [] _ = [] | |
| 164 | | cartesian_product (x :: xs) yss = | |
| 165 | map (cons x) yss @ cartesian_product xs yss | |
| 166 | (* 'a list list -> 'a list list *) | |
| 167 | fun n_fold_cartesian_product xss = fold_rev cartesian_product xss [[]] | |
| 168 | (* ''a list -> (''a * ''a) list *)
 | |
| 169 | fun all_distinct_unordered_pairs_of [] = [] | |
| 170 | | all_distinct_unordered_pairs_of (x :: xs) = | |
| 171 | map (pair x) xs @ all_distinct_unordered_pairs_of xs | |
| 172 | ||
| 173 | (* (int * int) list -> int -> int list *) | |
| 174 | val nth_combination = | |
| 175 | let | |
| 176 | (* (int * int) list -> int -> int list * int *) | |
| 177 | fun aux [] n = ([], n) | |
| 178 | | aux ((k, j0) :: xs) n = | |
| 179 | let val (js, n) = aux xs n in ((n mod k) + j0 :: js, n div k) end | |
| 180 | in fst oo aux end | |
| 181 | ||
| 182 | (* (int * int) list -> int list list *) | |
| 183 | val all_combinations = n_fold_cartesian_product o map (uncurry index_seq o swap) | |
| 184 | ||
| 185 | (* 'a list -> 'a list list *) | |
| 186 | fun all_permutations [] = [[]] | |
| 187 | | all_permutations xs = | |
| 188 | maps (fn j => map (cons (nth xs j)) (all_permutations (nth_drop j xs))) | |
| 189 | (index_seq 0 (length xs)) | |
| 190 | ||
| 191 | (* int -> 'a list -> 'a list list *) | |
| 192 | fun batch_list _ [] = [] | |
| 193 | | batch_list k xs = | |
| 194 | if length xs <= k then [xs] | |
| 195 | else List.take (xs, k) :: batch_list k (List.drop (xs, k)) | |
| 196 | ||
| 197 | (* int list -> 'a list -> 'a list list *) | |
| 198 | fun chunk_list_unevenly _ [] = [] | |
| 199 | | chunk_list_unevenly [] ys = map single ys | |
| 200 | | chunk_list_unevenly (k :: ks) ys = | |
| 201 | let val (ys1, ys2) = chop k ys in ys1 :: chunk_list_unevenly ks ys2 end | |
| 202 | ||
| 203 | (* ('a -> 'b -> 'c -> 'd) -> 'a list -> 'b list -> 'c list -> 'd list *)
 | |
| 204 | fun map3 _ [] [] [] = [] | |
| 205 | | map3 f (x :: xs) (y :: ys) (z :: zs) = f x y z :: map3 f xs ys zs | |
| 206 | | map3 _ _ _ _ = raise UnequalLengths | |
| 207 | ||
| 208 | (* ('a * 'a -> bool) -> ('a option * 'b) list -> 'a -> 'b option *)
 | |
| 209 | fun double_lookup eq ps key = | |
| 210 | case AList.lookup (fn (SOME x, SOME y) => eq (x, y) | _ => false) ps | |
| 211 | (SOME key) of | |
| 212 | SOME z => SOME z | |
| 213 | | NONE => ps |> find_first (is_none o fst) |> Option.map snd | |
| 214 | (* (''a * ''a -> bool) -> (''a option * 'b) list -> ''a -> 'b option *)
 | |
| 215 | fun triple_lookup eq ps key = | |
| 216 | case AList.lookup (op =) ps (SOME key) of | |
| 217 | SOME z => SOME z | |
| 218 | | NONE => double_lookup eq ps key | |
| 219 | ||
| 220 | (* string -> string -> bool *) | |
| 221 | fun is_substring_of needle stack = | |
| 222 | not (Substring.isEmpty (snd (Substring.position needle | |
| 223 | (Substring.full stack)))) | |
| 224 | ||
| 225 | (* string -> string list -> string list *) | |
| 226 | fun serial_commas _ [] = ["??"] | |
| 227 | | serial_commas _ [s] = [s] | |
| 228 | | serial_commas conj [s1, s2] = [s1, conj, s2] | |
| 229 | | serial_commas conj [s1, s2, s3] = [s1 ^ ",", s2 ^ ",", conj, s3] | |
| 230 | | serial_commas conj (s :: ss) = s ^ "," :: serial_commas conj ss | |
| 231 | ||
| 232 | (* int -> string *) | |
| 233 | fun plural_s n = if n = 1 then "" else "s" | |
| 234 | (* 'a list -> string *) | |
| 235 | fun plural_s_for_list xs = plural_s (length xs) | |
| 236 | ||
| 237 | (* polarity -> polarity *) | |
| 238 | fun flip_polarity Pos = Neg | |
| 239 | | flip_polarity Neg = Pos | |
| 240 | | flip_polarity Neut = Neut | |
| 241 | ||
| 242 | val prop_T = @{typ prop}
 | |
| 243 | val bool_T = @{typ bool}
 | |
| 244 | val nat_T = @{typ nat}
 | |
| 245 | val int_T = @{typ int}
 | |
| 246 | ||
| 247 | (* string -> string *) | |
| 248 | val subscript = implode o map (prefix "\<^isub>") o explode | |
| 249 | (* int -> string *) | |
| 250 | val nat_subscript = subscript o signed_string_of_int | |
| 251 | ||
| 252 | (* Time.time option -> ('a -> 'b) -> 'a -> 'b *)
 | |
| 253 | fun time_limit NONE f = f | |
| 254 | | time_limit (SOME delay) f = TimeLimit.timeLimit delay f | |
| 255 | ||
| 256 | (* Time.time option -> tactic -> tactic *) | |
| 257 | fun DETERM_TIMEOUT delay tac st = | |
| 258 | Seq.of_list (the_list (time_limit delay (fn () => SINGLE tac st) ())) | |
| 259 | ||
| 260 | (* ('a -> 'b) -> 'a -> 'b *)
 | |
| 261 | fun setmp_show_all_types f = | |
| 262 | setmp_CRITICAL show_all_types | |
| 263 | (! show_types orelse ! show_sorts orelse ! show_all_types) f | |
| 264 | ||
| 265 | val indent_size = 2 | |
| 266 | ||
| 267 | (* string -> Pretty.T list *) | |
| 268 | val pstrs = Pretty.breaks o map Pretty.str o space_explode " " | |
| 269 | ||
| 270 | (* XML.tree -> string *) | |
| 271 | fun plain_string_from_xml_tree t = | |
| 272 | Buffer.empty |> XML.add_content t |> Buffer.content | |
| 273 | (* string -> string *) | |
| 274 | val plain_string_from_yxml = plain_string_from_xml_tree o YXML.parse | |
| 275 | ||
| 276 | (* string -> bool *) | |
| 277 | val is_long_identifier = forall Syntax.is_identifier o space_explode "." | |
| 278 | (* string -> string *) | |
| 279 | fun maybe_quote y = | |
| 280 | let val s = plain_string_from_yxml y in | |
| 281 | y |> (not (is_long_identifier (perhaps (try (unprefix "'")) s)) | |
| 282 | orelse OuterKeyword.is_keyword s) ? quote | |
| 283 | end | |
| 284 | ||
| 285 | end; |