src/HOL/Nitpick_Examples/Refute_Nits.thy
changeset 38185 b51677438b3a
parent 36319 8feb2c4bef1a
child 40341 03156257040f
equal deleted inserted replaced
38184:6a221fffbc8a 38185:b51677438b3a
     9 
     9 
    10 theory Refute_Nits
    10 theory Refute_Nits
    11 imports Main
    11 imports Main
    12 begin
    12 begin
    13 
    13 
    14 nitpick_params [max_potential = 0, sat_solver = MiniSat_JNI, max_threads = 1,
    14 nitpick_params [card = 1\<midarrow>6, max_potential = 0,
    15                 timeout = 60 s]
    15                 sat_solver = MiniSat_JNI, max_threads = 1, timeout = 60 s]
    16 
    16 
    17 lemma "P \<and> Q"
    17 lemma "P \<and> Q"
    18 apply (rule conjI)
    18 apply (rule conjI)
    19 nitpick [expect = genuine] 1
    19 nitpick [expect = genuine] 1
    20 nitpick [expect = genuine] 2
    20 nitpick [expect = genuine] 2