diff -r 07025152dd80 -r 2bc2a8599369 src/Doc/Codegen/Computations.thy --- a/src/Doc/Codegen/Computations.thy Mon Jan 14 18:33:52 2019 +0000 +++ b/src/Doc/Codegen/Computations.thy Mon Jan 14 18:33:53 2019 +0000 @@ -16,7 +16,7 @@ datatype %quote form = T | F | And form form | Or form form (*<*) -(*>*) ML %quotetypewriter \ +(*>*) ML %quote \ fun eval_form @{code T} = true | eval_form @{code F} = false | eval_form (@{code And} (p, q)) = @@ -97,7 +97,7 @@ The following example illustrates its basic usage: \ -ML %quotetypewriter \ +ML %quote \ local fun int_of_nat @{code "0 :: nat"} = 0 @@ -164,7 +164,7 @@ \noindent A simple application: \ -ML_val %quotetypewriter \ +ML_val %quote \ comp_nat \<^context> \<^term>\sum_list [Suc 0, Suc (Suc 0)] * Suc (Suc 0)\ \ @@ -177,7 +177,7 @@ and for all: \ -ML %quotetypewriter \ +ML %quote \ local fun int_of_nat @{code "0 :: nat"} = 0 @@ -269,7 +269,7 @@ \secref{sec:literal_input}.} \ -ML %quotetypewriter \ +ML %quote \ local fun raw_dvd (b, ct) = Thm.mk_binop \<^cterm>\Pure.eq :: bool \ bool \ prop\ @@ -310,7 +310,7 @@ \<^item> Partiality makes the whole conversion fall back to reflexivity. \ (*<*) -(*>*) ML_val %quotetypewriter \ +(*>*) ML_val %quote \ conv_dvd \<^context> \<^cterm>\7 dvd ( 62437867527846782 :: int)\; conv_dvd \<^context> \<^cterm>\7 dvd (-62437867527846783 :: int)\; \ @@ -336,7 +336,7 @@ Code_Target_Int.positive (r * s)" by simp -ML %quotetypewriter \ +ML %quote \ local fun integer_of_int (@{code int_of_integer} k) = k @@ -365,7 +365,7 @@ end \ -ML_val %quotetypewriter \ +ML_val %quote \ conv_div \<^context> \<^cterm>\46782454343499999992777742432342242323423425 div (7 :: int)\ \ @@ -393,7 +393,7 @@ once and for all: \ -ML %quotetypewriter \ +ML %quote \ val check_nat = @{computation_check terms: Trueprop "less :: nat \ nat \ bool" "plus :: nat \ nat \ nat" "times :: nat \ nat \ nat" @@ -406,7 +406,7 @@ of type \<^typ>\bool\ into \<^typ>\prop\. \ -ML_val %quotetypewriter \ +ML_val %quote \ check_nat \<^context> \<^cprop>\less (Suc (Suc 0)) (Suc (Suc (Suc 0)))\ \ @@ -448,27 +448,27 @@ paragraph \An example for \<^typ>\nat\\ -ML %quotetypewriter \ +ML %quote \ val check_nat = @{computation_check terms: Trueprop "even :: nat \ bool" "plus :: nat \ nat \ nat" "0 :: nat" Suc "1 :: nat" "2 :: nat" "3 :: nat" } \ -ML_val %quotetypewriter \ +ML_val %quote \ check_nat \<^context> \<^cprop>\even (Suc 0 + 1 + 2 + 3 + 4 + 5)\ \ paragraph \An example for \<^typ>\int\\ -ML %quotetypewriter \ +ML %quote \ val check_int = @{computation_check terms: Trueprop "even :: int \ bool" "plus :: int \ int \ int" "0 :: int" "1 :: int" "2 :: int" "3 :: int" "-1 :: int" } \ -ML_val %quotetypewriter \ +ML_val %quote \ check_int \<^context> \<^cprop>\even ((0::int) + 1 + 2 + 3 + -1 + -2 + -3)\ \ @@ -478,13 +478,13 @@ where "is_cap_letter s \ (case String.asciis_of_literal s of [] \ False | k # _ \ 65 \ k \ k \ 90)" (*<*) -(*>*) ML %quotetypewriter \ +(*>*) ML %quote \ val check_literal = @{computation_check terms: Trueprop is_cap_letter datatypes: bool String.literal } \ -ML_val %quotetypewriter \ +ML_val %quote \ check_literal \<^context> \<^cprop>\is_cap_letter (STR ''Q'')\ \ @@ -523,13 +523,13 @@ whose concrete values are given in list \<^term>\vs\. \ -ML %quotetypewriter (*<*) \\ +ML %quote (*<*) \\ lemma "thm": fixes x y z :: "'a::order" shows "x < y \ x < z \ interp (And (Less (Suc 0) (Suc (Suc 0))) (Less (Suc 0) 0)) [z, x, y]" -ML_prf %quotetypewriter +ML_prf %quote (*>*) \val thm = Reification.conv \<^context> @{thms interp.simps} \<^cterm>\x < y \ x < z\\ (*<*) by (tactic \ALLGOALS (resolve_tac \<^context> [thm])\) -(*>*) +(*>*) text \ \noindent By virtue of @{fact interp.simps}, \<^ML>\Reification.conv\ provides a conversion