| author | wenzelm | 
| Fri, 03 Mar 2017 17:53:24 +0100 | |
| changeset 65093 | 5f08197206ce | 
| parent 64578 | 7b20f9f94f4e | 
| child 67443 | 3abf6a722518 | 
| permissions | -rw-r--r-- | 
| 59720 | 1  | 
(* Title: HOL/Library/Code_Test.thy  | 
| 64578 | 2  | 
Author: Andreas Lochbihler, ETH Zürich  | 
| 58039 | 3  | 
|
| 64578 | 4  | 
Test infrastructure for the code generator.  | 
| 58039 | 5  | 
*)  | 
6  | 
||
7  | 
theory Code_Test  | 
|
8  | 
imports Main  | 
|
| 
58348
 
2d47c7d10b62
add target language evaluators for the value command;
 
Andreas Lochbihler 
parents: 
58310 
diff
changeset
 | 
9  | 
keywords "test_code" :: diag  | 
| 58039 | 10  | 
begin  | 
11  | 
||
| 60500 | 12  | 
subsection \<open>YXML encoding for @{typ Code_Evaluation.term}\<close>
 | 
| 58039 | 13  | 
|
| 58626 | 14  | 
datatype (plugins del: code size "quickcheck") yxml_of_term = YXML  | 
| 58039 | 15  | 
|
16  | 
lemma yot_anything: "x = (y :: yxml_of_term)"  | 
|
17  | 
by(cases x y rule: yxml_of_term.exhaust[case_product yxml_of_term.exhaust])(simp)  | 
|
18  | 
||
19  | 
definition yot_empty :: yxml_of_term where [code del]: "yot_empty = YXML"  | 
|
20  | 
definition yot_literal :: "String.literal \<Rightarrow> yxml_of_term"  | 
|
21  | 
where [code del]: "yot_literal _ = YXML"  | 
|
22  | 
definition yot_append :: "yxml_of_term \<Rightarrow> yxml_of_term \<Rightarrow> yxml_of_term"  | 
|
23  | 
where [code del]: "yot_append _ _ = YXML"  | 
|
24  | 
definition yot_concat :: "yxml_of_term list \<Rightarrow> yxml_of_term"  | 
|
25  | 
where [code del]: "yot_concat _ = YXML"  | 
|
26  | 
||
| 60500 | 27  | 
text \<open>Serialise @{typ yxml_of_term} to native string of target language\<close>
 | 
| 58039 | 28  | 
|
29  | 
code_printing type_constructor yxml_of_term  | 
|
30  | 
\<rightharpoonup> (SML) "string"  | 
|
31  | 
and (OCaml) "string"  | 
|
32  | 
and (Haskell) "String"  | 
|
33  | 
and (Scala) "String"  | 
|
34  | 
| constant yot_empty  | 
|
35  | 
\<rightharpoonup> (SML) "\"\""  | 
|
36  | 
and (OCaml) "\"\""  | 
|
37  | 
and (Haskell) "\"\""  | 
|
38  | 
and (Scala) "\"\""  | 
|
39  | 
| constant yot_literal  | 
|
40  | 
\<rightharpoonup> (SML) "_"  | 
|
41  | 
and (OCaml) "_"  | 
|
42  | 
and (Haskell) "_"  | 
|
43  | 
and (Scala) "_"  | 
|
44  | 
| constant yot_append  | 
|
45  | 
\<rightharpoonup> (SML) "String.concat [(_), (_)]"  | 
|
46  | 
and (OCaml) "String.concat \"\" [(_); (_)]"  | 
|
47  | 
and (Haskell) infixr 5 "++"  | 
|
48  | 
and (Scala) infixl 5 "+"  | 
|
49  | 
| constant yot_concat  | 
|
50  | 
\<rightharpoonup> (SML) "String.concat"  | 
|
51  | 
and (OCaml) "String.concat \"\""  | 
|
52  | 
and (Haskell) "Prelude.concat"  | 
|
53  | 
and (Scala) "_.mkString(\"\")"  | 
|
54  | 
||
| 60500 | 55  | 
text \<open>  | 
| 63680 | 56  | 
Stripped-down implementations of Isabelle's XML tree with YXML encoding as  | 
57  | 
defined in \<^file>\<open>~~/src/Pure/PIDE/xml.ML\<close>, \<^file>\<open>~~/src/Pure/PIDE/yxml.ML\<close>  | 
|
58  | 
  sufficient to encode @{typ "Code_Evaluation.term"} as in
 | 
