theory Testimports "Parser"begin(* This file contains only the examples that are not supposed to work yet. *)atom_decl nameML {* val _ = recursive := false *}(* example 7 from Peter Sewell's bestiary *)(* dest_Const raisednominal_datatype exp7 = EVar' name| EUnit'| EPair' exp7 exp7| ELetRec' l::lrbs e::exp7 bind "b7s l" in e, bind "b7s l" in land lrb = Assign' name exp7and lrbs = Single' lrb| More' lrb lrbsbinder b7 :: "lrb \<Rightarrow> atom set" and b7s :: "lrbs \<Rightarrow> atom set"where "b7 (Assign x e) = {atom x}"| "b7s (Single a) = b7 a"| "b7s (More a as) = (b7 a) \<union> (b7s as)"thm alpha_exp7_raw_alpha_lrb_raw_alpha_lrbs_raw.intros*)(* example 8 from Peter Sewell's bestiary *)(**** fv_bn: recursive argument, but wrong datatype.*** At command "nominal_datatype".nominal_datatype exp8 = EVar' name| EUnit'| EPair' exp8 exp8| ELetRec' l::lrbs8 e::exp8 bind "b_lrbs8 l" in e, bind "b_lrbs8 l" in land fnclause = K' x::name p::pat8 e::exp8 bind "b_pat p" in eand fnclauses = S' fnclause| ORs' fnclause fnclausesand lrb8 = Clause' fnclausesand lrbs8 = Single' lrb8| More' lrb8 lrbs8and pat8 = PVar'' name| PUnit''| PPair'' pat8 pat8binder b_lrbs8 :: "lrbs8 \<Rightarrow> atom set" and b_pat :: "pat8 \<Rightarrow> atom set" and b_fnclauses :: "fnclauses \<Rightarrow> atom set" and b_fnclause :: "fnclause \<Rightarrow> atom set" and b_lrb8 :: "lrb8 \<Rightarrow> atom set"where "b_lrbs8 (Single' l) = b_lrb8 l"| "b_lrbs8 (More' l ls) = b_lrb8 l \<union> b_lrbs8 ls"| "b_pat (PVar'' x) = {atom x}"| "b_pat (PUnit'') = {}"| "b_pat (PPair'' p1 p2) = b_pat p1 \<union> b_pat p2"| "b_fnclauses (S' fc) = (b_fnclause fc)"| "b_fnclauses (ORs' fc fcs) = (b_fnclause fc) \<union> (b_fnclauses fcs)"| "b_lrb8 (Clause' fcs) = (b_fnclauses fcs)"| "b_fnclause (K' x pat exp8) = {atom x}"thm alpha_exp8_raw_alpha_fnclause_raw_alpha_fnclauses_raw_alpha_lrb8_raw_alpha_lrbs8_raw_alpha_pat8_raw.intros*)(* example 4 from Terms.thy *)(* fv_eqvt does not work, we need to repaire defined permute functions defined fv and defined alpha... *)(* lists-datastructure does not work yetnominal_datatype trm4 = Vr4 "name"| Ap4 "trm4" "trm4 list"| Lm4 x::"name" t::"trm4" bind x in tthm alpha_trm4_raw_alpha_trm4_raw_list.intros[no_vars]thm fv_trm4_raw_fv_trm4_raw_list.simps[no_vars]*)(* example 8 from Terms.thy *)(* Binding in a term under a bn, needs to fail *)(*nominal_datatype foo8 = Foo0 "name"| Foo1 b::"bar8" f::"foo8" bind "bv8 b" in f --"check fo error if this is called foo"and bar8 = Bar0 "name"| Bar1 "name" s::"name" b::"bar8" bind s in bbinder bv8where "bv8 (Bar0 x) = {}"| "bv8 (Bar1 v x b) = {atom v}"*)(* example 9 from Peter Sewell's bestiary *)(* run out of steam at the moment *)end