src/HOL/Nitpick_Examples/Special_Nits.thy
changeset 34083 652719832159
parent 33737 e441fede163d
child 34124 c4628a1dcf75
equal deleted inserted replaced
34082:61b7aa37f4b7 34083:652719832159
     9 
     9 
    10 theory Special_Nits
    10 theory Special_Nits
    11 imports Main
    11 imports Main
    12 begin
    12 begin
    13 
    13 
    14 nitpick_params [card = 4, debug, show_consts, timeout = 10 s]
    14 nitpick_params [sat_solver = MiniSatJNI, max_threads = 1, timeout = 60 s,
       
    15                 card = 4]
    15 
    16 
    16 fun f1 :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" where
    17 fun f1 :: "nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat \<Rightarrow> nat" where
    17 "f1 a b c d e = a + b + c + d + e"
    18 "f1 a b c d e = a + b + c + d + e"
    18 
    19 
    19 lemma "f1 0 0 0 0 0 = f1 0 0 0 0 (1 - 1)"
    20 lemma "f1 0 0 0 0 0 = f1 0 0 0 0 (1 - 1)"