(* * Copyright 2014, NICTA * * This software may be distributed and modified according to the terms of * the BSD 2-Clause license. Note that NO WARRANTY is provided. * See "LICENSE_BSD2.txt" for details. * * @TAG(NICTA_BSD) *) theory Crunch_Test imports Crunch Crunch_Test_Qualified GenericLib begin text {* Test cases for crunch *} definition "crunch_foo1 (x :: nat) \ do modify (op + x); modify (op + x) od" definition "crunch_foo2 \ do crunch_foo1 12; crunch_foo1 13 od" crunch (empty_fail) empty_fail: crunch_foo2 (ignore: modify bind) crunch_ignore (add: crunch_foo1) crunch gt: crunch_foo2 "\x. x > y" (ignore: modify bind ignore_del: crunch_foo1) crunch_ignore (del: crunch_foo1) definition "crunch_always_true (x :: nat) \ \y :: nat. True" lemma crunch_foo1_at_2: "\crunch_always_true 2 and crunch_always_true 3\ crunch_foo1 x \\rv. crunch_always_true 2\" by (simp add: crunch_always_true_def, wp) lemma crunch_foo1_at_3[wp]: "\crunch_always_true 3\ crunch_foo1 x \\rv. crunch_always_true 3\" by (simp add: crunch_always_true_def, wp) lemma crunch_foo1_no_fail: "no_fail (crunch_always_true 2 and crunch_always_true 3) (crunch_foo1 x)" apply (simp add:crunch_always_true_def crunch_foo1_def) apply (rule no_fail_pre) apply (wp, simp) done crunch (no_fail) no_fail: crunch_foo2 (ignore: modify bind wp: crunch_foo1_at_2) crunch (valid) at_2: crunch_foo2 "crunch_always_true 2" (ignore: modify bind wp: crunch_foo1_at_2) fun crunch_foo3 :: "nat => nat => 'a => (nat,unit) nondet_monad" where "crunch_foo3 0 x _ = crunch_foo1 x" | "crunch_foo3 (Suc n) x y = crunch_foo3 n x y" crunch gt2: crunch_foo3 "\x. x > y" (ignore: modify bind) crunch (empty_fail) empty_fail2: crunch_foo3 (ignore: modify bind) class foo_class = fixes stuff :: 'a begin fun crunch_foo4 :: "nat => nat => 'a => (nat,unit) nondet_monad" where "crunch_foo4 0 x _ = crunch_foo1 x" | "crunch_foo4 (Suc n) x y = crunch_foo4 n x y" definition "crunch_foo5 x y \ crunch_foo1 x" end lemma crunch_foo4_alt: "crunch_foo4 n x y \ crunch_foo1 x" apply (induct n) apply simp+ done (* FIXME: breaks, does not find induct_instance, wrong number of params crunch gt3: crunch_foo4 "\x. x > y" (ignore: modify bind) *) crunch (no_fail) no_fail2: crunch_foo4 (unfold: "(crunch_foo4, crunch_foo4_alt)" ignore: modify bind) crunch gt3: crunch_foo4 "\x. x > y" (unfold: "(crunch_foo4, crunch_foo4_alt)" ignore: modify bind) crunch gt4: crunch_foo5 "\x. x > y" (ignore: modify bind) (* Test cases for crunch in locales *) crunch_ignore (add: bind) definition "crunch_foo6 \ return () >>= (\_. return ())" locale test_locale = fixes fixed_return_unit :: "(unit, unit) nondet_monad" begin definition "crunch_foo7 \ return () >>= (\_. return ())" (* crunch works on a global constant within a locale *) crunch test[wp]: crunch_foo6 P (* crunch works on a locale constant *) crunch test[wp]: crunch_foo7 P definition "crunch_foo8 \ fixed_return_unit >>= (\_. fixed_return_unit)" definition "crunch_foo9 (x :: nat) \ do modify (op + x); modify (op + x) od" crunch test[wp]: crunch_foo9 "\x. x > y" (ignore: modify) (* crunch_ignore works within a locale *) crunch_ignore (add: modify) crunch test[wp]: crunch_foo9 "\x. x > y" end interpretation test_locale "return ()" . (* interpretation promotes the wp attribute from the locale *) lemma "\Q\ crunch_foo7 \\_. Q\" by wp (* crunch still works on an interpreted locale constant *) crunch test2: crunch_foo7 P locale test_sublocale sublocale test_sublocale < test_locale "return ()" . context test_sublocale begin (* crunch works on a locale constant with a fixed locale parameter *) crunch test[wp]: crunch_foo8 P end (* check that qualified names are handled properly. *) consts foo_const :: "(unit, unit) nondet_monad" defs foo_const_def: "foo_const \ Crunch_Test_Qualified.foo_const" crunch test: foo_const P end