src/HOL/HOL.thy
changeset 42420 8a09dfeb2cec
parent 42361 23f352990944
child 42422 6a2837ddde4b
--- a/src/HOL/HOL.thy	Tue Apr 19 14:52:22 2011 +0200
+++ b/src/HOL/HOL.thy	Wed Apr 20 07:44:23 2011 +0200
@@ -1915,18 +1915,22 @@
   (Haskell "Bool")
   (Scala "Boolean")
 
-code_const True and False and Not and HOL.conj and HOL.disj and If
+code_const True and False and Not and HOL.conj and HOL.disj and HOL.implies and If 
   (SML "true" and "false" and "not"
     and infixl 1 "andalso" and infixl 0 "orelse"
+    and "!(if (_)/ then (_)/ else true)"
     and "!(if (_)/ then (_)/ else (_))")
   (OCaml "true" and "false" and "not"
     and infixl 3 "&&" and infixl 2 "||"
+    and "!(if (_)/ then (_)/ else true)"
     and "!(if (_)/ then (_)/ else (_))")
   (Haskell "True" and "False" and "not"
     and infixr 3 "&&" and infixr 2 "||"
+    and "!(if (_)/ then (_)/ else True)"
     and "!(if (_)/ then (_)/ else (_))")
   (Scala "true" and "false" and "'! _"
     and infixl 3 "&&" and infixl 1 "||"
+    and "!(if ((_))/ (_)/ else true)"
     and "!(if ((_))/ (_)/ else (_))")
 
 code_reserved SML