Nominal/Ex/Ex2.thy
changeset 2309 13f20fe02ff3
parent 2301 8732ff59068b
child 2311 4da5c5c29009
equal deleted inserted replaced
2308:387fcbd33820 2309:13f20fe02ff3
     1 theory Ex2
     1 theory Ex2
     2 imports "../NewParser"
     2 imports "../NewParser"
     3 begin
     3 begin
     4 
     4 
     5 text {* example 2 *}
     5 text {* example 2 *}
     6 declare [[STEPS = 8]]
     6 declare [[STEPS = 9]]
     7 
     7 
     8 atom_decl name
     8 atom_decl name
     9 
     9 
    10 nominal_datatype trm =
    10 nominal_datatype trm =
    11   Var "name"
    11   Var "name"
    21 where
    21 where
    22   "f PN = {}"
    22   "f PN = {}"
    23 | "f (PD x y) = {atom x, atom y}"
    23 | "f (PD x y) = {atom x, atom y}"
    24 | "f (PS x) = {atom x}"
    24 | "f (PS x) = {atom x}"
    25 
    25 
       
    26 
    26 thm fv_trm_raw.simps[no_vars] fv_pat_raw.simps[no_vars] fv_f_raw.simps[no_vars] f_raw.simps[no_vars]
    27 thm fv_trm_raw.simps[no_vars] fv_pat_raw.simps[no_vars] fv_f_raw.simps[no_vars] f_raw.simps[no_vars]
    27 thm alpha_trm_raw_alpha_pat_raw_alpha_f_raw.intros[no_vars]
    28 thm alpha_trm_raw_alpha_pat_raw_alpha_f_raw.intros[no_vars]
    28 
    29 
    29 
    30 
    30 
    31