--- a/src/HOL/Lambda/WeakNorm.thy Tue Aug 11 10:43:43 2009 +0200
+++ b/src/HOL/Lambda/WeakNorm.thy Tue Aug 11 10:46:11 2009 +0200
@@ -444,7 +444,7 @@
"default" ("(error \"default\")")
"default :: 'a \<Rightarrow> 'b::default" ("(fn '_ => error \"default\")")
-(*code_module Norm
+code_module Norm
contains
test = "type_NF"
@@ -509,6 +509,6 @@
val ct2 = @{cterm "%f x. (%x. f x x) ((%x. f x x) ((%x. f x x) ((%x. f x x) ((%x. f x x) ((%x. f x x) x)))))"};
val (dB2, _) = Norm.type_NF (typing_of_term [] dummyf (term_of ct2));
val ct2' = cterm_of @{theory} (term_of_dB [] (#T (rep_cterm ct2)) dB2);
-*}*)
+*}
end
--- a/src/HOL/MicroJava/J/State.thy Tue Aug 11 10:43:43 2009 +0200
+++ b/src/HOL/MicroJava/J/State.thy Tue Aug 11 10:46:11 2009 +0200
@@ -21,10 +21,6 @@
init_vars :: "('a \<times> ty) list => ('a \<rightharpoonup> val)"
"init_vars == map_of o map (\<lambda>(n,T). (n,default_val T))"
-lemma [code] (*manual eta expansion*):
- "init_vars xs = map_of (map (\<lambda>(n, T). (n, default_val T)) xs)"
- by (simp add: init_vars_def)
-
types aheap = "loc \<rightharpoonup> obj" -- {* "@{text heap}" used in a translation below *}
locals = "vname \<rightharpoonup> val" -- "simple state, i.e. variable contents"