src/HOL/String.thy
changeset 39274 b17ffa965223
parent 39272 0b61951d2682
child 39302 d7728f65b353
--- a/src/HOL/String.thy	Fri Sep 10 10:59:07 2010 +0200
+++ b/src/HOL/String.thy	Fri Sep 10 10:59:09 2010 +0200
@@ -157,8 +157,6 @@
 typedef (open) literal = "UNIV :: string \<Rightarrow> bool"
   morphisms explode STR ..
 
-code_datatype STR
-
 instantiation literal :: size
 begin
 
@@ -183,6 +181,10 @@
 
 end
 
+lemma STR_inject' [simp]: "(STR xs = STR ys) = (xs = ys)"
+by(simp add: STR_inject)
+
+
 subsection {* Code generator *}
 
 use "Tools/string_code.ML"