src/HOL/Nitpick_Examples/Special_Nits.thy
changeset 40341 03156257040f
parent 36389 8228b3a4a2ba
child 41278 8e1cde88aae6
equal deleted inserted replaced
40340:d1c14898fd04 40341:03156257040f
    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, sat_solver = MiniSat_JNI, max_threads = 1,
    14 nitpick_params [card = 4, sat_solver = MiniSat_JNI, max_threads = 1,
    15                 timeout = 60 s]
    15                 timeout = 60]
    16 
    16 
    17 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
    18 "f1 a b c d e = a + b + c + d + e"
    18 "f1 a b c d e = a + b + c + d + e"
    19 
    19 
    20 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)"