Nominal/Ex/Ex2.thy
changeset 2050 8bd75f2fd7b0
parent 2030 43d7612f1429
child 2082 0854af516f14
equal deleted inserted replaced
2049:38bbccdf9ff9 2050:8bd75f2fd7b0
     1 theory Ex2
     1 theory Ex2
     2 imports "../Parser"
     2 imports "../NewParser"
     3 begin
     3 begin
     4 
     4 
     5 text {* example 2 *}
     5 text {* example 2 *}
     6 
     6 
     7 atom_decl name
     7 atom_decl name
     8 
     8 
     9 ML {* val _ = recursive := false  *}
     9 nominal_datatype trm =
    10 nominal_datatype trm' =
       
    11   Var "name"
    10   Var "name"
    12 | App "trm'" "trm'"
    11 | App "trm" "trm"
    13 | Lam x::"name" t::"trm'"          bind x in t
    12 | Lam x::"name" t::"trm"          bind_set x in t
    14 | Let p::"pat'" "trm'" t::"trm'"   bind "f p" in t
    13 | Let p::"pat" "trm" t::"trm"   bind_set "f p" in t
    15 and pat' =
    14 and pat =
    16   PN
    15   PN
    17 | PS "name"
    16 | PS "name"
    18 | PD "name" "name"
    17 | PD "name" "name"
    19 binder
    18 binder
    20   f::"pat' \<Rightarrow> atom set"
    19   f::"pat \<Rightarrow> atom set"
    21 where
    20 where
    22   "f PN = {}"
    21   "f PN = {}"
    23 | "f (PD x y) = {atom x, atom y}"
    22 | "f (PD x y) = {atom x, atom y}"
    24 | "f (PS x) = {atom x}"
    23 | "f (PS x) = {atom x}"
    25 
    24 thm trm_pat.bn
    26 thm trm'_pat'.fv
    25 thm trm_pat.perm
    27 thm trm'_pat'.eq_iff
    26 thm trm_pat.induct
    28 thm trm'_pat'.bn
    27 thm trm_pat.distinct
    29 thm trm'_pat'.perm
    28 thm trm_pat.fv[simplified trm_pat.supp(1-2)]
    30 thm trm'_pat'.induct
       
    31 thm trm'_pat'.distinct
       
    32 thm trm'_pat'.fv[simplified trm'_pat'.supp]
       
    33 
    29 
    34 end
    30 end
    35 
    31 
    36 
    32 
    37 
    33