| author | haftmann | 
| Sat, 24 Dec 2011 15:55:03 +0100 | |
| changeset 45978 | d3325de5f299 | 
| parent 45490 | 20c8c0cca555 | 
| child 46483 | 10a9c31a22b4 | 
| permissions | -rw-r--r-- | 
| 21759 | 1  | 
(* Title: HOL/Tools/string_syntax.ML  | 
2  | 
Author: Makarius  | 
|
3  | 
||
4  | 
Concrete syntax for hex chars and strings.  | 
|
5  | 
*)  | 
|
6  | 
||
7  | 
signature STRING_SYNTAX =  | 
|
8  | 
sig  | 
|
9  | 
val setup: theory -> theory  | 
|
10  | 
end;  | 
|
11  | 
||
| 35123 | 12  | 
structure String_Syntax: STRING_SYNTAX =  | 
| 21759 | 13  | 
struct  | 
14  | 
||
15  | 
||
16  | 
(* nibble *)  | 
|
17  | 
||
18  | 
val mk_nib =  | 
|
| 
42290
 
b1f544c84040
discontinued special treatment of structure Lexicon;
 
wenzelm 
parents: 
42248 
diff
changeset
 | 
19  | 
Ast.Constant o Lexicon.mark_const o  | 
| 35123 | 20  | 
fst o Term.dest_Const o HOLogic.mk_nibble;  | 
| 21759 | 21  | 
|
| 
42224
 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 
wenzelm 
parents: 
40627 
diff
changeset
 | 
22  | 
fun dest_nib (Ast.Constant s) =  | 
| 
42290
 
b1f544c84040
discontinued special treatment of structure Lexicon;
 
wenzelm 
parents: 
42248 
diff
changeset
 | 
23  | 
(case try Lexicon.unmark_const s of  | 
| 35256 | 24  | 
NONE => raise Match  | 
25  | 
| SOME c => (HOLogic.dest_nibble (Const (c, HOLogic.nibbleT)) handle TERM _ => raise Match));  | 
|
| 21759 | 26  | 
|
27  | 
||
28  | 
(* char *)  | 
|
29  | 
||
30  | 
fun mk_char s =  | 
|
31  | 
if Symbol.is_ascii s then  | 
|
| 
42224
 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 
wenzelm 
parents: 
40627 
diff
changeset
 | 
32  | 
    Ast.Appl [Ast.Constant @{const_syntax Char}, mk_nib (ord s div 16), mk_nib (ord s mod 16)]
 | 
| 21759 | 33  | 
  else error ("Non-ASCII symbol: " ^ quote s);
 | 
34  | 
||
| 
40627
 
becf5d5187cc
renamed raw "explode" function to "raw_explode" to emphasize its meaning;
 
wenzelm 
parents: 
35363 
diff
changeset
 | 
35  | 
val specials = raw_explode "\\\"`'";  | 
| 21759 | 36  | 
|
37  | 
fun dest_chr c1 c2 =  | 
|
38  | 
let val c = chr (dest_nib c1 * 16 + dest_nib c2) in  | 
|
39  | 
if not (member (op =) specials c) andalso Symbol.is_ascii c andalso Symbol.is_printable c  | 
|
40  | 
then c else raise Match  | 
|
41  | 
end;  | 
|
42  | 
||
| 
42224
 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 
wenzelm 
parents: 
40627 
diff
changeset
 | 
43  | 
fun dest_char (Ast.Appl [Ast.Constant @{const_syntax Char}, c1, c2]) = dest_chr c1 c2
 | 
| 21759 | 44  | 
| dest_char _ = raise Match;  | 
45  | 
||
| 29317 | 46  | 
fun syntax_string cs =  | 
| 
42290
 
b1f544c84040
discontinued special treatment of structure Lexicon;
 
wenzelm 
parents: 
42248 
diff
changeset
 | 
47  | 
  Ast.Appl [Ast.Constant @{syntax_const "_inner_string"},
 | 
| 
 
b1f544c84040
discontinued special treatment of structure Lexicon;
 
wenzelm 
parents: 
42248 
diff
changeset
 | 
48  | 
Ast.Variable (Lexicon.implode_xstr cs)];  | 
| 21759 | 49  | 
|
50  | 
||
| 
42224
 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 
wenzelm 
parents: 
40627 
diff
changeset
 | 
51  | 
fun char_ast_tr [Ast.Variable xstr] =  | 
| 
42290
 
b1f544c84040
discontinued special treatment of structure Lexicon;
 
wenzelm 
parents: 
42248 
diff
changeset
 | 
