Nominal/Ex/SingleLet.thy
changeset 2338 e1764a73c292
parent 2336 f2d545b77b31
child 2339 1e0b3992189c
equal deleted inserted replaced
2337:b151399bd2c3 2338:e1764a73c292
     2 imports "../NewParser"
     2 imports "../NewParser"
     3 begin
     3 begin
     4 
     4 
     5 atom_decl name
     5 atom_decl name
     6 
     6 
     7 declare [[STEPS = 15]]
     7 declare [[STEPS = 16]]
     8 
     8 
     9 nominal_datatype trm  =
     9 nominal_datatype trm  =
    10   Var "name"
    10   Var "name"
    11 | App "trm" "trm"
    11 | App "trm" "trm"
    12 | Lam x::"name" t::"trm"  bind_set x in t
    12 | Lam x::"name" t::"trm"  bind_set x in t
    18   As "name" x::"name" t::"trm" bind x in t
    18   As "name" x::"name" t::"trm" bind x in t
    19 binder
    19 binder
    20   bn::"assg \<Rightarrow> atom set"
    20   bn::"assg \<Rightarrow> atom set"
    21 where
    21 where
    22   "bn (As x y t) = {atom x}"
    22   "bn (As x y t) = {atom x}"
       
    23 
       
    24 term Var 
       
    25 term App
       
    26 term Baz
       
    27 
    23 
    28 
    24 typ trm
    29 typ trm
    25 typ assg
    30 typ assg
    26 
    31 
    27 thm trm_assg.fv
    32 thm trm_assg.fv