| author | huffman | 
| Tue, 17 Apr 2007 00:37:14 +0200 | |
| changeset 22720 | 296813d7d306 | 
| parent 22499 | 68c8a8390e16 | 
| child 22744 | 5cbe966d67a2 | 
| permissions | -rw-r--r-- | 
| 21046 | 1 | (* ID: $Id$ | 
| 2 | Author: Florian Haftmann, TU Muenchen | |
| 3 | *) | |
| 4 | ||
| 5 | header {* Setup of code generator tools *}
 | |
| 6 | ||
| 7 | theory Code_Generator | |
| 8 | imports HOL | |
| 9 | begin | |
| 10 | ||
| 22385 | 11 | subsection {* SML code generator setup *}
 | 
| 21046 | 12 | |
| 13 | types_code | |
| 14 |   "bool"  ("bool")
 | |
| 15 | attach (term_of) {*
 | |
| 16 | fun term_of_bool b = if b then HOLogic.true_const else HOLogic.false_const; | |
| 17 | *} | |
| 18 | attach (test) {*
 | |
| 19 | fun gen_bool i = one_of [false, true]; | |
| 20 | *} | |
| 21 |   "prop"  ("bool")
 | |
| 22 | attach (term_of) {*
 | |
| 23 | fun term_of_prop b = | |
| 24 | HOLogic.mk_Trueprop (if b then HOLogic.true_const else HOLogic.false_const); | |
| 25 | *} | |
| 26 | ||
| 27 | consts_code | |
| 28 |   "Trueprop" ("(_)")
 | |
| 29 |   "True"    ("true")
 | |
| 30 |   "False"   ("false")
 | |
| 31 |   "Not"     ("not")
 | |
| 32 |   "op |"    ("(_ orelse/ _)")
 | |
| 33 |   "op &"    ("(_ andalso/ _)")
 | |
| 34 |   "HOL.If"      ("(if _/ then _/ else _)")
 | |
