src/HOL/ex/Quickcheck_Narrowing_Examples.thy
changeset 42184 1d4fae76ba5e
parent 42024 51df23535105
child 43239 42f82fda796b
equal deleted inserted replaced
42178:b992c8e6394b 42184:1d4fae76ba5e
    30 oops
    30 oops
    31 
    31 
    32 lemma "rev xs = xs"
    32 lemma "rev xs = xs"
    33   quickcheck[tester = narrowing, finite_types = false, default_type = int, expect = counterexample]
    33   quickcheck[tester = narrowing, finite_types = false, default_type = int, expect = counterexample]
    34 oops
    34 oops
    35 
    35 (*
    36 lemma "rev xs = xs"
    36 lemma "rev xs = xs"
    37   quickcheck[tester = narrowing, finite_types = true, expect = counterexample]
    37   quickcheck[tester = narrowing, finite_types = true, expect = counterexample]
    38 oops
    38 oops
    39 
    39 *)
    40 subsection {* Simple examples with functions *}
    40 subsection {* Simple examples with functions *}
    41 
    41 
    42 declare [[quickcheck_finite_functions]]
    42 declare [[quickcheck_finite_functions]]
    43 
    43 (*
    44 lemma "map f xs = map g xs"
    44 lemma "map f xs = map g xs"
    45   quickcheck[tester = narrowing, finite_types = true, expect = counterexample]
    45   quickcheck[tester = narrowing, finite_types = true, expect = counterexample]
    46 oops
    46 oops
    47 
    47 
    48 lemma "map f xs = map f ys ==> xs = ys"
    48 lemma "map f xs = map f ys ==> xs = ys"
    55 oops
    55 oops
    56 
    56 
    57 lemma "map f xs = F f xs"
    57 lemma "map f xs = F f xs"
    58   quickcheck[tester = narrowing, finite_types = true, expect = counterexample]
    58   quickcheck[tester = narrowing, finite_types = true, expect = counterexample]
    59 oops
    59 oops
    60 
    60 *)
    61 lemma "map f xs = F f xs"
    61 lemma "map f xs = F f xs"
    62   quickcheck[tester = narrowing, finite_types = false, expect = counterexample]
    62   quickcheck[tester = narrowing, finite_types = false, expect = counterexample]
    63 oops
    63 oops
    64 
    64 
    65 (* requires some work...
    65 (* requires some work...
   161 
   161 
   162 subsubsection {* Invalid Lemma due to typo in lbal *}
   162 subsubsection {* Invalid Lemma due to typo in lbal *}
   163 
   163 
   164 lemma is_ord_l_bal:
   164 lemma is_ord_l_bal:
   165  "\<lbrakk> is_ord(MKT (x :: nat) l r h); height l = height r + 2 \<rbrakk> \<Longrightarrow> is_ord(l_bal(x,l,r))"
   165  "\<lbrakk> is_ord(MKT (x :: nat) l r h); height l = height r + 2 \<rbrakk> \<Longrightarrow> is_ord(l_bal(x,l,r))"
   166 quickcheck[tester = narrowing, finite_types = false, default_type = nat, size = 6, timeout = 100, expect = counterexample]
   166 quickcheck[tester = narrowing, finite_types = false, default_type = nat, size = 6, timeout = 1000, expect = counterexample]
   167 oops
   167 oops
   168 
   168 
   169 
   169 
   170 end
   170 end