# HG changeset patch # User bulwahn # Date 1299853273 -3600 # Node ID af501c14d8f0f1b79464c3fc36543178b2542384 # Parent 1e008cc4883a758a0a9af95fdbac7158506a8cd6 renaming LSC_Examples theory to Quickcheck_Narrowing diff -r 1e008cc4883a -r af501c14d8f0 src/HOL/ex/LSC_Examples.thy --- a/src/HOL/ex/LSC_Examples.thy Fri Mar 11 15:21:13 2011 +0100 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,141 +0,0 @@ -(* Title: HOL/ex/LSC_Examples.thy - Author: Lukas Bulwahn - Copyright 2011 TU Muenchen -*) - -header {* Examples for invoking lazysmallcheck (LSC) *} - -theory LSC_Examples -imports "~~/src/HOL/Library/LSC" -begin - -subsection {* Simple list examples *} - -lemma "rev xs = xs" -quickcheck[tester = lazy_exhaustive, finite_types = false, default_type = nat, expect = counterexample] -oops - -text {* Example fails due to some strange thing... *} -(* -lemma "rev xs = xs" -quickcheck[tester = lazy_exhaustive, finite_types = true] -oops -*) - -subsection {* AVL Trees *} - -datatype 'a tree = ET | MKT 'a "'a tree" "'a tree" nat - -primrec set_of :: "'a tree \ 'a set" -where -"set_of ET = {}" | -"set_of (MKT n l r h) = insert n (set_of l \ set_of r)" - -primrec height :: "'a tree \ nat" -where -"height ET = 0" | -"height (MKT x l r h) = max (height l) (height r) + 1" - -primrec avl :: "'a tree \ bool" -where -"avl ET = True" | -"avl (MKT x l r h) = - ((height l = height r \ height l = 1+height r \ height r = 1+height l) \ - h = max (height l) (height r) + 1 \ avl l \ avl r)" - -primrec is_ord :: "('a::order) tree \ bool" -where -"is_ord ET = True" | -"is_ord (MKT n l r h) = - ((\n' \ set_of l. n' < n) \ (\n' \ set_of r. n < n') \ is_ord l \ is_ord r)" - -primrec is_in :: "('a::order) \ 'a tree \ bool" -where - "is_in k ET = False" | - "is_in k (MKT n l r h) = (if k = n then True else - if k < n then (is_in k l) - else (is_in k r))" - -primrec ht :: "'a tree \ nat" -where -"ht ET = 0" | -"ht (MKT x l r h) = h" - -definition - mkt :: "'a \ 'a tree \ 'a tree \ 'a tree" where -"mkt x l r = MKT x l r (max (ht l) (ht r) + 1)" - -(* replaced MKT lrn lrl lrr by MKT lrr lrl *) -fun l_bal where -"l_bal(n, MKT ln ll lr h, r) = - (if ht ll < ht lr - then case lr of ET \ ET (* impossible *) - | MKT lrn lrr lrl lrh \ - mkt lrn (mkt ln ll lrl) (mkt n lrr r) - else mkt ln ll (mkt n lr r))" - -fun r_bal where -"r_bal(n, l, MKT rn rl rr h) = - (if ht rl > ht rr - then case rl of ET \ ET (* impossible *) - | MKT rln rll rlr h \ mkt rln (mkt n l rll) (mkt rn rlr rr) - else mkt rn (mkt n l rl) rr)" - -primrec insrt :: "'a::order \ 'a tree \ 'a tree" -where -"insrt x ET = MKT x ET ET 1" | -"insrt x (MKT n l r h) = - (if x=n - then MKT n l r h - else if x is_ord(MKT (x :: nat) l r h); height l = height r + 2 \ \ is_ord(l_bal(x,l,r))" -quickcheck[tester = lazy_exhaustive, finite_types = false, default_type = nat, size = 1, timeout = 100, expect = counterexample] -oops - - -end diff -r 1e008cc4883a -r af501c14d8f0 src/HOL/ex/Quickcheck_Narrowing.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/ex/Quickcheck_Narrowing.thy Fri Mar 11 15:21:13 2011 +0100 @@ -0,0 +1,141 @@ +(* Title: HOL/ex/LSC_Examples.thy + Author: Lukas Bulwahn + Copyright 2011 TU Muenchen +*) + +header {* Examples for invoking lazysmallcheck (LSC) *} + +theory LSC_Examples +imports "~~/src/HOL/Library/LSC" +begin + +subsection {* Simple list examples *} + +lemma "rev xs = xs" +quickcheck[tester = lazy_exhaustive, finite_types = false, default_type = nat, expect = counterexample] +oops + +text {* Example fails due to some strange thing... *} +(* +lemma "rev xs = xs" +quickcheck[tester = lazy_exhaustive, finite_types = true] +oops +*) + +subsection {* AVL Trees *} + +datatype 'a tree = ET | MKT 'a "'a tree" "'a tree" nat + +primrec set_of :: "'a tree \ 'a set" +where +"set_of ET = {}" | +"set_of (MKT n l r h) = insert n (set_of l \ set_of r)" + +primrec height :: "'a tree \ nat" +where +"height ET = 0" | +"height (MKT x l r h) = max (height l) (height r) + 1" + +primrec avl :: "'a tree \ bool" +where +"avl ET = True" | +"avl (MKT x l r h) = + ((height l = height r \ height l = 1+height r \ height r = 1+height l) \ + h = max (height l) (height r) + 1 \ avl l \ avl r)" + +primrec is_ord :: "('a::order) tree \ bool" +where +"is_ord ET = True" | +"is_ord (MKT n l r h) = + ((\n' \ set_of l. n' < n) \ (\n' \ set_of r. n < n') \ is_ord l \ is_ord r)" + +primrec is_in :: "('a::order) \ 'a tree \ bool" +where + "is_in k ET = False" | + "is_in k (MKT n l r h) = (if k = n then True else + if k < n then (is_in k l) + else (is_in k r))" + +primrec ht :: "'a tree \ nat" +where +"ht ET = 0" | +"ht (MKT x l r h) = h" + +definition + mkt :: "'a \ 'a tree \ 'a tree \ 'a tree" where +"mkt x l r = MKT x l r (max (ht l) (ht r) + 1)" + +(* replaced MKT lrn lrl lrr by MKT lrr lrl *) +fun l_bal where +"l_bal(n, MKT ln ll lr h, r) = + (if ht ll < ht lr + then case lr of ET \ ET (* impossible *) + | MKT lrn lrr lrl lrh \ + mkt lrn (mkt ln ll lrl) (mkt n lrr r) + else mkt ln ll (mkt n lr r))" + +fun r_bal where +"r_bal(n, l, MKT rn rl rr h) = + (if ht rl > ht rr + then case rl of ET \ ET (* impossible *) + | MKT rln rll rlr h \ mkt rln (mkt n l rll) (mkt rn rlr rr) + else mkt rn (mkt n l rl) rr)" + +primrec insrt :: "'a::order \ 'a tree \ 'a tree" +where +"insrt x ET = MKT x ET ET 1" | +"insrt x (MKT n l r h) = + (if x=n + then MKT n l r h + else if x is_ord(MKT (x :: nat) l r h); height l = height r + 2 \ \ is_ord(l_bal(x,l,r))" +quickcheck[tester = lazy_exhaustive, finite_types = false, default_type = nat, size = 1, timeout = 100, expect = counterexample] +oops + + +end