src/HOL/Code_Setup.thy
author wenzelm
Fri, 13 Mar 2009 19:58:26 +0100
changeset 30510 4120fc59dd85
parent 29105 8f38bf68d42e
child 30549 d2d7874648bd
permissions -rw-r--r--
unified type Proof.method and pervasive METHOD combinators;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
24285
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
     1
(*  Title:      HOL/Code_Setup.thy
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
     2
    ID:         $Id$
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
     3
    Author:     Florian Haftmann
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
     4
*)
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
     5
28400
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
     6
header {* Setup of code generators and related tools *}
24285
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
     7
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
     8
theory Code_Setup
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
     9
imports HOL
28400
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    10
begin
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    11
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    12
subsection {* Generic code generator foundation *}
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    13
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    14
text {* Datatypes *}
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    15
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    16
code_datatype True False
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    17
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    18
code_datatype "TYPE('a\<Colon>{})"
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    19
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    20
code_datatype Trueprop "prop"
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    21
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    22
text {* Code equations *}
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    23
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 28537
diff changeset
    24
lemma [code]:
28740
22a8125d66fa improved handling of !!/==> for eval and normalization
haftmann
parents: 28699
diff changeset
    25
  shows "(True \<Longrightarrow> PROP P) \<equiv> PROP P" 
22a8125d66fa improved handling of !!/==> for eval and normalization
haftmann
parents: 28699
diff changeset
    26
    and "(False \<Longrightarrow> Q) \<equiv> Trueprop True" 
22a8125d66fa improved handling of !!/==> for eval and normalization
haftmann
parents: 28699
diff changeset
    27
    and "(PROP P \<Longrightarrow> True) \<equiv> Trueprop True" 
22a8125d66fa improved handling of !!/==> for eval and normalization
haftmann
parents: 28699
diff changeset
    28
    and "(Q \<Longrightarrow> False) \<equiv> Trueprop (\<not> Q)" by (auto intro!: equal_intr_rule)
22a8125d66fa improved handling of !!/==> for eval and normalization
haftmann
parents: 28699
diff changeset
    29
22a8125d66fa improved handling of !!/==> for eval and normalization
haftmann
parents: 28699
diff changeset
    30
lemma [code]:
28400
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    31
  shows "False \<and> x \<longleftrightarrow> False"
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    32
    and "True \<and> x \<longleftrightarrow> x"
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    33
    and "x \<and> False \<longleftrightarrow> False"
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    34
    and "x \<and> True \<longleftrightarrow> x" by simp_all
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    35
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 28537
diff changeset
    36
lemma [code]:
28400
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    37
  shows "False \<or> x \<longleftrightarrow> x"
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    38
    and "True \<or> x \<longleftrightarrow> True"
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    39
    and "x \<or> False \<longleftrightarrow> x"
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    40
    and "x \<or> True \<longleftrightarrow> True" by simp_all
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    41
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 28537
diff changeset
    42
lemma [code]:
28400
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    43
  shows "\<not> True \<longleftrightarrow> False"
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    44
    and "\<not> False \<longleftrightarrow> True" by (rule HOL.simp_thms)+
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    45
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 28537
diff changeset
    46
lemmas [code] = Let_def if_True if_False
28400
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    47
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 28537
diff changeset
    48
lemmas [code, code unfold, symmetric, code post] = imp_conv_disj
28400
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    49
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    50
text {* Equality *}
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    51
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    52
context eq
24285
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
    53
begin
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
    54
28562
4e74209f113e `code func` now just `code`
haftmann
parents: 28537
diff changeset
    55
lemma equals_eq [code inline, code]: "op = \<equiv> eq"
28400
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    56
  by (rule eq_reflection) (rule ext, rule ext, rule sym, rule eq_equals)
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    57
28512
f29fecd6ddaa corrected SML undefined
haftmann
parents: 28400
diff changeset
    58
declare eq [code unfold, code inline del]
f29fecd6ddaa corrected SML undefined
haftmann
parents: 28400
diff changeset
    59
