| author | huffman | 
| Wed, 12 Oct 2005 03:02:18 +0200 | |
| changeset 17840 | 11bcd77cfa22 | 
| parent 17796 | 86daafee72d6 | 
| child 19475 | 8aa2b380614a | 
| permissions | -rw-r--r-- | 
| 15481 | 1 | (* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) | 
| 16179 | 2 | (* Title: Pure/IsaPlanner/isaplib.ML | 
| 3 | ID: $Id$ | |
| 15481 | 4 | Author: Lucas Dixon, University of Edinburgh | 
| 5 | lucasd@dai.ed.ac.uk | |
| 6 | *) | |
| 7 | (* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) | |
| 8 | (* DESCRIPTION: | |
| 9 | ||
| 10 | A few useful system-independent utilities. | |
| 11 | *) | |
| 12 | (* -=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=-=- *) | |
| 13 | signature ISAP_LIB = | |
| 14 | sig | |
| 15 | (* seq operations *) | |
| 16 |   val ALL_BUT_LAST : ('a -> 'b Seq.seq) -> 'a -> 'b Seq.seq
 | |
| 17 |   val FST : ('a -> 'b Seq.seq) -> 'a -> 'b Seq.seq
 | |
| 18 |   val NTH : int -> ('a -> 'b Seq.seq) -> 'a -> 'b Seq.seq
 | |
| 19 | val all_but_last_of_seq : 'a Seq.seq -> 'a Seq.seq | |
| 20 | val nat_seq : int Seq.seq | |
| 15531 | 21 | val nth_of_seq : int -> 'a Seq.seq -> 'a option | 
| 15481 | 22 |   val pair_seq : 'a Seq.seq -> 'b Seq.seq -> ('a * 'b) Seq.seq
 | 
| 23 | val seq_is_empty : 'a Seq.seq -> bool | |
| 24 | val number_seq : 'a Seq.seq -> (int * 'a) Seq.seq | |
| 25 | ||
| 16004 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 26 | datatype 'a skipseqT = | 
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 27 | skipmore of int | 
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 28 | | skipseq of 'a Seq.seq Seq.seq; | 
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 29 | val skipto_seqseq : int -> 'a Seq.seq Seq.seq -> 'a skipseqT | 
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 30 | val seqflat_seq : 'a Seq.seq Seq.seq -> 'a Seq.seq | 
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 31 | |
| 15481 | 32 | (* lists *) | 
| 33 | val mk_num_list : int -> int list | |
| 34 | val number_list : int -> 'a list -> (int * 'a) list | |
| 35 | val repeated_list : int -> 'a -> 'a list | |
| 36 |   val all_pairs : 'a list -> 'b list -> ('a * 'b) list
 | |
| 37 | 	val get_ends_of : ('a * 'a -> bool) ->
 | |
| 38 | 										('a * 'a) -> 'a list -> ('a * 'a)
 | |
| 39 |   val flatmap : ('a -> 'b list) -> 'a list -> 'b list
 | |
| 40 | ||
| 41 | 	val lrem : ('a * 'b -> bool) -> 'a list -> 'b list -> 'b list
 | |
