src/HOL/Nitpick_Examples/Typedef_Nits.thy
changeset 41278 8e1cde88aae6
parent 40590 b994d855dbd2
child 42142 d24a93025feb
equal deleted inserted replaced
41277:1369c27c6966 41278:8e1cde88aae6
     9 
     9 
    10 theory Typedef_Nits
    10 theory Typedef_Nits
    11 imports Complex_Main
    11 imports Complex_Main
    12 begin
    12 begin
    13 
    13 
    14 nitpick_params [card = 1\<midarrow>4, sat_solver = MiniSat_JNI, max_threads = 1,
    14 nitpick_params [verbose, card = 1\<midarrow>4, sat_solver = MiniSat_JNI, max_threads = 1,
    15                 timeout = 60]
    15                 timeout = 60]
    16 
    16 
    17 typedef three = "{0\<Colon>nat, 1, 2}"
    17 typedef three = "{0\<Colon>nat, 1, 2}"
    18 by blast
    18 by blast
    19 
    19