28400
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    60
declare equals_eq [symmetric, code post]
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    61
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    62
end
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    63
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    64
declare simp_thms(6) [code nbe]
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    65
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    66
hide (open) const eq
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    67
hide const eq
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    68
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    69
setup {*
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    70
  Code_Unit.add_const_alias @{thm equals_eq}
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    71
*}
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    72
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    73
text {* Cases *}
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    74
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    75
lemma Let_case_cert:
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    76
  assumes "CASE \<equiv> (\<lambda>x. Let x f)"
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    77
  shows "CASE x \<equiv> f x"
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    78
  using assms by simp_all
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    79
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    80
lemma If_case_cert:
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    81
  assumes "CASE \<equiv> (\<lambda>b. If b f g)"
28856
5e009a80fe6d Pure syntax: more coherent treatment of aprop, permanent TERM and &&&;
wenzelm
parents: 28740
diff changeset
    82
  shows "(CASE True \<equiv> f) &&& (CASE False \<equiv> g)"
28400
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    83
  using assms by simp_all
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    84
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    85
setup {*
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    86
  Code.add_case @{thm Let_case_cert}
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    87
  #> Code.add_case @{thm If_case_cert}
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    88
  #> Code.add_undefined @{const_name undefined}
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    89
*}
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    90
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    91
code_abort undefined
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    92
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    93
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    94
subsection {* Generic code generator preprocessor *}
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    95
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    96
setup {*
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    97
  Code.map_pre (K HOL_basic_ss)
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    98
  #> Code.map_post (K HOL_basic_ss)
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
    99
*}
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   100
24285
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   101
28400
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   102
subsection {* Generic code generator target languages *}
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   103
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   104
text {* type bool *}
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   105
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   106
code_type bool
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   107
  (SML "bool")
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   108
  (OCaml "bool")
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   109
  (Haskell "Bool")
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   110
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   111
code_const True and False and Not and "op &" and "op |" and If
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   112
  (SML "true" and "false" and "not"
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   113
    and infixl 1 "andalso" and infixl 0 "orelse"
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   114
    and "!(if (_)/ then (_)/ else (_))")
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   115
  (OCaml "true" and "false" and "not"
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   116
    and infixl 4 "&&" and infixl 2 "||"
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   117
    and "!(if (_)/ then (_)/ else (_))")
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   118
  (Haskell "True" and "False" and "not"
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   119
    and infixl 3 "&&" and infixl 2 "||"
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   120
    and "!(if (_)/ then (_)/ else (_))")
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   121
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   122
code_reserved SML
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   123
  bool true false not
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   124
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   125
code_reserved OCaml
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   126
  bool not
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   127
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   128
text {* using built-in Haskell equality *}
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   129
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   130
code_class eq
28687
150a8a1eae1a simplified syntax for class parameters
haftmann
parents: 28562
diff changeset
   131
  (Haskell "Eq")
28400
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   132
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   133
code_const "eq_class.eq"
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   134
  (Haskell infixl 4 "==")
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   135
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   136
code_const "op ="
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   137
  (Haskell infixl 4 "==")
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   138
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   139
text {* undefined *}
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   140
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   141
code_const undefined
28512
f29fecd6ddaa corrected SML undefined
haftmann
parents: 28400
diff changeset
   142
  (SML "!(raise/ Fail/ \"undefined\")")
