| author | blanchet | 
| Fri, 31 Jan 2014 18:43:16 +0100 | |
| changeset 55221 | ee90eebb8b73 | 
| parent 55108 | 0b7a0c1fdf7e | 
| child 58822 | 90a5e981af3e | 
| 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 =  | 
|
| 
55015
 
e33c5bd729ff
added \<newline> symbol, which is used for char/string literals in HOL;
 
wenzelm 
parents: 
52143 
diff
changeset
 | 
31  | 
let  | 
| 
 
e33c5bd729ff
added \<newline> symbol, which is used for char/string literals in HOL;
 
wenzelm 
parents: 
52143 
diff
changeset
 | 
32  | 
val c =  | 
| 
 
e33c5bd729ff
added \<newline> symbol, which is used for char/string literals in HOL;
 
wenzelm 
parents: 
52143 
diff
changeset
 | 
33  | 
if Symbol.is_ascii s then ord s  | 
| 
 
e33c5bd729ff
added \<newline> symbol, which is used for char/string literals in HOL;
 
wenzelm 
parents: 
52143 
diff
changeset
 | 
34  | 
else if s = "\<newline>" then 10  | 
| 
 
e33c5bd729ff
added \<newline> symbol, which is used for char/string literals in HOL;
 
wenzelm 
parents: 
52143 
diff
changeset
 | 
35  | 
      else error ("Bad character: " ^ quote s);
 | 
| 
 
e33c5bd729ff
added \<newline> symbol, which is used for char/string literals in HOL;
 
wenzelm 
parents: 
52143 
diff
changeset
 | 
36  | 
  in Ast.Appl [Ast.Constant @{const_syntax Char}, mk_nib (c div 16), mk_nib (c mod 16)] end;
 | 
| 21759 | 37  | 
|
| 
40627
 
becf5d5187cc
renamed raw "explode" function to "raw_explode" to emphasize its meaning;
 
wenzelm 
parents: 
35363 
diff
changeset
 | 
38  | 
val specials = raw_explode "\\\"`'";  | 
| 21759 | 39  | 
|
40  | 
fun dest_chr c1 c2 =  | 
|
| 
55108
 
0b7a0c1fdf7e
inner syntax token language allows regular quoted strings;
 
wenzelm 
parents: 
55015 
diff
changeset
 | 
41  | 
let val s = chr (dest_nib c1 * 16 + dest_nib c2) in  | 
| 
 
0b7a0c1fdf7e
inner syntax token language allows regular quoted strings;
 
wenzelm 
parents: 
55015 
diff
changeset
 | 
42  | 
if not (member (op =) specials s) andalso Symbol.is_ascii s andalso Symbol.is_printable s  | 
| 
 
0b7a0c1fdf7e
inner syntax token language allows regular quoted strings;
 
wenzelm 
parents: 
55015 
diff
changeset
 | 
43  | 
then s  | 
| 
 
0b7a0c1fdf7e
inner syntax token language allows regular quoted strings;
 
wenzelm 
parents: 
55015 
diff
changeset
 | 
44  | 
else if s = "\n" then "\<newline>"  | 
| 
55015
 
e33c5bd729ff
added \<newline> symbol, which is used for char/string literals in HOL;
 
wenzelm 
parents: 
52143 
diff
changeset
 | 
45  | 
else raise Match  | 
| 21759 | 46  | 
end;  | 
47  | 
||
| 
42224
 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 
wenzelm 
parents: 
40627 
diff
changeset
 | 
48  | 
fun dest_char (Ast.Appl [Ast.Constant @{const_syntax Char}, c1, c2]) = dest_chr c1 c2
 | 
| 21759 | 49  | 
| dest_char _ = raise Match;  | 
50  | 
||
| 
55108
 
0b7a0c1fdf7e
inner syntax token language allows regular quoted strings;
 
wenzelm 
parents: 
55015 
diff
changeset
 | 
51  | 
fun syntax_string ss =  | 
| 
42290
 
b1f544c84040
discontinued special treatment of structure Lexicon;
 
wenzelm 
parents: 
42248 
diff
changeset
 | 
52  | 
  Ast.Appl [Ast.Constant @{syntax_const "_inner_string"},
 | 
| 
55108
 
0b7a0c1fdf7e
inner syntax token language allows regular quoted strings;
 
wenzelm 
parents: 
55015 
diff
changeset
 | 
53  | 
Ast.Variable (Lexicon.implode_str ss)];  | 
| 21759 | 54  | 
|
55  | 
||
| 46483 | 56  | 
fun char_ast_tr [Ast.Variable str] =  | 
| 
55108
 
