author | haftmann |
Tue, 28 Oct 2008 12:29:57 +0100 | |
changeset 28702 | 4867f2fa0e48 |
parent 28701 | ca5840b1f7b3 |
child 28703 | aef727ef30e5 |
--- a/src/HOL/Library/Coinductive_List.thy Tue Oct 28 12:28:14 2008 +0100 +++ b/src/HOL/Library/Coinductive_List.thy Tue Oct 28 12:29:57 2008 +0100 @@ -216,7 +216,7 @@ CONS_type Abs_llist_inverse Rep_llist Rep_llist_inverse inj_Leaf) lemma llist_case_cert: - includes meta_conjunction_syntax + fixes meta_conjunction :: "prop => prop => prop" (infixr "&&" 2) assumes "CASE \<equiv> llist_case c d" shows "(CASE LNil \<equiv> c) && (CASE (LCons M N) \<equiv> d M N)" using assms by simp_all