28400
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   143
  (OCaml "failwith/ \"undefined\"")
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   144
  (Haskell "error/ \"undefined\"")
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   145
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   146
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   147
subsection {* SML code generator setup *}
24285
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   148
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   149
types_code
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   150
  "bool"  ("bool")
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   151
attach (term_of) {*
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   152
fun term_of_bool b = if b then HOLogic.true_const else HOLogic.false_const;
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   153
*}
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   154
attach (test) {*
25885
6fbc3f54f819 New interface for test data generators.
berghofe
parents: 25866
diff changeset
   155
fun gen_bool i =
6fbc3f54f819 New interface for test data generators.
berghofe
parents: 25866
diff changeset
   156
  let val b = one_of [false, true]
6fbc3f54f819 New interface for test data generators.
berghofe
parents: 25866
diff changeset
   157
  in (b, fn () => term_of_bool b) end;
24285
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   158
*}
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   159
  "prop"  ("bool")
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   160
attach (term_of) {*
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   161
fun term_of_prop b =
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   162
  HOLogic.mk_Trueprop (if b then HOLogic.true_const else HOLogic.false_const);
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   163
*}
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   164
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   165
consts_code
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   166
  "Trueprop" ("(_)")
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   167
  "True"    ("true")
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   168
  "False"   ("false")
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   169
  "Not"     ("Bool.not")
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   170
  "op |"    ("(_ orelse/ _)")
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   171
  "op &"    ("(_ andalso/ _)")
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   172
  "If"      ("(if _/ then _/ else _)")
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   173
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   174
setup {*
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   175
let
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   176
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28512
diff changeset
   177
fun eq_codegen thy defs dep thyname b t gr =
24285
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   178
    (case strip_comb t of
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   179
       (Const ("op =", Type (_, [Type ("fun", _), _])), _) => NONE
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   180
     | (Const ("op =", _), [t, u]) =>
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   181
          let
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28512
diff changeset
   182
            val (pt, gr') = Codegen.invoke_codegen thy defs dep thyname false t gr;
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28512
diff changeset
   183
            val (pu, gr'') = Codegen.invoke_codegen thy defs dep thyname false u gr';
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28512
diff changeset
   184
            val (_, gr''') = Codegen.invoke_tycodegen thy defs dep thyname false HOLogic.boolT gr'';
24285
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   185
          in
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28512
diff changeset
   186
            SOME (Codegen.parens
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28512
diff changeset
   187
              (Pretty.block [pt, Codegen.str " =", Pretty.brk 1, pu]), gr''')
24285
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   188
          end
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   189
     | (t as Const ("op =", _), ts) => SOME (Codegen.invoke_codegen
28537
1e84256d1a8a established canonical argument order in SML code generators
haftmann
parents: 28512
diff changeset
   190
         thy defs dep thyname b (Codegen.eta_expand t ts 2) gr)
24285
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   191
     | _ => NONE);
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   192
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   193
in
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   194
  Codegen.add_codegen "eq_codegen" eq_codegen
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   195
end
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   196
*}
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   197
24463
841c2e24761f Smaller size and fewer iterations for quickcheck.
berghofe
parents: 24293
diff changeset
   198
28400
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   199
subsection {* Evaluation and normalization by evaluation *}
24285
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   200
29105
8f38bf68d42e moved value.ML to src/Tools
haftmann
parents: 28856
diff changeset
   201
setup {*
8f38bf68d42e moved value.ML to src/Tools
haftmann
parents: 28856
diff changeset
   202
  Value.add_evaluator ("SML", Codegen.eval_term o ProofContext.theory_of)
8f38bf68d42e moved value.ML to src/Tools
haftmann
parents: 28856
diff changeset
   203
*}
8f38bf68d42e moved value.ML to src/Tools
haftmann
parents: 28856
diff changeset
   204
27103
d8549f4d900b major refactorings in code generator modules
haftmann
parents: 26975
diff changeset
   205
ML {*
d8549f4d900b major refactorings in code generator modules
haftmann
parents: 26975
diff changeset
   206
structure Eval_Method =
d8549f4d900b major refactorings in code generator modules
haftmann
parents: 26975
diff changeset
   207
struct
d8549f4d900b major refactorings in code generator modules
haftmann
parents: 26975
diff changeset
   208
d8549f4d900b major refactorings in code generator modules
haftmann
parents: 26975
diff changeset
   209
val eval_ref : (unit -> bool) option ref = ref NONE;
d8549f4d900b major refactorings in code generator modules
haftmann
parents: 26975
diff changeset
   210
d8549f4d900b major refactorings in code generator modules
haftmann
parents: 26975
diff changeset
   211
end;
d8549f4d900b major refactorings in code generator modules
haftmann
parents: 26975
diff changeset
   212
*}
d8549f4d900b major refactorings in code generator modules
haftmann
parents: 26975
diff changeset
   213
28290
4cc2b6046258 simplified oracle interface;
wenzelm
parents: 28054
diff changeset
   214
oracle eval_oracle = {* fn ct =>
4cc2b6046258 simplified oracle interface;
wenzelm
parents: 28054
diff changeset
   215
  let
4cc2b6046258 simplified oracle interface;
wenzelm
parents: 28054
diff changeset
   216
    val thy = Thm.theory_of_cterm ct;
4cc2b6046258 simplified oracle interface;
wenzelm
parents: 28054
diff changeset
   217
    val t = Thm.term_of ct;
28740
22a8125d66fa improved handling of !!/==> for eval and normalization
haftmann
parents: 28699
diff changeset
   218
    val dummy = @{cprop True};
22a8125d66fa improved handling of !!/==> for eval and normalization
haftmann
parents: 28699
diff changeset
   219
  in case try HOLogic.dest_Trueprop t
22a8125d66fa improved handling of !!/==> for eval and normalization
haftmann
parents: 28699
diff changeset
   220
   of SOME t' => if Code_ML.eval_term
22a8125d66fa improved handling of !!/==> for eval and normalization
haftmann
parents: 28699
diff changeset
   221
         ("Eval_Method.eval_ref", Eval_Method.eval_ref) thy t' [] 
22a8125d66fa improved handling of !!/==> for eval and normalization
haftmann
parents: 28699
diff changeset
   222
       then Thm.capply (Thm.capply @{cterm "op \<equiv> \<Colon> prop \<Rightarrow> prop \<Rightarrow> prop"} ct) dummy
22a8125d66fa improved handling of !!/==> for eval and normalization
haftmann
parents: 28699
diff changeset
   223
       else dummy
22a8125d66fa improved handling of !!/==> for eval and normalization
haftmann
parents: 28699
diff changeset
   224
    | NONE => dummy
28290
4cc2b6046258 simplified oracle interface;
wenzelm
parents: 28054
diff changeset
   225
  end
24285
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   226
*}
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   227
28740
22a8125d66fa improved handling of !!/==> for eval and normalization
haftmann
parents: 28699
diff changeset
   228
ML {*
30510
4120fc59dd85 unified type Proof.method and pervasive METHOD combinators;
wenzelm
parents: 29105
diff changeset
   229
fun gen_eval_method conv = Method.ctxt_args (fn ctxt => SIMPLE_METHOD'
28740
22a8125d66fa improved handling of !!/==> for eval and normalization
haftmann
parents: 28699
diff changeset
   230
  (CONVERSION (Conv.params_conv (~1) (K (Conv.concl_conv (~1) conv)) ctxt)
22a8125d66fa improved handling of !!/==> for eval and normalization
haftmann
parents: 28699
diff changeset
   231
    THEN' rtac TrueI))
22a8125d66fa improved handling of !!/==> for eval and normalization
haftmann
parents: 28699
diff changeset
   232
*}
24285
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   233
28740
22a8125d66fa improved handling of !!/==> for eval and normalization
haftmann
parents: 28699
diff changeset
   234
method_setup eval = {* gen_eval_method eval_oracle *}
22a8125d66fa improved handling of !!/==> for eval and normalization
haftmann
parents: 28699
diff changeset
   235
  "solve goal by evaluation"
24285
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   236
28740
22a8125d66fa improved handling of !!/==> for eval and normalization
haftmann
parents: 28699
diff changeset
   237
method_setup evaluation = {* gen_eval_method Codegen.evaluation_conv *}
22a8125d66fa improved handling of !!/==> for eval and normalization
haftmann
parents: 28699
diff changeset
   238
  "solve goal by evaluation"
22a8125d66fa improved handling of !!/==> for eval and normalization
haftmann
parents: 28699
diff changeset
   239
30510
4120fc59dd85 unified type Proof.method and pervasive METHOD combinators;
wenzelm
parents: 29105
diff changeset
   240
method_setup normalization = {* (Method.no_args o SIMPLE_METHOD')
28740
22a8125d66fa improved handling of !!/==> for eval and normalization
haftmann
parents: 28699
diff changeset
   241
  (CONVERSION Nbe.norm_conv THEN' (fn k => TRY (rtac TrueI k)))
24285
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   242
*} "solve goal by normalization"
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   243
28400
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   244
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   245
subsection {* Quickcheck *}
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   246
29105
8f38bf68d42e moved value.ML to src/Tools
haftmann
parents: 28856
diff changeset
   247
setup {*
8f38bf68d42e moved value.ML to src/Tools
haftmann
parents: 28856
diff changeset
   248
  Quickcheck.add_generator ("SML", Codegen.test_term)
8f38bf68d42e moved value.ML to src/Tools
haftmann
parents: 28856
diff changeset
   249
*}
8f38bf68d42e moved value.ML to src/Tools
haftmann
parents: 28856
diff changeset
   250
28400
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   251
quickcheck_params [size = 5, iterations = 50]
89904cfd41c3 polished code generator setup
haftmann
parents: 28346
diff changeset
   252
24285
066bb557570f added Code_Setup
haftmann
parents:
diff changeset
   253
end