Nominal/Ex/Classical.thy
changeset 2436 3885dc2669f9
parent 2434 92dc6cfa3a95
child 2454 9ffee4eb1ae1
equal deleted inserted replaced
2435:3772bb3bd7ce 2436:3885dc2669f9
     4 
     4 
     5 (* example from my Urban's PhD *)
     5 (* example from my Urban's PhD *)
     6 
     6 
     7 atom_decl name
     7 atom_decl name
     8 atom_decl coname
     8 atom_decl coname
     9 
       
    10 declare [[STEPS = 21]]
       
    11 
     9 
    12 nominal_datatype trm =
    10 nominal_datatype trm =
    13    Ax "name" "coname"
    11    Ax "name" "coname"
    14 |  Cut n::"coname" t1::"trm" c::"coname" t2::"trm"             bind n in t1, bind c in t2
    12 |  Cut n::"coname" t1::"trm" c::"coname" t2::"trm"             bind n in t1, bind c in t2
    15 |  AndR c1::"coname" t1::"trm" c2::"coname" t2::"trm" "coname" bind c1 in t1, bind c2 in t2
    13 |  AndR c1::"coname" t1::"trm" c2::"coname" t2::"trm" "coname" bind c1 in t1, bind c2 in t2
    16 |  AndL1 n::"name" t::"trm" "name"                             bind n in t
    14 |  AndL1 n::"name" t::"trm" "name"                             bind n in t
    17 |  AndL2 n::"name" t::"trm" "name"                             bind n in t
    15 |  AndL2 n::"name" t::"trm" "name"                             bind n in t
    18 |  ImpL c::"coname" t1::"trm" n::"name" t2::"trm" "name"       bind c in t1, bind n in t2
    16 |  ImpL c::"coname" t1::"trm" n::"name" t2::"trm" "name"       bind c in t1, bind n in t2
    19 |  ImpR c::"coname" n::"name" t::"trm" "coname"                bind n c in t
    17 |  ImpR c::"coname" n::"name" t::"trm" "coname"                bind n c in t
    20 
    18 
    21 thm distinct
    19 thm trm.distinct
    22 thm induct
    20 thm trm.induct
    23 thm exhaust
    21 thm trm.exhaust
    24 thm fv_defs
    22 thm trm.fv_defs
    25 thm bn_defs
    23 thm trm.bn_defs
    26 thm perm_simps
    24 thm trm.perm_simps
    27 thm eq_iff
    25 thm trm.eq_iff
    28 thm fv_bn_eqvt
    26 thm trm.fv_bn_eqvt
    29 thm size_eqvt
    27 thm trm.size_eqvt
    30 
    28 
    31 
    29 
    32 
    30 
    33 
    31 
    34 end
    32 end