# HG changeset patch # User Christian Urban # Date 1285417725 14400 # Node ID b4ea19604b0ba59e2845238685cd6e9bbc73148f # Parent 6bab47906dbea2ccd5a5dddd70d57796b95b8a4d cleaned up two examples diff -r 6bab47906dbe -r b4ea19604b0b Nominal/Ex/SingleLet.thy --- a/Nominal/Ex/SingleLet.thy Sat Sep 25 02:53:39 2010 +0200 +++ b/Nominal/Ex/SingleLet.thy Sat Sep 25 08:28:45 2010 -0400 @@ -36,131 +36,7 @@ thm single_let.fsupp thm single_let.supp -lemma test2: - assumes "fv_trm t = supp t" - shows "\p. fv_trm (p \ t) = supp (p \ t)" -apply(rule allI) -apply(rule_tac p="-p" in permute_boolE) -apply(perm_simp add: single_let.fv_bn_eqvt permute_minus_cancel) -apply(rule assms) -done -lemma supp_fv: - "fv_trm x = supp x" - "fv_assg xa = supp xa" - "fv_bn xa = supp_rel alpha_bn xa" -apply(induct rule: single_let.inducts) --- " 0A " -apply(simp only: single_let.fv_defs supp_Pair[symmetric]) -apply(simp (no_asm) only: supp_def) -apply(perm_simp add: single_let.perm_simps single_let.fv_bn_eqvt) -apply(simp (no_asm) only: single_let.eq_iff Abs_eq_iff) --- " 0B" -apply(simp only: single_let.fv_defs supp_Pair[symmetric]) -apply(simp (no_asm) only: supp_def) -apply(perm_simp add: single_let.perm_simps single_let.fv_bn_eqvt) -apply(simp (no_asm) only: single_let.eq_iff Abs_eq_iff) -apply(simp only: alphas prod_alpha_def prod_fv.simps prod_rel.simps permute_prod_def prod.recs prod.cases prod.inject) ---" 1 " -apply(simp only: single_let.fv_defs supp_Pair[symmetric]) -apply(tactic {* EqSubst.eqsubst_tac @{context} [1] @{thms supp_abs(3)[symmetric]} 1 *}) -apply(simp (no_asm) only: supp_def) -apply(perm_simp add: single_let.perm_simps single_let.fv_bn_eqvt) -apply(simp (no_asm) only: single_let.eq_iff Abs_eq_iff) -apply(simp only: alphas prod_alpha_def prod_fv.simps prod_rel.simps permute_prod_def prod.recs prod.cases prod.inject) -apply(drule test2) -apply(simp only:) --- " 2 " -apply(simp only: single_let.fv_defs supp_Pair[symmetric]) -apply(subst supp_abs(3)[symmetric]) -apply(simp (no_asm) only: supp_def supp_rel_def) -apply(perm_simp add: single_let.perm_simps single_let.fv_bn_eqvt) -apply(simp (no_asm) only: single_let.eq_iff Abs_eq_iff) -apply(simp only: de_Morgan_conj Collect_disj_eq finite_Un) -apply(simp only: alphas prod_alpha_def prod_fv.simps prod_rel.simps permute_prod_def prod.recs prod.cases prod.inject) -apply(drule test2) -apply(simp only:) --- " 3 " -apply(simp only: single_let.fv_defs supp_Pair[symmetric]) -apply(subst supp_abs(1)[symmetric]) -apply(simp (no_asm) only: supp_def supp_rel_def) -apply(perm_simp add: single_let.perm_simps single_let.fv_bn_eqvt) -apply(simp (no_asm) only: single_let.eq_iff Abs_eq_iff) -apply(simp only: alphas prod_alpha_def prod_fv.simps prod_rel.simps permute_prod_def prod.recs prod.cases prod.inject) -apply(drule test2)+ -apply(simp only: supp_Pair Un_assoc conj_assoc) --- " Bar " -apply(simp only: single_let.fv_defs supp_Pair[symmetric]) -apply(subst supp_abs(3)[symmetric]) -apply(simp (no_asm) only: supp_def) -apply(perm_simp add: single_let.perm_simps single_let.fv_bn_eqvt) -apply(simp (no_asm) only: single_let.eq_iff Abs_eq_iff) -apply(simp only: alphas prod_alpha_def prod_fv.simps prod_rel.simps permute_prod_def prod.recs prod.cases prod.inject) -apply(drule test2)+ -apply(simp only: supp_Pair Un_assoc conj_assoc) --- " Baz " -apply(simp only: single_let.fv_defs supp_Pair[symmetric]) -apply(subst supp_abs(1)[symmetric]) -apply(subst supp_abs(1)[symmetric]) -apply(simp (no_asm) only: supp_def) -apply(perm_simp add: single_let.perm_simps single_let.fv_bn_eqvt) -apply(simp (no_asm) only: single_let.eq_iff Abs_eq_iff) -apply(simp only: alphas prod_alpha_def prod_fv.simps prod_rel.simps permute_prod_def prod.recs prod.cases prod.inject) -apply(simp only: de_Morgan_conj Collect_disj_eq finite_Un) -apply(drule test2)+ -apply(simp only: supp_Pair Un_assoc conj_assoc) --- "last" -apply(simp only: single_let.fv_defs supp_Pair[symmetric]) -apply(tactic {* EqSubst.eqsubst_tac @{context} [1] @{thms supp_abs(3)[symmetric]} 1 *}) -apply(simp (no_asm) only: supp_def supp_rel_def) -apply(perm_simp add: single_let.perm_simps single_let.fv_bn_eqvt) -apply(simp (no_asm) only: single_let.eq_iff Abs_eq_iff) -apply(simp only: alphas prod_alpha_def prod_fv.simps prod_rel.simps permute_prod_def prod.recs prod.cases prod.inject) -apply(simp only: de_Morgan_conj Collect_disj_eq finite_Un) -apply(drule test2)+ -apply(simp only: supp_Pair Un_assoc conj_assoc) --- "other case" -apply(simp only: single_let.fv_defs supp_Pair[symmetric]) -apply(simp only: supp_abs(3)[symmetric]) -apply(simp (no_asm) only: supp_def supp_rel_def) -apply(perm_simp add: single_let.perm_simps single_let.fv_bn_eqvt) -apply(simp (no_asm) only: single_let.eq_iff Abs_eq_iff) -apply(simp only: alphas prod_alpha_def prod_fv.simps prod_rel.simps permute_prod_def prod.recs prod.cases prod.inject) -apply(simp only: de_Morgan_conj Collect_disj_eq finite_Un)? -apply(drule test2)+ -apply(simp only: supp_Pair Un_assoc conj_assoc) -done - - - - - -text {* *} - -(* -consts perm_bn_trm :: "perm \ trm \ trm" -consts perm_bn_assg :: "perm \ assg \ 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 trm -typ assg - -thm trm_assg.fv -thm trm_assg.supp -thm trm_assg.eq_iff -thm trm_assg.bn -thm trm_assg.perm -thm trm_assg.induct -thm trm_assg.inducts -thm trm_assg.distinct -ML {* Sign.of_sort @{theory} (@{typ trm}, @{sort fs}) *} -*) diff -r 6bab47906dbe -r b4ea19604b0b Nominal/Ex/TypeSchemes.thy --- a/Nominal/Ex/TypeSchemes.thy Sat Sep 25 02:53:39 2010 +0200 +++ b/Nominal/Ex/TypeSchemes.thy Sat Sep 25 08:28:45 2010 -0400 @@ -6,7 +6,7 @@ atom_decl name -declare [[STEPS = 100]] +(* defined as a single nominal datatype *) nominal_datatype ty = Var "name" @@ -26,6 +26,8 @@ thm ty_tys.supports thm ty_tys.fsupp +(* defined as two separate nominal datatypes *) + nominal_datatype ty2 = Var2 "name" | Fun2 "ty2" "ty2" @@ -50,8 +52,6 @@ text {* *} (* -ML {* Sign.of_sort @{theory} (@{typ ty}, @{sort fs}) *} - lemma strong_induct: assumes a1: "\name b. P b (Var name)" and a2: "\t1 t2 b. \\c. P c t1; \c. P c t2\ \ P b (Fun t1 t2)" @@ -74,7 +74,7 @@ s="pa \ (p \ All fset ty)" in subst) apply (rule supp_perm_eq) apply assumption - apply (simp only: ty_tys.perm) + apply (simp only: ty_tys.perm_simps) apply (rule a3) apply(erule_tac x="(pa + p)" in allE) apply simp @@ -89,7 +89,8 @@ apply (simp add: fresh_star_def) apply clarify apply (simp add: fresh_def) - apply (simp add: ty_tys_supp) + apply(auto) + apply (simp add: ty_tys.supp) done then have "P a (0 \ t) \ P' d (0 \ ts)" by blast then show ?thesis by simp