src/HOL/Nitpick_Examples/Typedef_Nits.thy
changeset 35075 6fd1052fe463
parent 35073 cc19e2aef17e
child 35284 9edc2bd6d2bd
equal deleted inserted replaced
35074:c1dac8ace020 35075:6fd1052fe463
     9 
     9 
    10 theory Typedef_Nits
    10 theory Typedef_Nits
    11 imports Main Rational
    11 imports Main Rational
    12 begin
    12 begin
    13 
    13 
    14 nitpick_params [card = 1\<midarrow>4, timeout = 30 s]
    14 nitpick_params [sat_solver = MiniSat_JNI, max_threads = 1, timeout = 60 s,
       
    15                 card = 1\<midarrow>4]
    15 
    16 
    16 typedef three = "{0\<Colon>nat, 1, 2}"
    17 typedef three = "{0\<Colon>nat, 1, 2}"
    17 by blast
    18 by blast
    18 
    19 
    19 definition A :: three where "A \<equiv> Abs_three 0"
    20 definition A :: three where "A \<equiv> Abs_three 0"