diff -r a80d8ec6c998 -r 3dda49e08b9d src/HOL/Library/Code_Test.thy --- a/src/HOL/Library/Code_Test.thy Fri Jan 04 21:49:06 2019 +0100 +++ b/src/HOL/Library/Code_Test.thy Fri Jan 04 23:22:53 2019 +0100 @@ -9,7 +9,7 @@ keywords "test_code" :: diag begin -subsection \YXML encoding for @{typ Code_Evaluation.term}\ +subsection \YXML encoding for \<^typ>\Code_Evaluation.term\\ datatype (plugins del: code size "quickcheck") yxml_of_term = YXML @@ -24,7 +24,7 @@ definition yot_concat :: "yxml_of_term list \ yxml_of_term" where [code del]: "yot_concat _ = YXML" -text \Serialise @{typ yxml_of_term} to native string of target language\ +text \Serialise \<^typ>\yxml_of_term\ to native string of target language\ code_printing type_constructor yxml_of_term \ (SML) "string" @@ -55,7 +55,7 @@ text \ Stripped-down implementations of Isabelle's XML tree with YXML encoding as defined in \<^file>\~~/src/Pure/PIDE/xml.ML\, \<^file>\~~/src/Pure/PIDE/yxml.ML\ - sufficient to encode @{typ "Code_Evaluation.term"} as in + sufficient to encode \<^typ>\Code_Evaluation.term\ as in \<^file>\~~/src/Pure/term_xml.ML\. \ @@ -114,7 +114,7 @@ where "yxml_string_of_body ts = foldr yxml_string_of_xml_tree ts yot_empty" text \ - Encoding @{typ Code_Evaluation.term} into XML trees as defined in + Encoding \<^typ>\Code_Evaluation.term\ into XML trees as defined in \<^file>\~~/src/Pure/term_xml.ML\. \ @@ -132,8 +132,8 @@ "xml_of_term (Code_Evaluation.Const x ty) = [xml.tagged (STR ''0'') (Some x) (xml_of_typ ty)]" "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)]]" "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)]]" - \ \FIXME: @{const Code_Evaluation.Free} is used only in @{theory "HOL.Quickcheck_Narrowing"} to represent - uninstantiated parameters in constructors. Here, we always translate them to @{ML Free} variables.\ + \ \FIXME: \<^const>\Code_Evaluation.Free\ is used only in \<^theory>\HOL.Quickcheck_Narrowing\ to represent + uninstantiated parameters in constructors. Here, we always translate them to \<^ML>\Free\ variables.\ "xml_of_term (Code_Evaluation.Free x ty) = [xml.tagged (STR ''1'') (Some x) (xml_of_typ ty)]" by(simp_all add: xml_of_term_def xml_tree_anything)