| 35 | ||
| 36 | setup {*
 | |
| 37 | let | |
| 38 | ||
| 39 | fun eq_codegen thy defs gr dep thyname b t = | |
| 40 | (case strip_comb t of | |
| 41 |        (Const ("op =", Type (_, [Type ("fun", _), _])), _) => NONE
 | |
| 42 |      | (Const ("op =", _), [t, u]) =>
 | |
| 43 | let | |
| 44 | val (gr', pt) = Codegen.invoke_codegen thy defs dep thyname false (gr, t); | |
| 45 | val (gr'', pu) = Codegen.invoke_codegen thy defs dep thyname false (gr', u); | |
| 46 | val (gr''', _) = Codegen.invoke_tycodegen thy defs dep thyname false (gr'', HOLogic.boolT) | |
| 47 | in | |
| 48 | SOME (gr''', Codegen.parens | |
| 49 | (Pretty.block [pt, Pretty.str " =", Pretty.brk 1, pu])) | |
| 50 | end | |
| 51 |      | (t as Const ("op =", _), ts) => SOME (Codegen.invoke_codegen
 | |
| 52 | thy defs dep thyname b (gr, Codegen.eta_expand t ts 2)) | |
| 53 | | _ => NONE); | |
| 54 | ||
| 55 | in | |
| 56 | ||
| 57 | Codegen.add_codegen "eq_codegen" eq_codegen | |
| 58 | ||
| 59 | end | |
| 60 | *} | |
| 61 | ||
| 62 | text {* Evaluation *}
 | |
| 63 | ||
| 22099 | 64 | method_setup evaluation = {*
 | 
| 21046 | 65 | let | 
| 66 | ||
| 67 | fun evaluation_tac i = Tactical.PRIMITIVE (Drule.fconv_rule | |
| 68 | (Drule.goals_conv (equal i) Codegen.evaluation_conv)); | |
| 69 | ||
| 22099 | 70 | in Method.no_args (Method.SIMPLE_METHOD' (evaluation_tac THEN' rtac TrueI)) end | 
| 71 | *} "solve goal by evaluation" | |
| 21046 | 72 | |
| 73 | ||
| 74 | subsection {* Generic code generator setup *}
 | |
| 75 | ||
| 21904 | 76 | text {* operational equality for code generation *}
 | 
| 77 | ||
| 22473 | 78 | class eq (attach "op =") = type | 
| 21904 | 79 | |
| 80 | ||
| 81 | text {* equality for Haskell *}
 | |
| 82 | ||
| 83 | code_class eq | |
| 84 | (Haskell "Eq" where "op =" \<equiv> "(==)") | |
| 85 | ||
| 86 | code_const "op =" | |
| 87 | (Haskell infixl 4 "==") | |
| 88 | ||
| 89 | ||
| 22423 | 90 | text {* type bool *}
 | 
| 91 | ||
| 92 | code_datatype True False | |
| 21904 | 93 | |
| 94 | lemma [code func]: | |
| 95 | shows "(False \<and> x) = False" | |
| 96 | and "(True \<and> x) = x" | |
| 97 | and "(x \<and> False) = False" | |
| 98 | and "(x \<and> True) = x" by simp_all | |
| 99 | ||
| 100 | lemma [code func]: | |
| 101 | shows "(False \<or> x) = x" | |
| 102 | and "(True \<or> x) = True" | |
| 103 | and "(x \<or> False) = x" | |
| 104 | and "(x \<or> True) = True" by simp_all | |
| 105 | ||
| 106 | lemma [code func]: | |
| 107 | shows "(\<not> True) = False" | |
| 108 | and "(\<not> False) = True" by (rule HOL.simp_thms)+ | |
| 109 | ||
| 110 | lemmas [code func] = imp_conv_disj | |
| 111 | ||
| 112 | lemmas [code func] = if_True if_False | |
| 113 | ||
| 114 | instance bool :: eq .. | |
| 115 | ||
| 116 | lemma [code func]: | |
| 117 | "True = P \<longleftrightarrow> P" by simp | |
| 118 | ||
| 119 | lemma [code func]: | |
| 120 | "False = P \<longleftrightarrow> \<not> P" by simp | |
| 121 | ||
| 122 | lemma [code func]: | |
| 123 | "P = True \<longleftrightarrow> P" by simp | |
| 124 | ||
| 125 | lemma [code func]: | |
| 126 | "P = False \<longleftrightarrow> \<not> P" by simp | |
| 127 | ||
| 128 | code_type bool | |
| 129 | (SML "bool") | |
| 130 | (OCaml "bool") | |
| 131 | (Haskell "Bool") | |
| 132 | ||
| 133 | code_instance bool :: eq | |
| 134 | (Haskell -) | |
| 135 | ||
| 136 | code_const "op = \<Colon> bool \<Rightarrow> bool \<Rightarrow> bool" | |
| 137 | (Haskell infixl 4 "==") | |
| 138 | ||
| 139 | code_const True and False and Not and "op &" and "op |" and If | |
| 140 | (SML "true" and "false" and "not" | |
| 141 | and infixl 1 "andalso" and infixl 0 "orelse" | |
| 142 | and "!(if (_)/ then (_)/ else (_))") | |
| 143 | (OCaml "true" and "false" and "not" | |
| 144 | and infixl 4 "&&" and infixl 2 "||" | |
| 145 | and "!(if (_)/ then (_)/ else (_))") | |
| 146 | (Haskell "True" and "False" and "not" | |
| 147 | and infixl 3 "&&" and infixl 2 "||" | |
| 148 | and "!(if (_)/ then (_)/ else (_))") | |
| 149 | ||
| 150 | code_reserved SML | |
| 151 | bool true false not | |
| 152 | ||
| 153 | code_reserved OCaml | |
| 154 | bool true false not | |
| 155 | ||
| 156 | ||
| 22423 | 157 | text {* type prop *}
 | 
| 158 | ||
| 159 | code_datatype Trueprop "prop" | |
| 160 | ||
| 161 | ||
| 162 | text {* type itself *}
 | |
| 21046 | 163 | |
| 22423 | 164 | code_datatype "TYPE('a)"
 | 
| 165 | ||
| 166 | ||
| 167 | text {* prevent unfolding of quantifier definitions *}
 | |
| 168 | ||
| 169 | lemma [code func]: | |
| 170 | "Ex = Ex" | |
| 171 | "All = All" | |
| 172 | by rule+ | |
| 21046 | 173 | |
| 174 | ||
| 22480 | 175 | text {* code generation for undefined as exception *}
 | 
| 21046 | 176 | |
| 22480 | 177 | code_const undefined | 
| 22499 | 178 | (SML "raise/ Fail/ \"undefined\"") | 
| 179 | (OCaml "failwith/ \"undefined\"") | |
| 22480 | 180 | (Haskell "error/ \"undefined\"") | 
| 21046 | 181 | |
| 21079 | 182 | code_reserved SML Fail | 
| 21904 | 183 | code_reserved OCaml failwith | 
| 21046 | 184 | |
| 185 | ||
| 21378 | 186 | subsection {* Evaluation oracle *}
 | 
| 187 | ||
| 22487 | 188 | oracle eval_oracle ("term") = {* fn thy => fn t => 
 | 
| 189 | if CodegenPackage.satisfies thy (HOLogic.dest_Trueprop t) [] | |
| 190 | then t | |
| 191 | else HOLogic.Trueprop $ (HOLogic.true_const) (*dummy*) | |
| 21869 | 192 | *} | 
| 21378 | 193 | |
| 22099 | 194 | method_setup eval = {*
 | 
| 195 | let | |
| 22487 | 196 | fun eval_tac thy = | 
| 197 | SUBGOAL (fn (t, i) => rtac (eval_oracle thy t) i) | |
| 198 | in | |
| 199 | Method.ctxt_args (fn ctxt => | |
| 200 | Method.SIMPLE_METHOD' (eval_tac (ProofContext.theory_of ctxt))) | |
| 201 | end | |
| 22099 | 202 | *} "solve goal by evaluation" | 
| 21378 | 203 | |
| 204 | ||
| 205 | subsection {* Normalization by evaluation *}
 | |
| 21046 | 206 | |
| 22099 | 207 | method_setup normalization = {*
 | 
| 21046 | 208 | let | 
| 209 | fun normalization_tac i = Tactical.PRIMITIVE (Drule.fconv_rule | |
| 21546 | 210 | (Drule.goals_conv (equal i) (HOLogic.Trueprop_conv | 
| 21149 | 211 | NBE.normalization_conv))); | 
| 21046 | 212 | in | 
| 22099 | 213 | Method.no_args (Method.SIMPLE_METHOD' (normalization_tac THEN' resolve_tac [TrueI, refl])) | 
| 214 | end | |
| 215 | *} "solve goal by normalization" | |
| 216 | ||
| 21046 | 217 | |
| 21059 | 218 | text {* lazy @{const If} *}
 | 
| 219 | ||
| 220 | definition | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21378diff
changeset | 221 | if_delayed :: "bool \<Rightarrow> (bool \<Rightarrow> 'a) \<Rightarrow> (bool \<Rightarrow> 'a) \<Rightarrow> 'a" where | 
| 21059 | 222 | "if_delayed b f g = (if b then f True else g False)" | 
| 223 | ||
| 224 | lemma [code func]: | |
| 225 | shows "if_delayed True f g = f True" | |
| 226 | and "if_delayed False f g = g False" | |
| 227 | unfolding if_delayed_def by simp_all | |
| 228 | ||
| 229 | lemma [normal pre, symmetric, normal post]: | |
| 230 | "(if b then x else y) = if_delayed b (\<lambda>_. x) (\<lambda>_. y)" | |
| 231 | unfolding if_delayed_def .. | |
| 232 | ||
| 233 | ||
| 21454 | 234 | hide (open) const if_delayed | 
| 21046 | 235 | |
| 22099 | 236 | end |