| 42 |   val forall_list : ('a -> bool) -> 'a list -> bool
 | |
| 43 | ||
| 44 | (* the list of lists with one of each of the original sublists *) | |
| 45 | val one_of_each : 'a list list -> 'a list list | |
| 46 | ||
| 47 | (* type changing *) | |
| 48 | exception NOT_A_DIGIT of string | |
| 49 | val int_of_string : string -> int | |
| 50 | ||
| 51 | (* string manipulations *) | |
| 52 | val remove_end_spaces : string -> string | |
| 53 | val str_indent : string -> string | |
| 54 | val string_of_intlist : int list -> string | |
| 55 |   val string_of_list : ('a -> string) -> 'a list -> string
 | |
| 56 | ||
| 57 | (* options *) | |
| 15531 | 58 |   val aptosome : 'a option -> ('a -> 'b) -> 'b option
 | 
| 59 |   val seq_mapfilter : ('a -> 'b option) -> 'a Seq.seq -> 'b Seq.seq
 | |
| 60 |   val seq_map_to_some_filter : ('a -> 'b) -> 'a option Seq.seq 
 | |
| 15481 | 61 | -> 'b Seq.seq | 
| 62 | end; | |
| 63 | ||
| 64 | ||
| 65 | ||
| 17796 | 66 | structure IsaPLib : ISAP_LIB = | 
| 15481 | 67 | struct | 
| 68 | ||
| 69 | (* Seq *) | |
| 70 | fun seq_map_to_some_filter f s0 = | |
| 71 | let | |
| 72 | fun recf s () = | |
| 73 | case (Seq.pull s) of | |
| 15531 | 74 | NONE => NONE | 
| 75 | | SOME (NONE,s') => recf s' () | |
| 76 | | SOME (SOME d, s') => | |
| 77 | SOME (f d, Seq.make (recf s')) | |
| 15481 | 78 | in Seq.make (recf s0) end; | 
| 79 | ||
| 80 | fun seq_mapfilter f s0 = | |
| 81 | let | |
| 82 | fun recf s () = | |
| 83 | case (Seq.pull s) of | |
| 15531 | 84 | NONE => NONE | 
| 85 | | SOME (a,s') => | |
| 86 | (case f a of NONE => recf s' () | |
| 87 | | SOME b => SOME (b, Seq.make (recf s'))) | |
| 15481 | 88 | in Seq.make (recf s0) end; | 
| 89 | ||
| 90 | ||
| 91 | ||
| 92 | (* a simple function to pair with each element of a list, a number *) | |
| 93 | fun number_list i [] = [] | |
| 94 | | number_list i (h::t) = | |
| 95 | (i,h)::(number_list (i+1) t) | |
| 96 | ||
| 97 | (* check to see if a sequence is empty *) | |
| 98 | fun seq_is_empty s = is_none (Seq.pull s); | |
| 99 | ||
| 100 | (* the sequence of natural numbers *) | |
| 101 | val nat_seq = | |
| 15531 | 102 | let fun nseq i () = SOME (i, Seq.make (nseq (i+1))) | 
| 15481 | 103 | in Seq.make (nseq 1) | 
| 104 | end; | |
| 105 | ||
| 106 | (* create a sequence of pairs of the elements of the two sequences | |
| 107 | If one sequence becomes empty, then so does the pairing of them. | |
| 108 | ||
| 109 | This is particularly useful if you wish to perform counting or | |
| 110 | other repeated operations on a sequence and you want to combvine | |
| 111 | this infinite sequence with a possibly finite one. | |
| 112 | ||
| 113 | behaviour: | |
| 114 | s1: a1, a2, ... an | |
| 115 | s2: b1, b2, ... bn ... | |
| 116 | ||
| 117 | pair_seq s1 s2: (a1,b1), (a2,b2), ... (an,bn) | |
| 118 | *) | |
| 119 | fun pair_seq seq1 seq2 = | |
| 120 | let | |
| 121 | fun pseq s1 s2 () = | |
| 122 | case Seq.pull s1 of | |
| 15531 | 123 | NONE => NONE | 
| 124 | | SOME (s1h, s1t) => | |
| 15481 | 125 | case Seq.pull s2 of | 
| 15531 | 126 | NONE => NONE | 
| 127 | | SOME (s2h, s2t) => | |
| 128 | SOME ((s1h, s2h), Seq.make (pseq s1t s2t)) | |
| 15481 | 129 | in | 
| 130 | Seq.make (pseq seq1 seq2) | |
| 131 | end; | |
| 132 | ||
| 133 | (* number a sequence *) | |
| 134 | fun number_seq s = pair_seq nat_seq s; | |
| 135 | ||
| 136 | (* cuts off the last element of a sequence *) | |
| 137 | fun all_but_last_of_seq s = | |
| 138 | let | |
| 139 | fun f () = | |
| 140 | case Seq.pull s of | |
| 15531 | 141 | NONE => NONE | 
| 142 | | SOME (a, s2) => | |
| 15481 | 143 | (case Seq.pull s2 of | 
| 15531 | 144 | NONE => NONE | 
| 145 | | SOME (a2,s3) => | |
| 146 | SOME (a, all_but_last_of_seq (Seq.cons (a2,s3)))) | |
| 15481 | 147 | in | 
| 148 | Seq.make f | |
| 149 | end; | |
| 150 | ||
| 151 | fun ALL_BUT_LAST r st = all_but_last_of_seq (r st); | |
| 152 | ||
| 153 | ||
| 154 | (* nth elem for sequenes, return none if out of bounds *) | |
| 155 | fun nth_of_seq i l = | |
| 15531 | 156 | if (seq_is_empty l) then NONE | 
| 157 | else if i <= 1 then SOME (Seq.hd l) | |
| 15481 | 158 | else nth_of_seq (i-1) (Seq.tl l); | 
| 159 | ||
| 160 | (* for use with tactics... gives no_tac if element isn't there *) | |
| 161 | fun NTH n f st = | |
| 162 | let val r = nth_of_seq n (f st) in | |
| 15570 | 163 | if (is_none r) then Seq.empty else (Seq.single (valOf r)) | 
| 15481 | 164 | end; | 
| 165 | ||
| 166 | (* First result of a tactic... uses NTH, so if there is no elements, | |
| 167 | ||
| 168 | then no_tac is returned. *) | |
| 169 | fun FST f = NTH 1 f; | |
| 170 | ||
| 16004 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 171 | datatype 'a skipseqT = skipmore of int | 
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 172 | | skipseq of 'a Seq.seq Seq.seq; | 
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 173 | (* given a seqseq, skip the first m non-empty seq's *) | 
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 174 | fun skipto_seqseq m s = | 
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 175 | let | 
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 176 | fun skip_occs n sq = | 
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 177 | case Seq.pull sq of | 
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 178 | NONE => skipmore n | 
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 179 | | SOME (h,t) => | 
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 180 | (case Seq.pull h of NONE => skip_occs n t | 
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 181 | | SOME _ => if n <= 1 then skipseq (Seq.cons (h, t)) | 
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 182 | else skip_occs (n - 1) t) | 
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 183 | in (skip_occs m s) end; | 
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 184 | |
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 185 | (* handy function for re-arranging Sequence operations *) | 
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 186 | (* [[a,b,c],[d,e],...] => flatten [[a,d,...], [b,e,...], [c, ...]] *) | 
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 187 | fun seqflat_seq ss = | 
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 188 | let | 
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 189 | fun pulltl LL () = | 
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 190 | (case Seq.pull LL of | 
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 191 | NONE => NONE | 
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 192 | | SOME (hL,tLL) => | 
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 193 | (case Seq.pull hL of | 
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 194 | NONE => pulltl tLL () | 
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 195 | | SOME (h,tL) => | 
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 196 | SOME (h, Seq.make (recf (tLL, (Seq.single tL)))))) | 
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 197 | and recf (fstLL,sndLL) () = | 
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 198 | (case Seq.pull fstLL of | 
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 199 | NONE => pulltl sndLL () | 
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 200 | | SOME (hL, tLL) => | 
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 201 | (case Seq.pull hL of | 
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 202 | NONE => recf (tLL, sndLL) () | 
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 203 | | SOME (h,tL) => | 
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 204 | SOME (h, Seq.make (recf (tLL, Seq.append (sndLL, Seq.single tL)))))) | 
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 205 | in Seq.make (pulltl ss) end; | 
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 206 | (* tested with: | 
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 207 | val ss = Seq.of_list [Seq.of_list [1,2,3], Seq.of_list [4,5], Seq.of_list [7,8,9,10]]; | 
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 208 | Seq.list_of (seqflat_seq ss); | 
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 209 | val it = [1, 4, 7, 2, 5, 8, 3, 9, 10] : int list | 
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 210 | *) | 
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 211 | |
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 212 | |
| 
031f56012483
lucas - fixed subst in assumptions to count redexes from left to right.
 dixon parents: 
15570diff
changeset | 213 | |
| 15481 | 214 | (* create a list opf the form (n, n-1, n-2, ... ) *) | 
| 215 | fun mk_num_list n = | |
| 216 | if n < 1 then [] else (n :: (mk_num_list (n-1))); | |
| 217 | ||
| 218 | fun repeated_list n a = | |
| 219 | if n < 1 then [] else (a :: (repeated_list (n-1) a)); | |
| 220 | ||
| 221 | (* create all possible pairs with fst element from the first list | |
| 222 | and snd element from teh second list *) | |
| 223 | fun all_pairs xs ys = | |
| 224 | let | |
| 225 | fun all_pairs_aux yss [] _ acc = acc | |
| 226 | | all_pairs_aux yss (x::xs) [] acc = | |
| 227 | all_pairs_aux yss xs yss acc | |
| 228 | | all_pairs_aux yss (xs as (x1::x1s)) (y::ys) acc = | |
| 229 | all_pairs_aux yss xs ys ((x1,y)::acc) | |
| 230 | in | |
| 231 | all_pairs_aux ys xs ys [] | |
| 232 | end; | |
| 233 | ||
| 234 | ||
| 235 | (* create all possible pairs with fst element from the first list | |
| 236 | and snd element from teh second list *) | |
| 237 | (* FIXME: make tail recursive and quick *) | |
| 238 | fun one_of_each [] = [] | |
| 239 | | one_of_each [[]] = [] | |
| 240 | | one_of_each [(h::t)] = [h] :: (one_of_each [t]) | |
| 241 | | one_of_each ([] :: xs) = [] | |
| 242 | | one_of_each ((h :: t) :: xs) = | |
| 243 | (map (fn z => h :: z) (one_of_each xs)) | |
| 244 | @ (one_of_each (t :: xs)); | |
| 245 | ||
| 246 | ||
| 247 | (* function to get the upper/lower bounds of a list | |
| 248 | given: | |
| 249 | gr : 'a * 'a -> bool = greater than check | |
| 250 | e as (u,l) : ('a * 'a) = upper and lower bits
 | |
| 251 | l : 'a list = the list to get the upper and lower bounds of | |
| 252 | returns a pair ('a * 'a) of the biggest and lowest value w.r.t "gr"
 | |
| 253 | *) | |
| 254 | fun get_ends_of gr (e as (u,l)) [] = e | |
| 255 | | get_ends_of gr (e as (u,l)) (h :: t) = | |
| 256 | if gr(h,u) then get_ends_of gr (h,l) t | |
| 257 | else if gr(l,h) then get_ends_of gr (u,h) t | |
| 258 | else get_ends_of gr (u,l) t; | |
| 259 | ||
| 15570 | 260 | fun flatmap f = List.concat o map f; | 
| 15481 | 261 | |
| 262 | (* quick removal of "rs" elements from "ls" when (eqf (r,l)) is true | |
| 263 | Ignores ordering. *) | |
| 264 | fun lrem eqf rs ls = | |
| 265 | let fun list_remove rs ([],[]) = [] | |
| 266 | | list_remove [] (xs,_) = xs | |
| 267 | | list_remove (r::rs) ([],leftovers) = | |
| 268 | list_remove rs (leftovers,[]) | |
| 269 | | list_remove (r1::rs) ((x::xs),leftovers) = | |
| 270 | if eqf (r1, x) then list_remove (r1::rs) (xs,leftovers) | |
| 271 | else list_remove (r1::rs) (xs, x::leftovers) | |
| 272 | in | |
| 273 | list_remove rs (ls,[]) | |
| 274 | end; | |
| 275 | ||
| 276 | fun forall_list f [] = true | |
| 277 | | forall_list f (a::t) = f a orelse forall_list f t; | |
| 278 | ||
| 279 | ||
| 280 | (* crappy string indeter *) | |
| 281 | fun str_indent s = | |
| 282 | implode (map (fn s => if s = "\n" then "\n " else s) (explode s)); | |
| 283 | ||
| 284 | ||
| 285 | fun remove_end_spaces s = | |
| 286 | let | |
| 287 | fun rem_starts [] = [] | |
| 288 | 					| rem_starts (" " :: t) = rem_starts t
 | |
| 289 | 					| rem_starts ("\t" :: t) = rem_starts t
 | |
| 290 | 					| rem_starts ("\n" :: t) = rem_starts t
 | |
| 291 | | rem_starts l = l | |
| 292 | fun rem_ends l = rev (rem_starts (rev l)) | |
| 293 | in | |
| 294 | implode (rem_ends (rem_starts (explode s))) | |
| 295 | end; | |
| 296 | ||
| 297 | (* convert a string to an integer *) | |
| 298 | exception NOT_A_DIGIT of string; | |
| 299 | ||
| 300 | fun int_of_string s = | |
| 301 | let | |
| 302 | fun digits_to_int [] x = x | |
| 303 | | digits_to_int (h :: t) x = digits_to_int t (x * 10 + h); | |
| 304 | ||
| 305 | fun char_to_digit c = | |
| 306 | case c of | |
| 307 | "0" => 0 | |
| 308 | | "1" => 1 | |
| 309 | | "2" => 2 | |
| 310 | | "3" => 3 | |
| 311 | | "4" => 4 | |
| 312 | | "5" => 5 | |
| 313 | | "6" => 6 | |
| 314 | | "7" => 7 | |
| 315 | | "8" => 8 | |
| 316 | | "9" => 9 | |
| 317 | | _ => raise NOT_A_DIGIT c | |
| 318 | in | |
| 319 | digits_to_int (map char_to_digit (explode (remove_end_spaces s))) 0 | |
| 320 | end; | |
| 321 | ||
| 322 | (* Debugging/printing code for this datatype *) | |
| 323 | fun string_of_list f l = | |
| 324 | let | |
| 325 | fun auxf [] = "" | |
| 326 | | auxf [a] = (f a) | |
| 327 | | auxf (h :: (t as (h2 :: t2))) = (f h) ^ ", " ^ (auxf t) | |
| 328 | in | |
| 329 | "[" ^ (auxf l) ^ "]" | |
| 330 | end; | |
| 331 | ||
| 332 | val string_of_intlist = string_of_list string_of_int; | |
| 333 | ||
| 334 | ||
| 335 | (* options *) | |
| 15531 | 336 | fun aptosome NONE f = NONE | 
| 337 | | aptosome (SOME x) f = SOME (f x); | |
| 15481 | 338 | |
| 339 | end; |