|
59  | 
\<^file>\<open>~~/src/Pure/term_xml.ML\<close>.  | 
|
| 60500 | 60  | 
\<close>  | 
| 58039 | 61  | 
|
| 58626 | 62  | 
datatype (plugins del: code size "quickcheck") xml_tree = XML_Tree  | 
| 58039 | 63  | 
|
64  | 
lemma xml_tree_anything: "x = (y :: xml_tree)"  | 
|
65  | 
by(cases x y rule: xml_tree.exhaust[case_product xml_tree.exhaust])(simp)  | 
|
66  | 
||
67  | 
context begin  | 
|
| 60500 | 68  | 
local_setup \<open>Local_Theory.map_background_naming (Name_Space.mandatory_path "xml")\<close>  | 
| 58039 | 69  | 
|
70  | 
type_synonym attributes = "(String.literal \<times> String.literal) list"  | 
|
71  | 
type_synonym body = "xml_tree list"  | 
|
72  | 
||
73  | 
definition Elem :: "String.literal \<Rightarrow> attributes \<Rightarrow> xml_tree list \<Rightarrow> xml_tree"  | 
|
74  | 
where [code del]: "Elem _ _ _ = XML_Tree"  | 
|
75  | 
||
76  | 
definition Text :: "String.literal \<Rightarrow> xml_tree"  | 
|
77  | 
where [code del]: "Text _ = XML_Tree"  | 
|
78  | 
||
79  | 
definition node :: "xml_tree list \<Rightarrow> xml_tree"  | 
|
80  | 
where "node ts = Elem (STR '':'') [] ts"  | 
|
81  | 
||
82  | 
definition tagged :: "String.literal \<Rightarrow> String.literal option \<Rightarrow> xml_tree list \<Rightarrow> xml_tree"  | 
|
83  | 
where "tagged tag x ts = Elem tag (case x of None \<Rightarrow> [] | Some x' \<Rightarrow> [(STR ''0'', x')]) ts"  | 
|
84  | 
||
85  | 
definition list where "list f xs = map (node \<circ> f) xs"  | 
|
86  | 
||
| 62597 | 87  | 
definition X :: yxml_of_term where "X = yot_literal (STR [Char (num.Bit1 (num.Bit0 num.One))])"  | 
88  | 
definition Y :: yxml_of_term where "Y = yot_literal (STR [Char (num.Bit0 (num.Bit1 num.One))])"  | 
|
| 58039 | 89  | 
definition XY :: yxml_of_term where "XY = yot_append X Y"  | 
90  | 
definition XYX :: yxml_of_term where "XYX = yot_append XY X"  | 
|
91  | 
||
92  | 
end  | 
|
93  | 
||
94  | 
code_datatype xml.Elem xml.Text  | 
|
95  | 
||
96  | 
definition yxml_string_of_xml_tree :: "xml_tree \<Rightarrow> yxml_of_term \<Rightarrow> yxml_of_term"  | 
|
97  | 
where [code del]: "yxml_string_of_xml_tree _ _ = YXML"  | 
|
98  | 
||
99  | 
lemma yxml_string_of_xml_tree_code [code]:  | 
|
100  | 
"yxml_string_of_xml_tree (xml.Elem name atts ts) rest =  | 
|
101  | 
yot_append xml.XY (  | 
|
102  | 
yot_append (yot_literal name) (  | 
|
| 64578 | 103  | 
foldr (\<lambda>(a, x) rest.  | 
| 58039 | 104  | 
yot_append xml.Y (  | 
105  | 
yot_append (yot_literal a) (  | 
|
106  | 
yot_append (yot_literal (STR ''='')) (  | 
|
107  | 
yot_append (yot_literal x) rest)))) atts (  | 
|
108  | 
foldr yxml_string_of_xml_tree ts (  | 
|
109  | 
yot_append xml.XYX rest))))"  | 
|
110  | 
"yxml_string_of_xml_tree (xml.Text s) rest = yot_append (yot_literal s) rest"  | 
|
111  | 
by(rule yot_anything)+  | 
|
112  | 
||
113  | 
definition yxml_string_of_body :: "xml.body \<Rightarrow> yxml_of_term"  | 
|
114  | 
where "yxml_string_of_body ts = foldr yxml_string_of_xml_tree ts yot_empty"  | 
|
115  | 
||
| 60500 | 116  | 
text \<open>  | 
| 63680 | 117  | 
  Encoding @{typ Code_Evaluation.term} into XML trees as defined in
 | 
