# HG changeset patch # User Christian Urban <urbanc@in.tum.de> # Date 1263610560 -3600 # Node ID 4646198988909cf35bf26b5dfbe71d15a2a4da88 # Parent 4e0b89d886ac7744d4b8ecd4f143080fd368962c used "new" alpha-equivalence relation (according to new scheme); proved equivalence theorems and so on diff -r 4e0b89d886ac -r 464619898890 Quot/Examples/LamEx.thy --- a/Quot/Examples/LamEx.thy Sat Jan 16 02:09:38 2010 +0100 +++ b/Quot/Examples/LamEx.thy Sat Jan 16 03:56:00 2010 +0100 @@ -29,6 +29,8 @@ end +declare perm_rlam.simps[eqvt] + instance rlam::pt_name apply(default) apply(induct_tac [!] x rule: rlam.induct) @@ -48,34 +50,126 @@ apply(simp add: supp_atm) done +declare set_diff_eqvt[eqvt] + +lemma rfv_eqvt[eqvt]: + fixes pi::"name prm" + shows "(pi\<bullet>rfv t) = rfv (pi\<bullet>t)" +apply(induct t) +apply(simp_all) +apply(simp add: perm_set_eq) +apply(simp add: union_eqvt) +apply(simp add: set_diff_eqvt) +apply(simp add: perm_set_eq) +done + inductive - alpha :: "rlam \<Rightarrow> rlam \<Rightarrow> bool" ("_ \<approx> _" [100, 100] 100) + alpha :: "rlam \<Rightarrow> rlam \<Rightarrow> bool" ("_ \<approx> _" [100, 100] 100) where a1: "a = b \<Longrightarrow> (rVar a) \<approx> (rVar b)" | a2: "\<lbrakk>t1 \<approx> t2; s1 \<approx> s2\<rbrakk> \<Longrightarrow> rApp t1 s1 \<approx> rApp t2 s2" -| a3: "\<lbrakk>t \<approx> ([(a,b)] \<bullet> s); a \<notin> rfv (rLam b t)\<rbrakk> \<Longrightarrow> rLam a t \<approx> rLam b s" +| a3: "\<exists>pi::name prm. (rfv t - {a} = rfv s - {b} \<and> (rfv t - {a})\<sharp>* pi \<and> (pi \<bullet> t) \<approx> s \<and> (pi \<bullet> a) = b) + \<Longrightarrow> rLam a t \<approx> rLam b s" -lemma helper: - fixes t::"rlam" - and a::"name" - shows "[(a, a)] \<bullet> t = t" -by (induct t) - (auto simp add: calc_atm) +(* should be automatic with new version of eqvt-machinery *) +lemma alpha_eqvt: + fixes pi::"name prm" + shows "t \<approx> s \<Longrightarrow> (pi \<bullet> t) \<approx> (pi \<bullet> s)" +apply(induct rule: alpha.induct) +apply(simp add: a1) +apply(simp add: a2) +apply(simp) +apply(rule a3) +apply(erule conjE) +apply(erule exE) +apply(erule conjE) +apply(rule_tac x="pi \<bullet> pia" in exI) +apply(rule conjI) +apply(rule_tac pi1="rev pi" in perm_bij[THEN iffD1]) +apply(perm_simp add: eqvts) +apply(rule conjI) +apply(rule_tac pi1="rev pi" in pt_fresh_star_bij(1)[OF pt_name_inst at_name_inst, THEN iffD1]) +apply(perm_simp add: eqvts) +apply(rule conjI) +apply(subst perm_compose[symmetric]) +apply(simp) +apply(subst perm_compose[symmetric]) +apply(simp) +done lemma alpha_refl: - fixes t::"rlam" shows "t \<approx> t" - apply(induct t rule: rlam.induct) - apply(simp add: a1) - apply(simp add: a2) - apply(rule a3) - apply(simp add: helper) - apply(simp) - done +apply(induct t rule: rlam.induct) +apply(simp add: a1) +apply(simp add: a2) +apply(rule a3) +apply(rule_tac x="[]" in exI) +apply(simp_all add: fresh_star_def fresh_list_nil) +done + +lemma alpha_sym: + shows "t \<approx> s \<Longrightarrow> s \<approx> t" +apply(induct rule: alpha.induct) +apply(simp add: a1) +apply(simp add: a2) +apply(rule a3) +apply(erule exE) +apply(rule_tac x="rev pi" in exI) +apply(simp) +apply(simp add: fresh_star_def fresh_list_rev) +apply(rule conjI) +apply(erule conjE)+ +apply(rotate_tac 3) +apply(drule_tac pi="rev pi" in alpha_eqvt) +apply(perm_simp) +apply(rule pt_bij2[OF pt_name_inst at_name_inst]) +apply(simp) +done + +lemma alpha_trans: + shows "t1 \<approx> t2 \<Longrightarrow> t2 \<approx> t3 \<Longrightarrow> t1 \<approx> t3" +apply(induct arbitrary: t3 rule: alpha.induct) +apply(erule alpha.cases) +apply(simp_all) +apply(simp add: a1) +apply(rotate_tac 4) +apply(erule alpha.cases) +apply(simp_all) +apply(simp add: a2) +apply(rotate_tac 1) +apply(erule alpha.cases) +apply(simp_all) +apply(erule conjE)+ +apply(erule exE)+ +apply(erule conjE)+ +apply(rule a3) +apply(rule_tac x="pia @ pi" in exI) +apply(simp add: fresh_star_def fresh_list_append) +apply(simp add: pt_name2) +apply(drule_tac x="rev pia \<bullet> sa" in spec) +apply(drule mp) +apply(rotate_tac 8) +apply(drule_tac pi="rev pia" in alpha_eqvt) +apply(perm_simp) +apply(rotate_tac 11) +apply(drule_tac pi="pia" in alpha_eqvt) +apply(perm_simp) +done lemma alpha_equivp: shows "equivp alpha" -sorry +apply(rule equivpI) +unfolding reflp_def symp_def transp_def +apply(auto intro: alpha_refl alpha_sym alpha_trans) +done + +lemma alpha_rfv: + shows "t \<approx> s \<Longrightarrow> rfv t = rfv s" +apply(induct rule: alpha.induct) +apply(simp) +apply(simp) +apply(simp) +done quotient_type lam = rlam / alpha by (rule alpha_equivp) @@ -114,13 +208,6 @@ end -lemma real_alpha: - assumes a: "t = [(a,b)]\<bullet>s" "a\<sharp>[b].s" - shows "Lam a t = Lam b s" -using a -unfolding fresh_def supp_def -sorry - lemma perm_rsp[quot_respect]: "(op = ===> alpha ===> alpha) op \<bullet> op \<bullet>" apply(auto) @@ -143,20 +230,18 @@ lemma rLam_rsp[quot_respect]: "(op = ===> alpha ===> alpha) rLam rLam" apply(auto) apply(rule a3) - apply(simp add: helper) - apply(simp) + apply(rule_tac x="[]" in exI) + unfolding fresh_star_def + apply(simp add: fresh_list_nil) + apply(simp add: alpha_rfv) done lemma rfv_rsp[quot_respect]: "(alpha ===> op =) rfv rfv" - sorry +apply(simp add: alpha_rfv) +done -lemma rvar_inject: "rVar a \<approx> rVar b = (a = b)" - apply (auto) - apply (erule alpha.cases) - apply (simp_all add: rlam.inject alpha_refl) - done - +section {* lifted theorems *} lemma lam_induct: "\<lbrakk>\<And>name. P (Var name); @@ -196,36 +281,46 @@ by (lifting a2) lemma a3: - "\<lbrakk>x = [(a, b)] \<bullet> xa; a \<notin> fv (Lam b x)\<rbrakk> \<Longrightarrow> Lam a x = Lam b xa" + "\<lbrakk>\<exists>pi::name prm. (fv t - {a} = fv s - {b} \<and> (fv t - {a})\<sharp>* pi \<and> (pi \<bullet> t) = s \<and> (pi \<bullet> a) = b)\<rbrakk> + \<Longrightarrow> Lam a t = Lam b s" by (lifting a3) lemma alpha_cases: "\<lbrakk>a1 = a2; \<And>a b. \<lbrakk>a1 = Var a; a2 = Var b; a = b\<rbrakk> \<Longrightarrow> P; \<And>x xa xb xc. \<lbrakk>a1 = App x xb; a2 = App xa xc; x = xa; xb = xc\<rbrakk> \<Longrightarrow> P; - \<And>x a b xa. \<lbrakk>a1 = Lam a x; a2 = Lam b xa; x = [(a, b)] \<bullet> xa; a \<notin> fv (Lam b x)\<rbrakk> \<Longrightarrow> P\<rbrakk> + \<And>t a s b. \<lbrakk>a1 = Lam a t; a2 = Lam b s; + \<exists>pi::name prm. fv t - {a} = fv s - {b} \<and> (fv t - {a}) \<sharp>* pi \<and> (pi \<bullet> t) = s \<and> pi \<bullet> a = b\<rbrakk> \<Longrightarrow> P\<rbrakk> \<Longrightarrow> P" by (lifting alpha.cases) lemma alpha_induct: "\<lbrakk>qx = qxa; \<And>a b. a = b \<Longrightarrow> qxb (Var a) (Var b); \<And>x xa xb xc. \<lbrakk>x = xa; qxb x xa; xb = xc; qxb xb xc\<rbrakk> \<Longrightarrow> qxb (App x xb) (App xa xc); - \<And>x a b xa. - \<lbrakk>x = [(a, b)] \<bullet> xa; qxb x ([(a, b)] \<bullet> xa); a \<notin> fv (Lam b x)\<rbrakk> \<Longrightarrow> qxb (Lam a x) (Lam b xa)\<rbrakk> + \<And>t a s b. + \<lbrakk>\<exists>pi::name prm. fv t - {a} = fv s - {b} \<and> + (fv t - {a}) \<sharp>* pi \<and> ((pi \<bullet> t) = s \<and> qxb (pi \<bullet> t) s) \<and> pi \<bullet> a = b\<rbrakk> \<Longrightarrow> qxb (Lam a t) (Lam b s)\<rbrakk> \<Longrightarrow> qxb qx qxa" by (lifting alpha.induct) -lemma var_inject: - "(Var a = Var b) = (a = b)" - by (lifting rvar_inject) - - -lemma app_inject: - "(App t1 t2 = App s1 s2) = (t1 = s1 \<and> t2 = s2)" -sorry +lemma lam_inject [simp]: + shows "(Var a = Var b) = (a = b)" + and "(App t1 t2 = App s1 s2) = (t1 = s1 \<and> t2 = s2)" +apply(lifting rlam.inject(1) rlam.inject(2)) +apply(auto) +apply(drule alpha.cases) +apply(simp_all) +apply(simp add: alpha.a1) +apply(drule alpha.cases) +apply(simp_all) +apply(drule alpha.cases) +apply(simp_all) +apply(rule alpha.a2) +apply(simp_all) +done lemma var_supp1: shows "(supp (Var a)) = ((supp a)::name set)" -apply(simp add: supp_def var_inject) +apply(simp add: supp_def) done lemma var_supp: @@ -236,13 +331,14 @@ lemma app_supp: shows "supp (App t1 t2) = (supp t1) \<union> ((supp t2)::name set)" -apply(simp only: perm_lam supp_def app_inject) +apply(simp only: perm_lam supp_def lam_inject) apply(simp add: Collect_imp_eq Collect_neg_eq) done lemma lam_supp: shows "supp (Lam x t) = ((supp ([x].t))::name set)" apply(simp add: supp_def) +apply(simp add: abs_perm) sorry @@ -251,7 +347,8 @@ apply(induct_tac x rule: lam_induct) apply(simp add: var_supp) apply(simp add: app_supp) -sorry +apply(simp add: lam_supp abs_supp) +done lemma fresh_lam: "(a \<sharp> Lam b t) \<longleftrightarrow> (a = b) \<or> (a \<noteq> b \<and> a \<sharp> t)"