changed parser so that the binding mode is indicated as "bind (list)", "bind (set)" or "bind (res)"; if only "bind" is given, then bind (list) is assumed as default
theory SingleLetimports "../NewParser"beginatom_decl namedeclare [[STEPS = 20]]nominal_datatype trm = Var "name"| App "trm" "trm"| Lam x::"name" t::"trm" bind x in t| Let a::"assg" t::"trm" bind (set) "bn a" in t| Foo x::"name" y::"name" t::"trm" t1::"trm" t2::"trm" bind (set) x in y t t1 t2| Bar x::"name" y::"name" t::"trm" bind y x in t x y| Baz x::"name" t1::"trm" t2::"trm" bind x in t1, bind x in t2 and assg = As "name" x::"name" t::"trm" bind x in tbinder bn::"assg \<Rightarrow> atom set"where "bn (As x y t) = {atom x}"ML {* Function.prove_termination *}text {* can lift *}thm distinctthm trm_raw_assg_raw.inductsthm trm_raw.exhaustthm assg_raw.exhaustthm fv_defsthm perm_simpsthm perm_lawsthm trm_raw_assg_raw.size(9 - 16)thm eq_iffthm eq_iff_simpsthm bn_defsthm fv_eqvtthm bn_eqvtthm size_eqvtML {* val thms_d = map (lift_thm [@{typ trm}, @{typ assg}] @{context}) @{thms distinct}*}ML {* val thms_i = map (lift_thm [@{typ trm}, @{typ assg}] @{context}) @{thms trm_raw_assg_raw.inducts}*}ML {* val thms_i = map (lift_thm [@{typ trm}, @{typ assg}] @{context}) @{thms trm_raw.exhaust}*}ML {* val thms_i = map (lift_thm [@{typ trm}, @{typ assg}] @{context}) @{thms assg_raw.exhaust}*}ML {* val thms_f = map (lift_thm [@{typ trm}, @{typ assg}] @{context}) @{thms fv_defs}*}ML {* val thms_i = map (lift_thm [@{typ trm}, @{typ assg}] @{context}) @{thms trm_raw_assg_raw.size(9 - 16)}*}ML {* val thms_p = map (lift_thm [@{typ trm}, @{typ assg}] @{context}) @{thms perm_simps}*}ML {* val thms_f = map (lift_thm [@{typ trm}, @{typ assg}] @{context}) @{thms perm_laws}*}ML {* val thms_e = map (lift_thm [@{typ trm}, @{typ assg}] @{context}) @{thms eq_iff[unfolded alphas permute_prod.simps prod_fv.simps prod_alpha_def prod_rel.simps prod.cases]}*}ML {* val thms_e = map (lift_thm [@{typ trm}, @{typ assg}] @{context}) @{thms eq_iff_simps[unfolded alphas permute_prod.simps prod_fv.simps prod_alpha_def prod_rel.simps prod.cases]}*}ML {* val thms_f = map (lift_thm [@{typ trm}, @{typ assg}] @{context}) @{thms bn_defs}*}ML {* val thms_f = map (lift_thm [@{typ trm}, @{typ assg}] @{context}) @{thms bn_eqvt}*}ML {* val thms_f = map (lift_thm [@{typ trm}, @{typ assg}] @{context}) @{thms fv_eqvt}*}ML {* val thms_f = map (lift_thm [@{typ trm}, @{typ assg}] @{context}) @{thms size_eqvt}*}lemma supp_fv: "supp t = fv_trm t" "supp b = fv_bn b"apply(induct t and b rule: i1)apply(simp_all add: f1)apply(simp_all add: supp_def)apply(simp_all add: b1)sorryconsts perm_bn_trm :: "perm \<Rightarrow> trm \<Rightarrow> trm"consts perm_bn_assg :: "perm \<Rightarrow> assg \<Rightarrow> assg"lemma y: "perm_bn_trm p (Var x) = (Var x)" "perm_bn_trm p (App t1 t2) = (App t1 t2)" "perm_bn_trm p ("typ trmtyp assgthm trm_assg.fvthm trm_assg.suppthm trm_assg.eq_iffthm trm_assg.bnthm trm_assg.permthm trm_assg.inductthm trm_assg.inductsthm trm_assg.distinctML {* Sign.of_sort @{theory} (@{typ trm}, @{sort fs}) *}(* TEMPORARYthm trm_assg.fv[simplified trm_assg.supp(1-2)]*)end