equal
deleted
inserted
replaced
219 "is_ord (MKT n l r h) = ((\<forall>n' \<in> set (set_of' l). n' < n) \<and> (\<forall>n' \<in> set (set_of' r). n < n') \<and> is_ord l \<and> is_ord r)" |
219 "is_ord (MKT n l r h) = ((\<forall>n' \<in> set (set_of' l). n' < n) \<and> (\<forall>n' \<in> set (set_of' r). n < n') \<and> is_ord l \<and> is_ord r)" |
220 by (simp add: set_of') |
220 by (simp add: set_of') |
221 |
221 |
222 declare is_ord.simps(1)[code] is_ord_mkt[code] |
222 declare is_ord.simps(1)[code] is_ord_mkt[code] |
223 |
223 |
224 subsubsection \<open>Invalid Lemma due to typo in @{const l_bal}\<close> |
224 subsubsection \<open>Invalid Lemma due to typo in \<^const>\<open>l_bal\<close>\<close> |
225 |
225 |
226 lemma is_ord_l_bal: |
226 lemma is_ord_l_bal: |
227 "\<lbrakk> is_ord(MKT (x :: nat) l r h); height l = height r + 2 \<rbrakk> \<Longrightarrow> is_ord(l_bal(x,l,r))" |
227 "\<lbrakk> is_ord(MKT (x :: nat) l r h); height l = height r + 2 \<rbrakk> \<Longrightarrow> is_ord(l_bal(x,l,r))" |
228 quickcheck[tester = narrowing, finite_types = false, default_type = nat, size = 6, expect = counterexample] |
228 quickcheck[tester = narrowing, finite_types = false, default_type = nat, size = 6, expect = counterexample] |
229 oops |
229 oops |