0b7a0c1fdf7e
inner syntax token language allows regular quoted strings;
 
wenzelm 
parents: 
55015 
diff
changeset
 | 
57  | 
(case Lexicon.explode_str (str, Position.none) of  | 
| 
 
0b7a0c1fdf7e
inner syntax token language allows regular quoted strings;
 
wenzelm 
parents: 
55015 
diff
changeset
 | 
58  | 
[(s, _)] => mk_char s  | 
| 46483 | 59  | 
      | _ => error ("Single character expected: " ^ str))
 | 
| 45490 | 60  | 
  | char_ast_tr [Ast.Appl [Ast.Constant @{syntax_const "_constrain"}, ast1, ast2]] =
 | 
61  | 
      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
 | 
62  | 
  | char_ast_tr asts = raise Ast.AST ("char_ast_tr", asts);
 | 
| 21759 | 63  | 
|
| 35115 | 64  | 
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
 | 
65  | 
      Ast.Appl [Ast.Constant @{syntax_const "_Char"}, syntax_string [dest_chr c1 c2]]
 | 
| 21759 | 66  | 
| char_ast_tr' _ = raise Match;  | 
67  | 
||
68  | 
||
69  | 
(* string *)  | 
|
70  | 
||
| 
42224
 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 
wenzelm 
parents: 
40627 
diff
changeset
 | 
71  | 
fun mk_string [] = Ast.Constant @{const_syntax Nil}
 | 
| 
55108
 
0b7a0c1fdf7e
inner syntax token language allows regular quoted strings;
 
wenzelm 
parents: 
55015 
diff
changeset
 | 
72  | 
| mk_string (s :: ss) =  | 
| 
 
0b7a0c1fdf7e
inner syntax token language allows regular quoted strings;
 
wenzelm 
parents: 
55015 
diff
changeset
 | 
73  | 
      Ast.Appl [Ast.Constant @{const_syntax Cons}, mk_char s, mk_string ss];
 | 
| 21759 | 74  | 
|
| 46483 | 75  | 
fun string_ast_tr [Ast.Variable str] =  | 
| 
55108
 
0b7a0c1fdf7e
inner syntax token language allows regular quoted strings;
 
wenzelm 
parents: 
55015 
diff
changeset
 | 
76  | 
(case Lexicon.explode_str (str, Position.none) of  | 
| 
42224
 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 
wenzelm 
parents: 
40627 
diff
changeset
 | 
77  | 
[] =>  | 
| 
 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 
wenzelm 
parents: 
40627 
diff
changeset
 | 
78  | 
Ast.Appl  | 
| 42248 | 79  | 
            [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
 | 
80  | 
              Ast.Constant @{const_syntax Nil}, Ast.Constant @{type_syntax string}]
 | 
| 
55108
 
0b7a0c1fdf7e
inner syntax token language allows regular quoted strings;
 
wenzelm 
parents: 
55015 
diff
changeset
 | 
81  | 
| ss => mk_string (map Symbol_Pos.symbol ss))  | 
| 45490 | 82  | 
  | string_ast_tr [Ast.Appl [Ast.Constant @{syntax_const "_constrain"}, ast1, ast2]] =
 | 
83  | 
      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
 | 
84  | 
  | string_ast_tr asts = raise Ast.AST ("string_tr", asts);
 | 
| 21759 | 85  | 
|
| 35115 | 86  | 
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
 | 
87  | 
      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
 | 
88  | 
        syntax_string (map dest_char (Ast.unfold_ast @{syntax_const "_args"} args))]
 | 
| 
51160
 
599ff65b85e2
systematic conversions between nat and nibble/char;
 
haftmann 
parents: 
46483 
diff
changeset
 | 
89  | 
| list_ast_tr' _ = raise Match;  | 
| 21759 | 90  | 
|
91  | 
||
92  | 
(* theory setup *)  | 
|
93  | 
||
94  | 
val setup =  | 
|
| 52143 | 95  | 
Sign.parse_ast_translation  | 
96  | 
   [(@{syntax_const "_Char"}, K char_ast_tr),
 | 
|
97  | 
    (@{syntax_const "_String"}, K string_ast_tr)] #>
 | 
|
98  | 
Sign.print_ast_translation  | 
|
99  | 
   [(@{const_syntax Char}, K char_ast_tr'),
 | 
|
100  | 
    (@{syntax_const "_list"}, K list_ast_tr')];
 | 
|
| 21759 | 101  | 
|
102  | 
end;  |