| 77847 |      1 | (*  Title:      Pure/General/string.ML
 | 
|  |      2 |     Author:     Makarius
 | 
|  |      3 | 
 | 
|  |      4 | Additional operations for structure String, following Isabelle/ML conventions.
 | 
|  |      5 | 
 | 
|  |      6 | NB: Isabelle normally works with Symbol.symbol, which is represented as a
 | 
|  |      7 | small string fragment. Raw type char is rarely useful.
 | 
|  |      8 | *)
 | 
|  |      9 | 
 | 
|  |     10 | signature STRING =
 | 
|  |     11 | sig
 | 
|  |     12 |   include STRING
 | 
|  |     13 |   val nth: string -> int -> char
 | 
|  |     14 |   val exists: (char -> bool) -> string -> bool
 | 
|  |     15 |   val forall: (char -> bool) -> string -> bool
 | 
|  |     16 |   val member: string -> char -> bool
 | 
|  |     17 |   val fold: (char -> 'a -> 'a) -> string -> 'a -> 'a
 | 
|  |     18 |   val fold_rev: (char -> 'a -> 'a) -> string -> 'a -> 'a
 | 
|  |     19 | end;
 | 
|  |     20 | 
 | 
|  |     21 | structure String: STRING =
 | 
|  |     22 | struct
 | 
|  |     23 | 
 | 
|  |     24 | open String;
 | 
|  |     25 | 
 | 
|  |     26 | fun nth str i = sub (str, i);
 | 
|  |     27 | 
 | 
|  |     28 | fun exists pred str =
 | 
|  |     29 |   let
 | 
|  |     30 |     val n = size str;
 | 
|  |     31 |     fun ex i = Int.< (i, n) andalso (pred (nth str i) orelse ex (i + 1));
 | 
|  |     32 |   in ex 0 end;
 | 
|  |     33 | 
 | 
|  |     34 | fun forall pred = not o exists (not o pred);
 | 
|  |     35 | 
 | 
|  |     36 | fun member str c = exists (fn c' => c = c') str;
 | 
|  |     37 | 
 | 
|  |     38 | fun fold f str x0 =
 | 
|  |     39 |   let
 | 
|  |     40 |     val n = size str;
 | 
|  |     41 |     fun iter (x, i) = if Int.< (i, n) then iter (f (nth str i) x, i + 1) else x;
 | 
|  |     42 |   in iter (x0, 0) end;
 | 
|  |     43 | 
 | 
|  |     44 | fun fold_rev f str x0 =
 | 
|  |     45 |   let
 | 
|  |     46 |     val n = size str;
 | 
|  |     47 |     fun iter (x, i) = if Int.>= (i, 0) then iter (f (nth str i) x, i - 1) else x;
 | 
|  |     48 |   in iter (x0, n - 1) end;
 | 
|  |     49 | 
 | 
|  |     50 | end;
 |