52  | 
(case Lexicon.explode_xstr xstr of  | 
| 
42224
 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 
wenzelm 
parents: 
40627 
diff
changeset
 | 
53  | 
[c] => mk_char c  | 
| 
 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 
wenzelm 
parents: 
40627 
diff
changeset
 | 
54  | 
      | _ => error ("Single character expected: " ^ xstr))
 | 
| 45490 | 55  | 
  | char_ast_tr [Ast.Appl [Ast.Constant @{syntax_const "_constrain"}, ast1, ast2]] =
 | 
56  | 
      Ast.Appl [Ast.Constant @{syntax_const "_constrain"}, char_ast_tr [ast1], ast2]
 | 
|
| 
42224
 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 
wenzelm 
parents: 
40627 
diff
changeset
 | 
57  | 
  | char_ast_tr asts = raise Ast.AST ("char_ast_tr", asts);
 | 
| 21759 | 58  | 
|
| 35115 | 59  | 
fun char_ast_tr' [c1, c2] =  | 
| 
42224
 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 
wenzelm 
parents: 
40627 
diff
changeset
 | 
60  | 
      Ast.Appl [Ast.Constant @{syntax_const "_Char"}, syntax_string [dest_chr c1 c2]]
 | 
| 21759 | 61  | 
| char_ast_tr' _ = raise Match;  | 
62  | 
||
63  | 
||
64  | 
(* string *)  | 
|
65  | 
||
| 
42224
 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 
wenzelm 
parents: 
40627 
diff
changeset
 | 
66  | 
fun mk_string [] = Ast.Constant @{const_syntax Nil}
 | 
| 35115 | 67  | 
| mk_string (c :: cs) =  | 
| 
42224
 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 
wenzelm 
parents: 
40627 
diff
changeset
 | 
68  | 
      Ast.Appl [Ast.Constant @{const_syntax Cons}, mk_char c, mk_string cs];
 | 
| 21759 | 69  | 
|
| 
42224
 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 
wenzelm 
parents: 
40627 
diff
changeset
 | 
70  | 
fun string_ast_tr [Ast.Variable xstr] =  | 
| 
42290
 
b1f544c84040
discontinued special treatment of structure Lexicon;
 
wenzelm 
parents: 
42248 
diff
changeset
 | 
71  | 
(case Lexicon.explode_xstr xstr of  | 
| 
42224
 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 
wenzelm 
parents: 
40627 
diff
changeset
 | 
72  | 
[] =>  | 
| 
 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 
wenzelm 
parents: 
40627 
diff
changeset
 | 
73  | 
Ast.Appl  | 
| 42248 | 74  | 
            [Ast.Constant @{syntax_const "_constrain"},
 | 
| 
42224
 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 
wenzelm 
parents: 
40627 
diff
changeset
 | 
75  | 
              Ast.Constant @{const_syntax Nil}, Ast.Constant @{type_syntax string}]
 | 
| 
 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 
wenzelm 
parents: 
40627 
diff
changeset
 | 
76  | 
| cs => mk_string cs)  | 
| 45490 | 77  | 
  | string_ast_tr [Ast.Appl [Ast.Constant @{syntax_const "_constrain"}, ast1, ast2]] =
 | 
78  | 
      Ast.Appl [Ast.Constant @{syntax_const "_constrain"}, string_ast_tr [ast1], ast2]
 | 
|
| 
42224
 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 
wenzelm 
parents: 
40627 
diff
changeset
 | 
79  | 
  | string_ast_tr asts = raise Ast.AST ("string_tr", asts);
 | 
| 21759 | 80  | 
|
| 35115 | 81  | 
fun list_ast_tr' [args] =  | 
| 
42224
 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 
wenzelm 
parents: 
40627 
diff
changeset
 | 
82  | 
      Ast.Appl [Ast.Constant @{syntax_const "_String"},
 | 
| 
 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 
wenzelm 
parents: 
40627 
diff
changeset
 | 
83  | 
        syntax_string (map dest_char (Ast.unfold_ast @{syntax_const "_args"} args))]
 | 
| 21759 | 84  | 
| list_ast_tr' ts = raise Match;  | 
85  | 
||
86  | 
||
87  | 
(* theory setup *)  | 
|
88  | 
||
89  | 
val setup =  | 
|
| 
24712
 
64ed05609568
proper Sign operations instead of Theory aliases;
 
wenzelm 
parents: 
21775 
diff
changeset
 | 
90  | 
Sign.add_trfuns  | 
| 35115 | 91  | 
   ([(@{syntax_const "_Char"}, char_ast_tr), (@{syntax_const "_String"}, string_ast_tr)], [], [],
 | 
92  | 
    [(@{const_syntax Char}, char_ast_tr'), (@{syntax_const "_list"}, list_ast_tr')]);
 | 
|
| 21759 | 93  | 
|
94  | 
end;  |