Nominal/Ex/SingleLet.thy
changeset 2213 231a20534950
parent 2200 31f1ec832d39
child 2284 7b83e1c8ba2e
child 2312 ad03df7e8056
child 2326 b51532dd5689
equal deleted inserted replaced
2212:79cebcc230d6 2213:231a20534950
     6 
     6 
     7 nominal_datatype trm =
     7 nominal_datatype trm =
     8   Var "name"
     8   Var "name"
     9 | App "trm" "trm"
     9 | App "trm" "trm"
    10 | Lam x::"name" t::"trm"  bind_set x in t
    10 | Lam x::"name" t::"trm"  bind_set x in t
    11 | Let a::"assg" t::"trm"  bind_set "bn a" in t
    11 | Let a::"assg" t1::"trm" t2::"trm" bind_set "bn a" in t1 t2
       
    12  
    12 and assg =
    13 and assg =
    13   As "name" "trm"
    14   As "name" "trm"
    14 binder
    15 binder
    15   bn::"assg \<Rightarrow> atom set"
    16   bn::"assg \<Rightarrow> atom set"
    16 where
    17 where