src/HOL/SMT/Examples/SMT_Examples.thy
changeset 36082 8f10b0e77d94
parent 36081 70deefb6c093
--- a/src/HOL/SMT/Examples/SMT_Examples.thy	Wed Apr 07 19:48:58 2010 +0200
+++ b/src/HOL/SMT/Examples/SMT_Examples.thy	Wed Apr 07 20:38:11 2010 +0200
@@ -539,7 +539,8 @@
 
 lemma "i \<noteq> i1 \<and> i \<noteq> i2 \<Longrightarrow> (f (i1 := v1, i2 := v2)) i = f i" by smt
 
-lemma "(f g x = (g x \<and> True)) \<or> (f g x = True) \<or> (g x = True)" by smt
+lemma "(f g (x::'a::type) = (g x \<and> True)) \<or> (f g x = True) \<or> (g x = True)"
+  by smt
 
 lemma "id 3 = 3 \<and> id True = True" by (smt id_def)