118  | 
\<^file>\<open>~~/src/Pure/term_xml.ML\<close>.  | 
|
| 60500 | 119  | 
\<close>  | 
| 58039 | 120  | 
|
121  | 
definition xml_of_typ :: "Typerep.typerep \<Rightarrow> xml.body"  | 
|
122  | 
where [code del]: "xml_of_typ _ = [XML_Tree]"  | 
|
123  | 
||
124  | 
definition xml_of_term :: "Code_Evaluation.term \<Rightarrow> xml.body"  | 
|
125  | 
where [code del]: "xml_of_term _ = [XML_Tree]"  | 
|
126  | 
||
127  | 
lemma xml_of_typ_code [code]:  | 
|
128  | 
"xml_of_typ (typerep.Typerep t args) = [xml.tagged (STR ''0'') (Some t) (xml.list xml_of_typ args)]"  | 
|
129  | 
by(simp add: xml_of_typ_def xml_tree_anything)  | 
|
130  | 
||
131  | 
lemma xml_of_term_code [code]:  | 
|
132  | 
"xml_of_term (Code_Evaluation.Const x ty) = [xml.tagged (STR ''0'') (Some x) (xml_of_typ ty)]"  | 
|
133  | 
"xml_of_term (Code_Evaluation.App t1 t2) = [xml.tagged (STR ''5'') None [xml.node (xml_of_term t1), xml.node (xml_of_term t2)]]"  | 
|
134  | 
"xml_of_term (Code_Evaluation.Abs x ty t) = [xml.tagged (STR ''4'') (Some x) [xml.node (xml_of_typ ty), xml.node (xml_of_term t)]]"  | 
|
| 61585 | 135  | 
\<comment> \<open>  | 
| 58637 | 136  | 
    FIXME: @{const Code_Evaluation.Free} is used only in @{theory Quickcheck_Narrowing} to represent
 | 
| 58039 | 137  | 
    uninstantiated parameters in constructors. Here, we always translate them to @{ML Free} variables.
 | 
| 60500 | 138  | 
\<close>  | 
| 58039 | 139  | 
"xml_of_term (Code_Evaluation.Free x ty) = [xml.tagged (STR ''1'') (Some x) (xml_of_typ ty)]"  | 
140  | 
by(simp_all add: xml_of_term_def xml_tree_anything)  | 
|
141  | 
||
142  | 
definition yxml_string_of_term :: "Code_Evaluation.term \<Rightarrow> yxml_of_term"  | 
|
143  | 
where "yxml_string_of_term = yxml_string_of_body \<circ> xml_of_term"  | 
|
144  | 
||
| 60500 | 145  | 
subsection \<open>Test engine and drivers\<close>  | 
| 58039 | 146  | 
|
147  | 
ML_file "code_test.ML"  | 
|
148  | 
||
149  | 
end  |