diff -r c91530f18d9c -r f3615235dc4d src/HOL/Nominal/Examples/Class.thy --- a/src/HOL/Nominal/Examples/Class.thy Thu May 24 12:00:47 2007 +0200 +++ b/src/HOL/Nominal/Examples/Class.thy Thu May 24 12:09:38 2007 +0200 @@ -1,7 +1,7 @@ (* $Id$ *) -theory Class -imports "../Nominal" +theory ClassWork +imports "Nominal" begin section {* Term-Calculus from Urban's PhD *} @@ -11,11 +11,34 @@ text {* types *} nominal_datatype ty = - PROP "string" + PR "string" | NOT "ty" | AND "ty" "ty" ("_ AND _" [100,100] 100) | OR "ty" "ty" ("_ OR _" [100,100] 100) - | IMP "ty" "ty" ("_ IMPL _" [100,100] 100) + | IMP "ty" "ty" ("_ IMP _" [100,100] 100) + +instance ty :: size .. + +nominal_primrec + "size (PR s) = (1::nat)" + "size (NOT T) = 1 + size T" + "size (T1 AND T2) = 1 + size T1 + size T2" + "size (T1 OR T2) = 1 + size T1 + size T2" + "size (T1 IMP T2) = 1 + size T1 + size T2" +by (rule TrueI)+ + +lemma ty_cases: + fixes T::ty + shows "(\s. T=PR s) \ (\T'. T=NOT T') \ (\S U. T=S OR U) \ (\S U. T=S AND U) \ (\S U. T=S IMP U)" +by (induct T rule:ty.weak_induct) (auto) + +lemma fresh_ty: + fixes a::"coname" + and x::"name" + and T::"ty" + shows "a\T" and "x\T" +by (nominal_induct T rule: ty.induct) + (auto simp add: fresh_string) text {* terms *} @@ -33,12 +56,6 @@ | ImpR "\name\(\coname\trm)" "coname" ("ImpR (_).<_>._ _" [100,100,100,100] 100) | ImpL "\coname\trm" "\name\trm" "name" ("ImpL <_>._ (_)._ _" [100,100,100,100,100] 100) -text {* Induction principles *} - -thm trm.weak_induct --"weak" -thm trm.induct --"strong" -thm trm.induct' --"strong with explicit context (rarely needed)" - text {* named terms *} nominal_datatype ntrm = Na "\name\trm" ("((_):_)" [100,100] 100) @@ -97,7 +114,6 @@ lemmas eq_bij = pt_bij[OF pt_name_inst, OF at_name_inst] pt_bij[OF pt_coname_inst, OF at_coname_inst] - lemma crename_name_eqvt[eqvt]: fixes pi::"name prm" shows "pi\(M[d\c>e]) = (pi\M)[(pi\d)\c>(pi\e)]" @@ -105,7 +121,6 @@ apply(auto simp add: fresh_bij eq_bij) done - lemma crename_coname_eqvt[eqvt]: fixes pi::"coname prm" shows "pi\(M[d\c>e]) = (pi\M)[(pi\d)\c>(pi\e)]" @@ -127,46 +142,2769 @@ apply(auto simp add: fresh_bij eq_bij) done +lemmas rename_eqvts = crename_name_eqvt crename_coname_eqvt + nrename_name_eqvt nrename_coname_eqvt +lemma nrename_fresh: + assumes a: "x\M" + shows "M[x\n>y] = M" +using a +by (nominal_induct M avoiding: x y rule: trm.induct) + (auto simp add: trm.inject fresh_atm abs_fresh fin_supp abs_supp) + +lemma crename_fresh: + assumes a: "a\M" + shows "M[a\c>b] = M" +using a +by (nominal_induct M avoiding: a b rule: trm.induct) + (auto simp add: trm.inject fresh_atm abs_fresh) + +lemma nrename_nfresh: + fixes x::"name" + shows "x\y\x\M[x\n>y]" +by (nominal_induct M avoiding: x y rule: trm.induct) + (auto simp add: fresh_atm abs_fresh abs_supp fin_supp) + + lemma crename_nfresh: + fixes x::"name" + shows "x\M\x\M[a\c>b]" +by (nominal_induct M avoiding: a b rule: trm.induct) + (auto simp add: fresh_atm abs_fresh abs_supp fin_supp) + +lemma crename_cfresh: + fixes a::"coname" + shows "a\b\a\M[a\c>b]" +by (nominal_induct M avoiding: a b rule: trm.induct) + (auto simp add: fresh_atm abs_fresh abs_supp fin_supp) + +lemma nrename_cfresh: + fixes c::"coname" + shows "c\M\c\M[x\n>y]" +by (nominal_induct M avoiding: x y rule: trm.induct) + (auto simp add: fresh_atm abs_fresh abs_supp fin_supp) + +lemma nrename_nfresh': + fixes x::"name" + shows "x\(M,z,y)\x\M[z\n>y]" +by (nominal_induct M avoiding: x z y rule: trm.induct) + (auto simp add: fresh_prod fresh_atm abs_fresh abs_supp fin_supp) + +lemma crename_cfresh': + fixes a::"coname" + shows "a\(M,b,c)\a\M[b\c>c]" +by (nominal_induct M avoiding: a b c rule: trm.induct) + (auto simp add: fresh_prod fresh_atm abs_fresh abs_supp fin_supp) + +lemma nrename_rename: + assumes a: "x'\M" + shows "([(x',x)]\M)[x'\n>y]= M[x\n>y]" +using a +apply(nominal_induct M avoiding: x x' y rule: trm.induct) +apply(auto simp add: abs_fresh fresh_bij fresh_atm fresh_prod fresh_right calc_atm abs_supp fin_supp) +apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm) +done + +lemma crename_rename: + assumes a: "a'\M" + shows "([(a',a)]\M)[a'\c>b]= M[a\c>b]" +using a +apply(nominal_induct M avoiding: a a' b rule: trm.induct) +apply(auto simp add: abs_fresh fresh_bij fresh_atm fresh_prod fresh_right calc_atm abs_supp fin_supp) +apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm) +done + +lemmas rename_fresh = nrename_fresh crename_fresh + nrename_nfresh crename_nfresh crename_cfresh nrename_cfresh + nrename_nfresh' crename_cfresh' + nrename_rename crename_rename + +lemma better_nrename_Cut: + assumes a: "x\(u,v)" + shows "(Cut .M (x).N)[u\n>v] = Cut .(M[u\n>v]) (x).(N[u\n>v])" +proof - + obtain x'::"name" where fs1: "x'\(M,N,a,x,u,v)" by (rule exists_fresh(1), rule fin_supp, blast) + obtain a'::"coname" where fs2: "a'\(M,N,a,x,u,v)" by (rule exists_fresh(2), rule fin_supp, blast) + have eq1: "(Cut .M (x).N) = (Cut .([(a',a)]\M) (x').([(x',x)]\N))" + using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm) + have "(Cut .([(a',a)]\M) (x').([(x',x)]\N))[u\n>v] = + Cut .(([(a',a)]\M)[u\n>v]) (x').(([(x',x)]\N)[u\n>v])" + using fs1 fs2 + apply - + apply(rule nrename.simps) + apply(simp add: fresh_left calc_atm) + apply(simp add: fresh_left calc_atm) + done + also have "\ = Cut .(M[u\n>v]) (x).(N[u\n>v])" using fs1 fs2 a + apply - + apply(simp add: trm.inject alpha fresh_atm fresh_prod rename_eqvts) + apply(simp add: calc_atm) + apply(simp add: rename_fresh fresh_atm) + done + finally show "(Cut .M (x).N)[u\n>v] = Cut .(M[u\n>v]) (x).(N[u\n>v])" using eq1 + by simp +qed + +lemma better_crename_Cut: + assumes a: "a\(b,c)" + shows "(Cut .M (x).N)[b\c>c] = Cut .(M[b\c>c]) (x).(N[b\c>c])" +proof - + obtain x'::"name" where fs1: "x'\(M,N,a,x,b,c)" by (rule exists_fresh(1), rule fin_supp, blast) + obtain a'::"coname" where fs2: "a'\(M,N,a,x,b,c)" by (rule exists_fresh(2), rule fin_supp, blast) + have eq1: "(Cut .M (x).N) = (Cut .([(a',a)]\M) (x').([(x',x)]\N))" + using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm) + have "(Cut .([(a',a)]\M) (x').([(x',x)]\N))[b\c>c] = + Cut .(([(a',a)]\M)[b\c>c]) (x').(([(x',x)]\N)[b\c>c])" + using fs1 fs2 + apply - + apply(rule crename.simps) + apply(simp add: fresh_left calc_atm) + apply(simp add: fresh_left calc_atm) + done + also have "\ = Cut .(M[b\c>c]) (x).(N[b\c>c])" using fs1 fs2 a + apply - + apply(simp add: trm.inject alpha fresh_atm fresh_prod rename_eqvts) + apply(simp add: calc_atm) + apply(simp add: rename_fresh fresh_atm) + done + finally show "(Cut .M (x).N)[b\c>c] = Cut .(M[b\c>c]) (x).(N[b\c>c])" using eq1 + by simp +qed + +lemma crename_id: + shows "M[a\c>a] = M" +by (nominal_induct M avoiding: a rule: trm.induct) (auto) + +lemma nrename_id: + shows "M[x\n>x] = M" +by (nominal_induct M avoiding: x rule: trm.induct) (auto) + +lemma nrename_swap: + shows "x\M \ [(x,y)]\M = M[y\n>x]" +by (nominal_induct M avoiding: x y rule: trm.induct) + (simp_all add: calc_atm fresh_atm trm.inject alpha abs_fresh abs_supp fin_supp) + +lemma crename_swap: + shows "a\M \ [(a,b)]\M = M[b\c>a]" +by (nominal_induct M avoiding: a b rule: trm.induct) + (simp_all add: calc_atm fresh_atm trm.inject alpha abs_fresh abs_supp fin_supp) + +lemma crename_ax: + assumes a: "M[a\c>b] = Ax x c" "c\a" "c\b" + shows "M = Ax x c" +using a +apply(nominal_induct M avoiding: a b x c rule: trm.induct) +apply(simp_all add: trm.inject split: if_splits) +done + +lemma nrename_ax: + assumes a: "M[x\n>y] = Ax z a" "z\x" "z\y" + shows "M = Ax z a" +using a +apply(nominal_induct M avoiding: x y z a rule: trm.induct) +apply(simp_all add: trm.inject split: if_splits) +done + text {* substitution functions *} +lemma fresh_perm_coname: + fixes c::"coname" + and pi::"coname prm" + and M::"trm" + assumes a: "c\pi" "c\M" + shows "c\(pi\M)" +using a +apply - +apply(simp add: fresh_left) +apply(simp add: at_prm_fresh[OF at_coname_inst] fresh_list_rev) +done + +lemma fresh_perm_name: + fixes x::"name" + and pi::"name prm" + and M::"trm" + assumes a: "x\pi" "x\M" + shows "x\(pi\M)" +using a +apply - +apply(simp add: fresh_left) +apply(simp add: at_prm_fresh[OF at_name_inst] fresh_list_rev) +done + +lemma fresh_fun_simp_NotL: + assumes a: "x'\P" "x'\M" + shows "fresh_fun (\x'. Cut .P (x').NotL .M x') = Cut .P (x').NotL .M x'" +using a +apply - +apply(rule fresh_fun_app) +apply(rule pt_name_inst) +apply(rule at_name_inst) +apply(finite_guess) +apply(subgoal_tac "\n::name. n\(c,P,a,M)") +apply(erule exE) +apply(rule_tac x="n" in exI) +apply(simp add: fresh_prod abs_fresh) +apply(fresh_guess) +apply(rule exists_fresh') +apply(simp add: fin_supp) +apply(fresh_guess) +done + +lemma fresh_fun_NotL[eqvt_force]: + fixes pi1::"name prm" + and pi2::"coname prm" + shows "pi1\fresh_fun (\x'. Cut .P (x').NotL .M x')= + fresh_fun (pi1\(\x'. Cut .P (x').NotL .M x'))" + and "pi2\fresh_fun (\x'. Cut .P (x').NotL .M x')= + fresh_fun (pi2\(\x'. Cut .P (x').NotL .M x'))" +apply - +apply(perm_simp) +apply(generate_fresh "name") +apply(auto simp add: fresh_prod) +apply(simp add: fresh_fun_simp_NotL) +apply(rule sym) +apply(rule trans) +apply(rule fresh_fun_simp_NotL) +apply(rule fresh_perm_name) +apply(assumption) +apply(assumption) +apply(rule fresh_perm_name) +apply(assumption) +apply(assumption) +apply(simp add: at_prm_fresh[OF at_name_inst]) +apply(perm_simp) +apply(subgoal_tac "\n::name. n\(P,M,pi2\P,pi2\M,pi2)") +apply(simp add: fresh_prod) +apply(auto) +apply(simp add: fresh_fun_simp_NotL calc_atm) +apply(rule exists_fresh') +apply(simp add: fin_supp) +done + +lemma fresh_fun_simp_AndL1: + assumes a: "z'\P" "z'\M" "z'\x" + shows "fresh_fun (\z'. Cut .P (z').AndL1 (x).M z') = Cut .P (z').AndL1 (x).M z'" +using a +apply - +apply(rule fresh_fun_app) +apply(rule pt_name_inst) +apply(rule at_name_inst) +apply(finite_guess) +apply(subgoal_tac "\n::name. n\(c,P,x,M)") +apply(erule exE) +apply(rule_tac x="n" in exI) +apply(simp add: fresh_prod abs_fresh) +apply(fresh_guess) +apply(rule exists_fresh') +apply(simp add: fin_supp) +apply(fresh_guess) +done + +lemma fresh_fun_AndL1[eqvt_force]: + fixes pi1::"name prm" + and pi2::"coname prm" + shows "pi1\fresh_fun (\z'. Cut .P (z').AndL1 (x).M z')= + fresh_fun (pi1\(\z'. Cut .P (z').AndL1 (x).M z'))" + and "pi2\fresh_fun (\z'. Cut .P (z').AndL1 (x).M z')= + fresh_fun (pi2\(\z'. Cut .P (z').AndL1 (x).M z'))" +apply - +apply(perm_simp) +apply(subgoal_tac "\n::name. n\(P,M,x,pi1\P,pi1\M,pi1\x,pi1)") +apply(simp add: fresh_prod) +apply(auto) +apply(simp add: fresh_fun_simp_AndL1 at_prm_fresh[OF at_name_inst]) +apply(rule exists_fresh') +apply(simp add: fin_supp) +apply(perm_simp) +apply(subgoal_tac "\n::name. n\(P,M,x,pi2\P,pi2\M,pi2\x,pi2)") +apply(simp add: fresh_prod) +apply(auto) +apply(simp add: fresh_fun_simp_AndL1 calc_atm) +apply(rule exists_fresh') +apply(simp add: fin_supp) +done + +lemma fresh_fun_simp_AndL2: + assumes a: "z'\P" "z'\M" "z'\x" + shows "fresh_fun (\z'. Cut .P (z').AndL2 (x).M z') = Cut .P (z').AndL2 (x).M z'" +using a +apply - +apply(rule fresh_fun_app) +apply(rule pt_name_inst) +apply(rule at_name_inst) +apply(finite_guess) +apply(subgoal_tac "\n::name. n\(c,P,x,M)") +apply(erule exE) +apply(rule_tac x="n" in exI) +apply(simp add: fresh_prod abs_fresh) +apply(fresh_guess) +apply(rule exists_fresh') +apply(simp add: fin_supp) +apply(fresh_guess) +done + +lemma fresh_fun_AndL2[eqvt_force]: + fixes pi1::"name prm" + and pi2::"coname prm" + shows "pi1\fresh_fun (\z'. Cut .P (z').AndL2 (x).M z')= + fresh_fun (pi1\(\z'. Cut .P (z').AndL2 (x).M z'))" + and "pi2\fresh_fun (\z'. Cut .P (z').AndL2 (x).M z')= + fresh_fun (pi2\(\z'. Cut .P (z').AndL2 (x).M z'))" +apply - +apply(perm_simp) +apply(subgoal_tac "\n::name. n\(P,M,x,pi1\P,pi1\M,pi1\x,pi1)") +apply(simp add: fresh_prod) +apply(auto) +apply(simp add: fresh_fun_simp_AndL2 at_prm_fresh[OF at_name_inst]) +apply(rule exists_fresh') +apply(simp add: fin_supp) +apply(perm_simp) +apply(subgoal_tac "\n::name. n\(P,M,x,pi2\P,pi2\M,pi2\x,pi2)") +apply(simp add: fresh_prod) +apply(auto) +apply(simp add: fresh_fun_simp_AndL2 calc_atm) +apply(rule exists_fresh') +apply(simp add: fin_supp) +done + +lemma fresh_fun_simp_OrL: + assumes a: "z'\P" "z'\M" "z'\N" "z'\u" "z'\x" + shows "fresh_fun (\z'. Cut .P (z').OrL (x).M (u).N z') = Cut .P (z').OrL (x).M (u).N z'" +using a +apply - +apply(rule fresh_fun_app) +apply(rule pt_name_inst) +apply(rule at_name_inst) +apply(finite_guess) +apply(subgoal_tac "\n::name. n\(c,P,x,M,u,N)") +apply(erule exE) +apply(rule_tac x="n" in exI) +apply(simp add: fresh_prod abs_fresh) +apply(fresh_guess) +apply(rule exists_fresh') +apply(simp add: fin_supp) +apply(fresh_guess) +done + +lemma fresh_fun_OrL[eqvt_force]: + fixes pi1::"name prm" + and pi2::"coname prm" + shows "pi1\fresh_fun (\z'. Cut .P (z').OrL (x).M (u).N z')= + fresh_fun (pi1\(\z'. Cut .P (z').OrL (x).M (u).N z'))" + and "pi2\fresh_fun (\z'. Cut .P (z').OrL (x).M (u).N z')= + fresh_fun (pi2\(\z'. Cut .P (z').OrL (x).M (u).N z'))" +apply - +apply(perm_simp) +apply(subgoal_tac "\n::name. n\(P,M,N,x,u,pi1\P,pi1\M,pi1\N,pi1\x,pi1\u,pi1)") +apply(simp add: fresh_prod) +apply(auto) +apply(simp add: fresh_fun_simp_OrL at_prm_fresh[OF at_name_inst]) +apply(rule exists_fresh') +apply(simp add: fin_supp) +apply(perm_simp) +apply(subgoal_tac "\n::name. n\(P,M,N,x,u,pi2\P,pi2\M,pi2\N,pi2\x,pi2\u,pi2)") +apply(simp add: fresh_prod) +apply(auto) +apply(simp add: fresh_fun_simp_OrL calc_atm) +apply(rule exists_fresh') +apply(simp add: fin_supp) +done + +lemma fresh_fun_simp_ImpL: + assumes a: "z'\P" "z'\M" "z'\N" "z'\x" + shows "fresh_fun (\z'. Cut .P (z').ImpL .M (x).N z') = Cut .P (z').ImpL .M (x).N z'" +using a +apply - +apply(rule fresh_fun_app) +apply(rule pt_name_inst) +apply(rule at_name_inst) +apply(finite_guess) +apply(subgoal_tac "\n::name. n\(c,P,x,M,N)") +apply(erule exE) +apply(rule_tac x="n" in exI) +apply(simp add: fresh_prod abs_fresh) +apply(fresh_guess) +apply(rule exists_fresh') +apply(simp add: fin_supp) +apply(fresh_guess) +done + +lemma fresh_fun_ImpL[eqvt_force]: + fixes pi1::"name prm" + and pi2::"coname prm" + shows "pi1\fresh_fun (\z'. Cut .P (z').ImpL .M (x).N z')= + fresh_fun (pi1\(\z'. Cut .P (z').ImpL .M (x).N z'))" + and "pi2\fresh_fun (\z'. Cut .P (z').ImpL .M (x).N z')= + fresh_fun (pi2\(\z'. Cut .P (z').ImpL .M (x).N z'))" +apply - +apply(perm_simp) +apply(subgoal_tac "\n::name. n\(P,M,N,x,pi1\P,pi1\M,pi1\N,pi1\x,pi1)") +apply(simp add: fresh_prod) +apply(auto) +apply(simp add: fresh_fun_simp_ImpL at_prm_fresh[OF at_name_inst]) +apply(rule exists_fresh') +apply(simp add: fin_supp) +apply(perm_simp) +apply(subgoal_tac "\n::name. n\(P,M,N,x,pi2\P,pi2\M,pi2\N,pi2\x,pi2)") +apply(simp add: fresh_prod) +apply(auto) +apply(simp add: fresh_fun_simp_ImpL calc_atm) +apply(rule exists_fresh') +apply(simp add: fin_supp) +done + +lemma fresh_fun_simp_NotR: + assumes a: "a'\P" "a'\M" + shows "fresh_fun (\a'. Cut .(NotR (y).M a') (x).P) = Cut .(NotR (y).M a') (x).P" +using a +apply - +apply(rule fresh_fun_app) +apply(rule pt_coname_inst) +apply(rule at_coname_inst) +apply(finite_guess) +apply(subgoal_tac "\n::coname. n\(x,P,y,M)") +apply(erule exE) +apply(rule_tac x="n" in exI) +apply(simp add: fresh_prod abs_fresh) +apply(fresh_guess) +apply(rule exists_fresh') +apply(simp add: fin_supp) +apply(fresh_guess) +done + +lemma fresh_fun_NotR[eqvt_force]: + fixes pi1::"name prm" + and pi2::"coname prm" + shows "pi1\fresh_fun (\a'. Cut .(NotR (y).M a') (x).P)= + fresh_fun (pi1\(\a'. Cut .(NotR (y).M a') (x).P))" + and "pi2\fresh_fun (\a'. Cut .(NotR (y).M a') (x).P)= + fresh_fun (pi2\(\a'. Cut .(NotR (y).M a') (x).P))" +apply - +apply(perm_simp) +apply(subgoal_tac "\n::coname. n\(P,M,pi1\P,pi1\M,pi1)") +apply(simp add: fresh_prod) +apply(auto) +apply(simp add: fresh_fun_simp_NotR calc_atm) +apply(rule exists_fresh') +apply(simp add: fin_supp) +apply(perm_simp) +apply(subgoal_tac "\n::coname. n\(P,M,pi2\P,pi2\M,pi2)") +apply(simp add: fresh_prod) +apply(auto) +apply(simp add: fresh_fun_simp_NotR at_prm_fresh[OF at_coname_inst]) +apply(rule exists_fresh') +apply(simp add: fin_supp) +done + +lemma fresh_fun_simp_AndR: + assumes a: "a'\P" "a'\M" "a'\N" "a'\b" "a'\c" + shows "fresh_fun (\a'. Cut .(AndR .M .N a') (x).P) = Cut .(AndR .M .N a') (x).P" +using a +apply - +apply(rule fresh_fun_app) +apply(rule pt_coname_inst) +apply(rule at_coname_inst) +apply(finite_guess) +apply(subgoal_tac "\n::coname. n\(x,P,b,M,c,N)") +apply(erule exE) +apply(rule_tac x="n" in exI) +apply(simp add: fresh_prod abs_fresh) +apply(fresh_guess) +apply(rule exists_fresh') +apply(simp add: fin_supp) +apply(fresh_guess) +done + +lemma fresh_fun_AndR[eqvt_force]: + fixes pi1::"name prm" + and pi2::"coname prm" + shows "pi1\fresh_fun (\a'. Cut .(AndR .M .N a') (x).P)= + fresh_fun (pi1\(\a'. Cut .(AndR .M .N a') (x).P))" + and "pi2\fresh_fun (\a'. Cut .(AndR .M .N a') (x).P)= + fresh_fun (pi2\(\a'. Cut .(AndR .M .N a') (x).P))" +apply - +apply(perm_simp) +apply(subgoal_tac "\n::coname. n\(P,M,N,b,c,pi1\P,pi1\M,pi1\N,pi1\b,pi1\c,pi1)") +apply(simp add: fresh_prod) +apply(auto) +apply(simp add: fresh_fun_simp_AndR calc_atm) +apply(rule exists_fresh') +apply(simp add: fin_supp) +apply(perm_simp) +apply(subgoal_tac "\n::coname. n\(P,M,N,b,c,pi2\P,pi2\M,pi2\N,pi2\b,pi2\c,pi2)") +apply(simp add: fresh_prod) +apply(auto) +apply(simp add: fresh_fun_simp_AndR at_prm_fresh[OF at_coname_inst]) +apply(rule exists_fresh') +apply(simp add: fin_supp) +done + +lemma fresh_fun_simp_OrR1: + assumes a: "a'\P" "a'\M" "a'\b" + shows "fresh_fun (\a'. Cut .(OrR1 .M a') (x).P) = Cut .(OrR1 .M a') (x).P" +using a +apply - +apply(rule fresh_fun_app) +apply(rule pt_coname_inst) +apply(rule at_coname_inst) +apply(finite_guess) +apply(subgoal_tac "\n::coname. n\(x,P,b,M)") +apply(erule exE) +apply(rule_tac x="n" in exI) +apply(simp add: fresh_prod abs_fresh) +apply(fresh_guess) +apply(rule exists_fresh') +apply(simp add: fin_supp) +apply(fresh_guess) +done + +lemma fresh_fun_OrR1[eqvt_force]: + fixes pi1::"name prm" + and pi2::"coname prm" + shows "pi1\fresh_fun (\a'. Cut .(OrR1 .M a') (x).P)= + fresh_fun (pi1\(\a'. Cut .(OrR1 .M a') (x).P))" + and "pi2\fresh_fun (\a'. Cut .(OrR1 .M a') (x).P)= + fresh_fun (pi2\(\a'. Cut .(OrR1 .M a') (x).P))" +apply - +apply(perm_simp) +apply(subgoal_tac "\n::coname. n\(P,M,b,pi1\P,pi1\M,pi1\b,pi1)") +apply(simp add: fresh_prod) +apply(auto) +apply(simp add: fresh_fun_simp_OrR1 calc_atm) +apply(rule exists_fresh') +apply(simp add: fin_supp) +apply(perm_simp) +apply(subgoal_tac "\n::coname. n\(P,M,b,pi2\P,pi2\M,pi2\b,pi2)") +apply(simp add: fresh_prod) +apply(auto) +apply(simp add: fresh_fun_simp_OrR1 at_prm_fresh[OF at_coname_inst]) +apply(rule exists_fresh') +apply(simp add: fin_supp) +done + +lemma fresh_fun_simp_OrR2: + assumes a: "a'\P" "a'\M" "a'\b" + shows "fresh_fun (\a'. Cut .(OrR2 .M a') (x).P) = Cut .(OrR2 .M a') (x).P" +using a +apply - +apply(rule fresh_fun_app) +apply(rule pt_coname_inst) +apply(rule at_coname_inst) +apply(finite_guess) +apply(subgoal_tac "\n::coname. n\(x,P,b,M)") +apply(erule exE) +apply(rule_tac x="n" in exI) +apply(simp add: fresh_prod abs_fresh) +apply(fresh_guess) +apply(rule exists_fresh') +apply(simp add: fin_supp) +apply(fresh_guess) +done + +lemma fresh_fun_OrR2[eqvt_force]: + fixes pi1::"name prm" + and pi2::"coname prm" + shows "pi1\fresh_fun (\a'. Cut .(OrR2 .M a') (x).P)= + fresh_fun (pi1\(\a'. Cut .(OrR2 .M a') (x).P))" + and "pi2\fresh_fun (\a'. Cut .(OrR2 .M a') (x).P)= + fresh_fun (pi2\(\a'. Cut .(OrR2 .M a') (x).P))" +apply - +apply(perm_simp) +apply(subgoal_tac "\n::coname. n\(P,M,b,pi1\P,pi1\M,pi1\b,pi1)") +apply(simp add: fresh_prod) +apply(auto) +apply(simp add: fresh_fun_simp_OrR2 calc_atm) +apply(rule exists_fresh') +apply(simp add: fin_supp) +apply(perm_simp) +apply(subgoal_tac "\n::coname. n\(P,M,b,pi2\P,pi2\M,pi2\b,pi2)") +apply(simp add: fresh_prod) +apply(auto) +apply(simp add: fresh_fun_simp_OrR2 at_prm_fresh[OF at_coname_inst]) +apply(rule exists_fresh') +apply(simp add: fin_supp) +done + +lemma fresh_fun_simp_ImpR: + assumes a: "a'\P" "a'\M" "a'\b" + shows "fresh_fun (\a'. Cut .(ImpR (y)..M a') (x).P) = Cut .(ImpR (y)..M a') (x).P" +using a +apply - +apply(rule fresh_fun_app) +apply(rule pt_coname_inst) +apply(rule at_coname_inst) +apply(finite_guess) +apply(subgoal_tac "\n::coname. n\(x,P,y,b,M)") +apply(erule exE) +apply(rule_tac x="n" in exI) +apply(simp add: fresh_prod abs_fresh) +apply(fresh_guess) +apply(rule exists_fresh') +apply(simp add: fin_supp) +apply(fresh_guess) +done + +lemma fresh_fun_ImpR[eqvt_force]: + fixes pi1::"name prm" + and pi2::"coname prm" + shows "pi1\fresh_fun (\a'. Cut .(ImpR (y)..M a') (x).P)= + fresh_fun (pi1\(\a'. Cut .(ImpR (y)..M a') (x).P))" + and "pi2\fresh_fun (\a'. Cut .(ImpR (y)..M a') (x).P)= + fresh_fun (pi2\(\a'. Cut .(ImpR (y)..M a') (x).P))" +apply - +apply(perm_simp) +apply(subgoal_tac "\n::coname. n\(P,M,b,pi1\P,pi1\M,pi1\b,pi1)") +apply(simp add: fresh_prod) +apply(auto) +apply(simp add: fresh_fun_simp_ImpR calc_atm) +apply(rule exists_fresh') +apply(simp add: fin_supp) +apply(perm_simp) +apply(subgoal_tac "\n::coname. n\(P,M,b,pi2\P,pi2\M,pi2\b,pi2)") +apply(simp add: fresh_prod) +apply(auto) +apply(simp add: fresh_fun_simp_ImpR at_prm_fresh[OF at_coname_inst]) +apply(rule exists_fresh') +apply(simp add: fin_supp) +done + consts - substn :: "trm \ name \ coname \ trm \ trm" ("_[_:=<_>._]" [100,100,100,100] 100) - substc :: "trm \ coname \ name \ trm \ trm" ("_[_:=(_)._]" [100,100,100,100] 100) + substn :: "trm \ name \ coname \ trm \ trm" ("_{_:=<_>._}" [100,100,100,100] 100) + substc :: "trm \ coname \ name \ trm \ trm" ("_{_:=(_)._}" [100,100,100,100] 100) nominal_primrec (freshness_context: "(y::name,c::coname,P::trm)") - "(Ax x a)[y:=.P] = (if x=y then P[c\c>a] else Ax x a)" - "\a\(c,P,N);x\(y,c,P,M)\ \ - (Cut .M (x).N)[y:=.P] = Cut .(M[y:=.P]) (x).(N[y:=.P])" - "x\(y,c,P) \ - (NotR (x).M a)[y:=.P] = NotR (x).(M[y:=.P]) a" - "a\(c,P)\ - (NotL .M x)[y:=.P] = (if x=y then Cut .P (x).(NotL . (M[y:=.P]) x) - else NotL . (M[y:=.P]) x)" + "(Ax x a){y:=.P} = (if x=y then Cut .P (y).Ax y a else Ax x a)" + "\a\(c,P,N);x\(y,P,M)\ \ (Cut .M (x).N){y:=.P} = + (if M=Ax y a then Cut .P (x).(N{y:=.P}) else Cut .(M{y:=.P}) (x).(N{y:=.P}))" + "x\(y,P) \ (NotR (x).M a){y:=.P} = NotR (x).(M{y:=.P}) a" + "a\(c,P) \ (NotL .M x){y:=.P} = + (if x=y then fresh_fun (\x'. Cut .P (x').NotL .(M{y:=.P}) x') else NotL .(M{y:=.P}) x)" "\a\(c,P,N,d);b\(c,P,M,d);b\a\ \ - (AndR .M .N d)[y:=.P] = AndR .(M[y:=.P]) .(N[y:=.P]) d" - "x\(y,c,P,z) \ - (AndL1 (x).M z)[y:=.P] = (if z=y then Cut .P (z).AndL1 (x).(M[y:=.P]) z - else AndL1 (x).(M[y:=.P]) z)" - "x\(y,c,P,z) \ - (AndL2 (x).M z)[y:=.P] = (if z=y then Cut .P (z).AndL2 (x).(M[y:=.P]) z - else AndL2 (x).(M[y:=.P]) z)" - "a\(c,P,b) \ - (OrR1 .M b)[y:=.P] = OrR1 .(M[y:=.P]) b" - "a\(c,P,b) \ - (OrR2 .M b)[y:=.P] = OrR2 .(M[y:=.P]) b" - "\x\(y,N,c,P,z);u\(y,M,c,P,z);x\u\ \ - (OrL (x).M (u).N z)[y:=.P] = (if z=y then Cut .P (z).(OrL (x).(M[y:=.P]) (u).(N[y:=.P]) z) - else OrL (x).(M[y:=.P]) (u).(N[y:=.P]) z)" - "\a\(b,c,P); x\(y,c,P)\ \ - (ImpR (x)..M b)[y:=.P] = ImpR (x)..(M[y:=.P]) b" - "\a\(N,c,P);x\(y,c,P,M,z)\ \ - (ImpL .M (x).N z)[y:=.P] = (if y=z then Cut .P (z).(ImpL .(M[y:=.P]) (x).(N[y:=.P]) z) - else ImpL .(M[y:=.P]) (x).(N[y:=.P]) z)" + (AndR .M .N d){y:=.P} = AndR .(M{y:=.P}) .(N{y:=.P}) d" + "x\(y,P,z) \ (AndL1 (x).M z){y:=.P} = + (if z=y then fresh_fun (\z'. Cut .P (z').AndL1 (x).(M{y:=.P}) z') + else AndL1 (x).(M{y:=.P}) z)" + "x\(y,P,z) \ (AndL2 (x).M z){y:=.P} = + (if z=y then fresh_fun (\z'. Cut .P (z').AndL2 (x).(M{y:=.P}) z') + else AndL2 (x).(M{y:=.P}) z)" + "a\(c,P,b) \ (OrR1 .M b){y:=.P} = OrR1 .(M{y:=.P}) b" + "a\(c,P,b) \ (OrR2 .M b){y:=.P} = OrR2 .(M{y:=.P}) b" + "\x\(y,N,P,z);u\(y,M,P,z);x\u\ \ (OrL (x).M (u).N z){y:=.P} = + (if z=y then fresh_fun (\z'. Cut .P (z').OrL (x).(M{y:=.P}) (u).(N{y:=.P}) z') + else OrL (x).(M{y:=.P}) (u).(N{y:=.P}) z)" + "\a\(b,c,P); x\(y,P)\ \ (ImpR (x)..M b){y:=.P} = ImpR (x)..(M{y:=.P}) b" + "\a\(N,c,P);x\(y,P,M,z)\ \ (ImpL .M (x).N z){y:=.P} = + (if y=z then fresh_fun (\z'. Cut .P (z').ImpL .(M{y:=.P}) (x).(N{y:=.P}) z') + else ImpL .(M{y:=.P}) (x).(N{y:=.P}) z)" +apply(finite_guess)+ +apply(rule TrueI)+ +apply(simp add: abs_fresh abs_supp)+ +apply(rule impI) +apply(subgoal_tac "\x::name. x\(x1,P,y1)", erule exE, simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_NotL abs_fresh fresh_atm) +apply(rule exists_fresh', simp add: fin_supp) +apply(simp add: abs_fresh abs_supp)+ +apply(rule impI) +apply(subgoal_tac "\x::name. x\(x1,P,y1)", erule exE, simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndL1 abs_fresh fresh_atm) +apply(rule exists_fresh', simp add: fin_supp) +apply(simp add: abs_fresh abs_supp)+ +apply(rule impI) +apply(subgoal_tac "\x::name. x\(x1,P,y1)", erule exE, simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndL2 abs_fresh fresh_atm) +apply(rule exists_fresh', simp add: fin_supp) +apply(simp add: abs_fresh abs_supp)+ +apply(rule impI) +apply(subgoal_tac "\x::name. x\(x1,P,y1,x3,y2)", erule exE, simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrL abs_fresh fresh_atm) +apply(rule exists_fresh', simp add: fin_supp) +apply(simp add: abs_fresh abs_supp)+ +apply(rule impI) +apply(subgoal_tac "\x::name. x\(x1,P,y1,x3,y2)", erule exE, simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrL abs_fresh fresh_atm) +apply(rule exists_fresh', simp add: fin_supp) +apply(simp add: abs_fresh abs_supp)+ +apply(rule impI) +apply(subgoal_tac "\x::name. x\(x3,P,y1,y2)", erule exE, simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_ImpL abs_fresh fresh_atm) +apply(rule exists_fresh', simp add: fin_supp) +apply(simp add: abs_fresh abs_supp)+ +apply(rule impI) +apply(subgoal_tac "\x::name. x\(x3,P,y1,y2)", erule exE, simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_ImpL abs_fresh fresh_atm) +apply(rule exists_fresh', simp add: fin_supp) +apply(fresh_guess)+ +done + +nominal_primrec (freshness_context: "(d::name,z::coname,P::trm)") + "(Ax x a){d:=(z).P} = (if d=a then Cut .(Ax x a) (z).P else Ax x a)" + "\a\(d,P,N);x\(z,P,M)\ \ (Cut .M (x).N){d:=(z).P} = + (if N=Ax x d then Cut .(M{d:=(z).P}) (z).P else Cut .(M{d:=(z).P}) (x).(N{d:=(z).P}))" + "x\(z,P) \ (NotR (x).M a){d:=(z).P} = + (if d=a then fresh_fun (\a'. Cut .NotR (x).(M{d:=(z).P}) a' (z).P) else NotR (x).(M{d:=(z).P}) a)" + "a\(d,P) \ (NotL .M x){d:=(z).P} = NotL .(M{d:=(z).P}) x" + "\a\(P,c,N,d);b\(P,c,M,d);b\a\ \ (AndR .M .N c){d:=(z).P} = + (if d=c then fresh_fun (\a'. Cut .(AndR .(M{d:=(z).P}) .(N{d:=(z).P}) a') (z).P) + else AndR .(M{d:=(z).P}) .(N{d:=(z).P}) c)" + "x\(y,z,P) \ (AndL1 (x).M y){d:=(z).P} = AndL1 (x).(M{d:=(z).P}) y" + "x\(y,P,z) \ (AndL2 (x).M y){d:=(z).P} = AndL2 (x).(M{d:=(z).P}) y" + "a\(d,P,b) \ (OrR1 .M b){d:=(z).P} = + (if d=b then fresh_fun (\a'. Cut .OrR1 .(M{d:=(z).P}) a' (z).P) else OrR1 .(M{d:=(z).P}) b)" + "a\(d,P,b) \ (OrR2 .M b){d:=(z).P} = + (if d=b then fresh_fun (\a'. Cut .OrR2 .(M{d:=(z).P}) a' (z).P) else OrR2 .(M{d:=(z).P}) b)" + "\x\(N,z,P,u);y\(M,z,P,u);x\y\ \ (OrL (x).M (y).N u){d:=(z).P} = + OrL (x).(M{d:=(z).P}) (y).(N{d:=(z).P}) u" + "\a\(b,d,P); x\(z,P)\ \ (ImpR (x)..M b){d:=(z).P} = + (if d=b then fresh_fun (\a'. Cut .ImpR (x)..(M{d:=(z).P}) a' (z).P) + else ImpR (x)..(M{d:=(z).P}) b)" + "\a\(N,d,P);x\(y,z,P,M)\ \ (ImpL .M (x).N y){d:=(z).P} = + ImpL .(M{d:=(z).P}) (x).(N{d:=(z).P}) y" apply(finite_guess)+ apply(rule TrueI)+ apply(simp add: abs_fresh abs_supp fs_name1 fs_coname1)+ -apply(fresh_guess)+ -done +apply(rule impI) +apply(subgoal_tac "\x::coname. x\(x1,P,y1)", erule exE, simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_NotR abs_fresh fresh_atm) +apply(rule exists_fresh', simp add: fin_supp) +apply(simp add: abs_fresh abs_supp)+ +apply(rule impI) +apply(subgoal_tac "\x::coname. x\(x1,P,y1,x3,y2)", erule exE, simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndR abs_fresh fresh_atm) +apply(rule exists_fresh', simp add: fin_supp) +apply(simp add: abs_fresh abs_supp)+ +apply(rule impI) +apply(subgoal_tac "\x::coname. x\(x1,P,y1,x3,y2)", erule exE, simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndR abs_fresh fresh_atm) +apply(rule exists_fresh', simp add: fin_supp) +apply(simp add: abs_fresh abs_supp)+ +apply(rule impI) +apply(subgoal_tac "\x::coname. x\(x1,P,y1)", erule exE, simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrR1 abs_fresh fresh_atm) +apply(rule exists_fresh', simp add: fin_supp) +apply(simp add: abs_fresh abs_supp)+ +apply(rule impI) +apply(subgoal_tac "\x::coname. x\(x1,P,y1)", erule exE, simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrR2 abs_fresh fresh_atm) +apply(rule exists_fresh', simp add: fin_supp) +apply(simp add: abs_fresh abs_supp)+ +apply(rule impI) +apply(subgoal_tac "\x::coname. x\(x1,P,x2,y1)", erule exE, simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_ImpR abs_fresh fresh_atm abs_supp) +apply(rule exists_fresh', simp add: fin_supp) +apply(simp add: abs_fresh abs_supp)+ +apply(rule impI) +apply(subgoal_tac "\x::coname. x\(x1,P,x2,y1)", erule exE, simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_ImpR abs_fresh fresh_atm) +apply(rule exists_fresh', simp add: fin_supp) +apply(simp add: abs_fresh abs_supp)+ +apply(fresh_guess add: abs_fresh fresh_prod)+ +done + +lemma csubst_eqvt[eqvt]: + fixes pi1::"name prm" + and pi2::"coname prm" + shows "pi1\(M{c:=(x).N}) = (pi1\M){(pi1\c):=(pi1\x).(pi1\N)}" + and "pi2\(M{c:=(x).N}) = (pi2\M){(pi2\c):=(pi2\x).(pi2\N)}" +apply(nominal_induct M avoiding: c x N rule: trm.induct) +apply(auto simp add: eq_bij fresh_bij eqvts) +apply(perm_simp)+ +done + +lemma nsubst_eqvt[eqvt]: + fixes pi1::"name prm" + and pi2::"coname prm" + shows "pi1\(M{x:=.N}) = (pi1\M){(pi1\x):=<(pi1\c)>.(pi1\N)}" + and "pi2\(M{x:=.N}) = (pi2\M){(pi2\x):=<(pi2\c)>.(pi2\N)}" +apply(nominal_induct M avoiding: c x N rule: trm.induct) +apply(auto simp add: eq_bij fresh_bij eqvts) +apply(perm_simp)+ +done + +lemma supp_subst1: + shows "supp (M{y:=.P}) \ ((supp M) - {y}) \ (supp P)" +apply(nominal_induct M avoiding: y P c rule: trm.induct) +apply(auto) +apply(auto simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(blast)+ +apply(subgoal_tac "\x'::name. x'\(trm{y:=.P},P)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_NotL abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm) +apply(blast) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(subgoal_tac "\x'::name. x'\(trm{y:=.P},P)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_NotL abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm) +apply(blast) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(blast)+ +apply(subgoal_tac "\x'::name. x'\(trm{y:=.P},P,name1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndL1 abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm) +apply(blast) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(subgoal_tac "\x'::name. x'\(trm{y:=.P},P,name1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndL1 abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(subgoal_tac "\x'::name. x'\(trm{y:=.P},P,name1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndL1 abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm) +apply(blast) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(subgoal_tac "\x'::name. x'\(trm{y:=.P},P,name1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndL2 abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm) +apply(blast) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(subgoal_tac "\x'::name. x'\(trm{y:=.P},P,name1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndL2 abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(subgoal_tac "\x'::name. x'\(trm{y:=.P},P,name1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndL2 abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm) +apply(blast) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(subgoal_tac "\x'::name. x'\(trm1{y:=.P},P,name1,trm2{y:=.P},name2)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrL abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm) +apply(blast) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(subgoal_tac "\x'::name. x'\(trm1{y:=.P},P,name1,trm2{y:=.P},name2)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrL abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm) +apply(blast) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(subgoal_tac "\x'::name. x'\(trm1{y:=.P},P,name1,trm2{y:=.P},name2)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrL abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm) +apply(blast) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(subgoal_tac "\x'::name. x'\(trm1{y:=.P},P,name1,trm2{y:=.P},name2)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrL abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm) +apply(blast) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(blast)+ +apply(subgoal_tac "\x'::name. x'\(trm1{name2:=.P},P,name1,trm2{name2:=.P})") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_ImpL abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm) +apply(blast) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(subgoal_tac "\x'::name. x'\(trm1{name2:=.P},P,name1,trm2{name2:=.P})") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_ImpL abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm) +apply(blast) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(subgoal_tac "\x'::name. x'\(trm1{name2:=.P},P,name1,trm2{name2:=.P})") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_ImpL abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm) +apply(blast) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(blast)+ +done + +lemma supp_subst2: + shows "supp (M{y:=.P}) \ supp (M) \ ((supp P) - {c})" +apply(nominal_induct M avoiding: y P c rule: trm.induct) +apply(auto) +apply(auto simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(blast)+ +apply(subgoal_tac "\x'::name. x'\(trm{y:=.P},P)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_NotL abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm) +apply(blast) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(subgoal_tac "\x'::name. x'\(trm{y:=.P},P)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_NotL abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(subgoal_tac "\x'::name. x'\(trm{y:=.P},P)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_NotL abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm) +apply(blast) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(blast)+ +apply(subgoal_tac "\x'::name. x'\(trm{y:=.P},P,name1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndL1 abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm) +apply(blast) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(subgoal_tac "\x'::name. x'\(trm{y:=.P},P,name1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndL1 abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm) +apply(blast) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(subgoal_tac "\x'::name. x'\(trm{y:=.P},P,name1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndL2 abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm) +apply(blast) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(subgoal_tac "\x'::name. x'\(trm{y:=.P},P,name1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndL2 abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm) +apply(blast) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(subgoal_tac "\x'::name. x'\(trm1{y:=.P},P,name1,trm2{y:=.P},name2)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrL abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm) +apply(blast) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(subgoal_tac "\x'::name. x'\(trm1{y:=.P},P,name1,trm2{y:=.P},name2)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrL abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm) +apply(blast) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(blast)+ +apply(subgoal_tac "\x'::name. x'\(trm1{name2:=.P},P,name1,trm2{name2:=.P})") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_ImpL abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm) +apply(blast) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(subgoal_tac "\x'::name. x'\(trm1{name2:=.P},P,name1,trm2{name2:=.P})") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_ImpL abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm) +apply(blast) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(subgoal_tac "\x'::name. x'\(trm1{name2:=.P},P,name1,trm2{name2:=.P})") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_ImpL abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm) +apply(blast) +apply(subgoal_tac "\x'::name. x'\(trm1{name2:=.P},P,name1,trm2{name2:=.P})") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_ImpL abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm) +apply(blast) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(blast)+ +done + +lemma supp_subst3: + shows "supp (M{c:=(x).P}) \ ((supp M) - {c}) \ (supp P)" +apply(nominal_induct M avoiding: x P c rule: trm.induct) +apply(auto) +apply(auto simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(blast)+ +apply(subgoal_tac "\x'::coname. x'\(trm{coname:=(x).P},P)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_NotR abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(blast) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(subgoal_tac "\x'::coname. x'\(trm{coname:=(x).P},P)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_NotR abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(blast) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(subgoal_tac "\x'::coname. x'\(trm1{coname3:=(x).P},P,trm2{coname3:=(x).P},coname1,coname2)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndR abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(blast) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(subgoal_tac "\x'::coname. x'\(trm1{coname3:=(x).P},P,trm2{coname3:=(x).P},coname1,coname2)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndR abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(blast) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(subgoal_tac "\x'::coname. x'\(trm1{coname3:=(x).P},P,trm2{coname3:=(x).P},coname1,coname2)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndR abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(blast) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(subgoal_tac "\x'::coname. x'\(trm1{coname3:=(x).P},P,trm2{coname3:=(x).P},coname1,coname2)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndR abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(blast) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(blast)+ +apply(subgoal_tac "\x'::coname. x'\(trm{coname2:=(x).P},P,coname1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrR1 abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(blast) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(subgoal_tac "\x'::coname. x'\(trm{coname2:=(x).P},P,coname1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrR1 abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(subgoal_tac "\x'::coname. x'\(trm{coname2:=(x).P},P,coname1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrR1 abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(blast) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(subgoal_tac "\x'::coname. x'\(trm{coname2:=(x).P},P,coname1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrR2 abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(blast) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(subgoal_tac "\x'::coname. x'\(trm{coname2:=(x).P},P,coname1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrR2 abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(subgoal_tac "\x'::coname. x'\(trm{coname2:=(x).P},P,coname1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrR2 abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(blast) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(blast)+ +apply(subgoal_tac "\x'::coname. x'\(trm{coname2:=(x).P},P,coname1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_ImpR abs_fresh abs_supp fin_supp fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(blast) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(subgoal_tac "\x'::coname. x'\(trm{coname2:=(x).P},P,coname1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_ImpR abs_fresh abs_supp fin_supp fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(subgoal_tac "\x'::coname. x'\(trm{coname2:=(x).P},P,coname1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_ImpR abs_fresh abs_supp fin_supp fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(blast) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(blast)+ +done + +lemma supp_subst4: + shows "supp (M{c:=(x).P}) \ (supp M) \ ((supp P) - {x})" +apply(nominal_induct M avoiding: x P c rule: trm.induct) +apply(auto) +apply(auto simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(blast)+ +apply(subgoal_tac "\x'::coname. x'\(trm{coname:=(x).P},P)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_NotR abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(blast) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(subgoal_tac "\x'::coname. x'\(trm{coname:=(x).P},P)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_NotR abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(subgoal_tac "\x'::coname. x'\(trm{coname:=(x).P},P)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_NotR abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(blast) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(subgoal_tac "\x'::coname. x'\(trm1{coname3:=(x).P},P,trm2{coname3:=(x).P},coname1,coname2)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndR abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(blast) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(subgoal_tac "\x'::coname. x'\(trm1{coname3:=(x).P},P,trm2{coname3:=(x).P},coname1,coname2)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndR abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(blast) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(blast)+ +apply(subgoal_tac "\x'::coname. x'\(trm{coname2:=(x).P},P,coname1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrR1 abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(blast) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(subgoal_tac "\x'::coname. x'\(trm{coname2:=(x).P},P,coname1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrR1 abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(blast) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(subgoal_tac "\x'::coname. x'\(trm{coname2:=(x).P},P,coname1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrR2 abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(blast) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(subgoal_tac "\x'::coname. x'\(trm{coname2:=(x).P},P,coname1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrR2 abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(blast) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(blast)+ +apply(subgoal_tac "\x'::coname. x'\(trm{coname2:=(x).P},P,coname1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_ImpR abs_fresh abs_supp fin_supp fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(blast) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(subgoal_tac "\x'::coname. x'\(trm{coname2:=(x).P},P,coname1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_ImpR abs_fresh abs_supp fin_supp fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(subgoal_tac "\x'::coname. x'\(trm{coname2:=(x).P},P,coname1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_ImpR abs_fresh abs_supp fin_supp fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(blast) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(blast)+ +done + +lemma supp_subst5: + shows "(supp M - {y}) \ supp (M{y:=.P})" +apply(nominal_induct M avoiding: y P c rule: trm.induct) +apply(auto) +apply(auto simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(blast)+ +apply(subgoal_tac "\x'::name. x'\(trm{y:=.P},P)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_NotL abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm) +apply(blast) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(blast) +apply(subgoal_tac "\x'::name. x'\(trm{y:=.P},P,name1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndL1 abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm) +apply(blast) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(subgoal_tac "\x'::name. x'\(trm{y:=.P},P,name1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndL2 abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm) +apply(blast) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(subgoal_tac "\x'::name. x'\(trm1{y:=.P},P,name1,trm2{y:=.P},name2)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrL abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm) +apply(blast) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(subgoal_tac "\x'::name. x'\(trm1{y:=.P},P,name1,trm2{y:=.P},name2)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrL abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm) +apply(blast) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(blast) +apply(subgoal_tac "\x'::name. x'\(trm1{name2:=.P},P,name1,trm2{name2:=.P})") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_ImpL abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm) +apply(blast) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(subgoal_tac "\x'::name. x'\(trm1{name2:=.P},P,name1,trm2{name2:=.P})") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_ImpL abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm) +apply(blast) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(blast) +done + +lemma supp_subst6: + shows "(supp M) \ ((supp (M{y:=.P}))::coname set)" +apply(nominal_induct M avoiding: y P c rule: trm.induct) +apply(auto) +apply(auto simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(blast)+ +apply(subgoal_tac "\x'::name. x'\(trm{y:=.P},P)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_NotL abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm) +apply(blast) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(blast) +apply(subgoal_tac "\x'::name. x'\(trm{y:=.P},P,name1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndL1 abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm) +apply(blast) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(subgoal_tac "\x'::name. x'\(trm{y:=.P},P,name1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndL2 abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm) +apply(blast) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(subgoal_tac "\x'::name. x'\(trm1{y:=.P},P,name1,trm2{y:=.P},name2)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrL abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm) +apply(blast) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(subgoal_tac "\x'::name. x'\(trm1{y:=.P},P,name1,trm2{y:=.P},name2)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrL abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm) +apply(blast) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(blast) +apply(subgoal_tac "\x'::name. x'\(trm1{name2:=.P},P,name1,trm2{name2:=.P})") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_ImpL abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm) +apply(blast) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(subgoal_tac "\x'::name. x'\(trm1{name2:=.P},P,name1,trm2{name2:=.P})") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_ImpL abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm) +apply(blast) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(blast) +done + +lemma supp_subst7: + shows "(supp M - {c}) \ supp (M{c:=(x).P})" +apply(nominal_induct M avoiding: x P c rule: trm.induct) +apply(auto) +apply(auto simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(blast)+ +apply(subgoal_tac "\x'::coname. x'\(trm{coname:=(x).P},P)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_NotR abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(blast) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(subgoal_tac "\x'::coname. x'\(trm1{coname3:=(x).P},P,trm2{coname3:=(x).P},coname1,coname2)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndR abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(blast) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(subgoal_tac "\x'::coname. x'\(trm1{coname3:=(x).P},P,trm2{coname3:=(x).P},coname1,coname2)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndR abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(blast) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(blast)+ +apply(subgoal_tac "\x'::coname. x'\(trm{coname2:=(x).P},P,coname1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrR1 abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(blast) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(subgoal_tac "\x'::coname. x'\(trm{coname2:=(x).P},P,coname1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrR2 abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(blast) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(blast)+ +apply(subgoal_tac "\x'::coname. x'\(trm{coname2:=(x).P},P,coname1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_ImpR abs_fresh abs_supp fin_supp fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(blast) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(blast) +done + +lemma supp_subst8: + shows "(supp M) \ ((supp (M{c:=(x).P}))::name set)" +apply(nominal_induct M avoiding: x P c rule: trm.induct) +apply(auto) +apply(auto simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(blast)+ +apply(subgoal_tac "\x'::coname. x'\(trm{coname:=(x).P},P)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_NotR abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(blast) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(subgoal_tac "\x'::coname. x'\(trm1{coname3:=(x).P},P,trm2{coname3:=(x).P},coname1,coname2)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndR abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(blast) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(subgoal_tac "\x'::coname. x'\(trm1{coname3:=(x).P},P,trm2{coname3:=(x).P},coname1,coname2)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndR abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(blast) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(blast)+ +apply(subgoal_tac "\x'::coname. x'\(trm{coname2:=(x).P},P,coname1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrR1 abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(blast) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(subgoal_tac "\x'::coname. x'\(trm{coname2:=(x).P},P,coname1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrR2 abs_fresh fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(blast) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(blast) +apply(subgoal_tac "\x'::coname. x'\(trm{coname2:=(x).P},P,coname1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_ImpR abs_fresh abs_supp fin_supp fresh_atm) +apply(simp add: fresh_def abs_supp trm.supp supp_atm fin_supp) +apply(blast) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(blast)+ +done + +lemmas subst_supp = supp_subst1 supp_subst2 supp_subst3 supp_subst4 + supp_subst5 supp_subst6 supp_subst7 supp_subst8 +lemma subst_fresh: + fixes x::"name" + and c::"coname" + shows "x\P \ x\M{x:=.P}" + and "b\P \ b\M{b:=(y).P}" + and "x\(M,P) \ x\M{y:=.P}" + and "x\M \ x\M{c:=(x).P}" + and "x\(M,P) \ x\M{c:=(y).P}" + and "b\(M,P) \ b\M{c:=(y).P}" + and "b\M \ b\M{y:=.P}" + and "b\(M,P) \ b\M{y:=.P}" +apply - +apply(insert subst_supp) +apply(simp_all add: fresh_def supp_prod) +apply(blast)+ +done + +lemma forget: + shows "x\M \ M{x:=.P} = M" + and "c\M \ M{c:=(x).P} = M" +apply(nominal_induct M avoiding: x c P rule: trm.induct) +apply(auto simp add: fresh_atm abs_fresh abs_supp fin_supp) +done + +lemma substc_rename1: + assumes a: "c\(M,a)" + shows "M{a:=(x).N} = ([(c,a)]\M){c:=(x).N}" +using a +proof(nominal_induct M avoiding: c a x N rule: trm.induct) + case (Ax z d) + then show ?case by (auto simp add: fresh_prod fresh_atm calc_atm trm.inject alpha) +next + case NotL + then show ?case by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod) +next + case (NotR y M d) + then show ?case + apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod) + apply(subgoal_tac "\a'::coname. a'\(N,M{d:=(x).N},([(c,d)]\M){c:=(x).N})") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_NotR) + apply(simp add: trm.inject alpha) + apply(rule exists_fresh'(2)[OF fs_coname1]) + done +next + case (AndR c1 M c2 M' c3) + then show ?case + apply(simp) + apply(auto) + apply(simp add: fresh_prod calc_atm fresh_atm abs_fresh) + apply(simp add: fresh_prod calc_atm fresh_atm abs_fresh fresh_left) + apply(subgoal_tac "\a'::coname. a'\(N,M{c3:=(x).N}, + M'{c3:=(x).N},c1,c2,c3,([(c,c3)]\M){c:=(x).N},([(c,c3)]\M'){c:=(x).N})") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_AndR) + apply (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh subst_fresh) + apply(rule exists_fresh'(2)[OF fs_coname1]) + apply(simp add: fresh_prod calc_atm fresh_atm abs_fresh fresh_left) + apply(simp add: fresh_prod calc_atm fresh_atm abs_fresh fresh_left) + apply(auto simp add: trm.inject alpha) + done +next + case AndL1 + then show ?case by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod) +next + case AndL2 + then show ?case by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod) +next + case (OrR1 d M e) + then show ?case + apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod) + apply(subgoal_tac "\a'::coname. a'\(N,M{e:=(x).N},([(c,e)]\M){c:=(x).N},d,e)") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_OrR1) + apply(simp add: trm.inject alpha) + apply(rule exists_fresh'(2)[OF fs_coname1]) + done +next + case (OrR2 d M e) + then show ?case + apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod) + apply(subgoal_tac "\a'::coname. a'\(N,M{e:=(x).N},([(c,e)]\M){c:=(x).N},d,e)") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_OrR2) + apply(simp add: trm.inject alpha) + apply(rule exists_fresh'(2)[OF fs_coname1]) + done +next + case (OrL x1 M x2 M' x3) + then show ?case + by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) +next + case ImpL + then show ?case + by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) +next + case (ImpR y d M e) + then show ?case + apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) + apply(subgoal_tac "\a'::coname. a'\(N,M{e:=(x).N},([(c,e)]\M){c:=(x).N},d,e)") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_ImpR) + apply(simp add: trm.inject alpha) + apply(rule exists_fresh'(2)[OF fs_coname1]) + done +next + case (Cut d M y M') + then show ?case + apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) + apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst]) + apply(simp add: calc_atm) + done +qed + +lemma substc_rename2: + assumes a: "y\(N,x)" + shows "M{a:=(x).N} = M{a:=(y).([(y,x)]\N)}" +using a +proof(nominal_induct M avoiding: a x y N rule: trm.induct) + case (Ax z d) + then show ?case + by (auto simp add: fresh_prod fresh_atm calc_atm trm.inject alpha perm_swap fresh_left) +next + case NotL + then show ?case + by (auto simp add: fresh_prod fresh_atm calc_atm trm.inject alpha perm_swap fresh_left) +next + case (NotR y M d) + then show ?case + apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) + apply(subgoal_tac "\a'::coname. a'\(N,M{d:=(y).([(y,x)]\N)},[(y,x)]\N)") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_NotR) + apply(simp add: trm.inject alpha perm_swap fresh_left calc_atm) + apply(rule exists_fresh'(2)[OF fs_coname1]) + done +next + case (AndR c1 M c2 M' c3) + then show ?case + apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) + apply(subgoal_tac + "\a'::coname. a'\(N,M{c3:=(y).([(y,x)]\N)},M'{c3:=(y).([(y,x)]\N)},[(y,x)]\N,c1,c2,c3)") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_AndR) + apply (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh subst_fresh perm_swap fresh_left) + apply(rule exists_fresh'(2)[OF fs_coname1]) + done +next + case AndL1 + then show ?case by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) +next + case AndL2 + then show ?case by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) +next + case (OrR1 d M e) + then show ?case + apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) + apply(subgoal_tac "\a'::coname. a'\(N,M{e:=(y).([(y,x)]\N)},[(y,x)]\N,d,e)") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_OrR1) + apply(simp add: trm.inject alpha perm_swap fresh_left calc_atm) + apply(rule exists_fresh'(2)[OF fs_coname1]) + done +next + case (OrR2 d M e) + then show ?case + apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) + apply(subgoal_tac "\a'::coname. a'\(N,M{e:=(y).([(y,x)]\N)},[(y,x)]\N,d,e)") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_OrR2) + apply(simp add: trm.inject alpha perm_swap fresh_left calc_atm) + apply(rule exists_fresh'(2)[OF fs_coname1]) + done +next + case (OrL x1 M x2 M' x3) + then show ?case + by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) +next + case ImpL + then show ?case + by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) +next + case (ImpR y d M e) + then show ?case + apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) + apply(subgoal_tac "\a'::coname. a'\(N,M{e:=(y).([(y,x)]\N)},[(y,x)]\N,d,e)") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_ImpR) + apply(simp add: trm.inject alpha perm_swap fresh_left calc_atm) + apply(rule exists_fresh'(2)[OF fs_coname1]) + done +next + case (Cut d M y M') + then show ?case + by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left perm_swap) +qed + +lemma substn_rename3: + assumes a: "y\(M,x)" + shows "M{x:=.N} = ([(y,x)]\M){y:=.N}" +using a +proof(nominal_induct M avoiding: a x y N rule: trm.induct) + case (Ax z d) + then show ?case by (auto simp add: fresh_prod fresh_atm calc_atm trm.inject alpha) +next + case NotR + then show ?case by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod) +next + case (NotL d M z) + then show ?case + apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod) + apply(subgoal_tac "\a'::name. a'\(N,M{x:=.N},([(y,x)]\M){y:=.N})") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_NotL) + apply(simp add: trm.inject alpha) + apply(rule exists_fresh'(1)[OF fs_name1]) + done +next + case (AndR c1 M c2 M' c3) + then show ?case + by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) +next + case OrR1 + then show ?case by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod) +next + case OrR2 + then show ?case by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod) +next + case (AndL1 u M v) + then show ?case + apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod) + apply(subgoal_tac "\a'::name. a'\(N,M{x:=.N},([(y,x)]\M){y:=.N},u,v)") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_AndL1) + apply(simp add: trm.inject alpha) + apply(rule exists_fresh'(1)[OF fs_name1]) + done +next + case (AndL2 u M v) + then show ?case + apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod) + apply(subgoal_tac "\a'::name. a'\(N,M{x:=.N},([(y,x)]\M){y:=.N},u,v)") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_AndL2) + apply(simp add: trm.inject alpha) + apply(rule exists_fresh'(1)[OF fs_name1]) + done +next + case (OrL x1 M x2 M' x3) + then show ?case + apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) + apply(subgoal_tac + "\a'::name. a'\(N,M{x:=.N},M'{x:=.N},([(y,x)]\M){y:=.N},([(y,x)]\M'){y:=.N},x1,x2)") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_OrL) + apply(simp add: trm.inject alpha) + apply(rule exists_fresh'(1)[OF fs_name1]) + done +next + case ImpR + then show ?case + by(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_left abs_supp fin_supp fresh_prod) +next + case (ImpL d M v M' u) + then show ?case + apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) + apply(subgoal_tac + "\a'::name. a'\(N,M{u:=.N},M'{u:=.N},([(y,u)]\M){y:=.N},([(y,u)]\M'){y:=.N},d,v)") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_ImpL) + apply(simp add: trm.inject alpha) + apply(rule exists_fresh'(1)[OF fs_name1]) + done +next + case (Cut d M y M') + then show ?case + apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) + apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst]) + apply(simp add: calc_atm) + done +qed + +lemma substn_rename4: + assumes a: "c\(N,a)" + shows "M{x:=.N} = M{x:=.([(c,a)]\N)}" +using a +proof(nominal_induct M avoiding: x c a N rule: trm.induct) + case (Ax z d) + then show ?case + by (auto simp add: fresh_prod fresh_atm calc_atm trm.inject alpha perm_swap fresh_left) +next + case NotR + then show ?case + by (auto simp add: fresh_prod fresh_atm calc_atm trm.inject alpha perm_swap fresh_left) +next + case (NotL d M y) + then show ?case + apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) + apply(subgoal_tac "\a'::name. a'\(N,M{x:=.([(c,a)]\N)},[(c,a)]\N)") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_NotL) + apply(simp add: trm.inject alpha perm_swap fresh_left calc_atm) + apply(rule exists_fresh'(1)[OF fs_name1]) + done +next + case (OrL x1 M x2 M' x3) + then show ?case + apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) + apply(subgoal_tac + "\a'::name. a'\(N,M{x:=.([(c,a)]\N)},M'{x:=.([(c,a)]\N)},[(c,a)]\N,x1,x2,x3)") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_OrL) + apply (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh subst_fresh perm_swap fresh_left) + apply(rule exists_fresh'(1)[OF fs_name1]) + done +next + case OrR1 + then show ?case by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) +next + case OrR2 + then show ?case by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) +next + case (AndL1 u M v) + then show ?case + apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) + apply(subgoal_tac "\a'::name. a'\(N,M{x:=.([(c,a)]\N)},[(c,a)]\N,u,v)") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_AndL1) + apply(simp add: trm.inject alpha perm_swap fresh_left calc_atm) + apply(rule exists_fresh'(1)[OF fs_name1]) + done +next + case (AndL2 u M v) + then show ?case + apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) + apply(subgoal_tac "\a'::name. a'\(N,M{x:=.([(c,a)]\N)},[(c,a)]\N,u,v)") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_AndL2) + apply(simp add: trm.inject alpha perm_swap fresh_left calc_atm) + apply(rule exists_fresh'(1)[OF fs_name1]) + done +next + case (AndR c1 M c2 M' c3) + then show ?case + by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) +next + case ImpR + then show ?case + by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) +next + case (ImpL d M y M' u) + then show ?case + apply(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) + apply(subgoal_tac "\a'::name. a'\(N,M{u:=.([(c,a)]\N)},M'{u:=.([(c,a)]\N)},[(c,a)]\N,y,u)") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_ImpL) + apply(simp add: trm.inject alpha perm_swap fresh_left calc_atm) + apply(rule exists_fresh'(1)[OF fs_name1]) + done +next + case (Cut d M y M') + then show ?case + by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left perm_swap) +qed + +lemma subst_rename5: + assumes a: "c'\(c,N)" "x'\(x,M)" + shows "M{x:=.N} = ([(x',x)]\M){x':=.([(c',c)]\N)}" +proof - + have "M{x:=.N} = ([(x',x)]\M){x':=.N}" using a by (simp add: substn_rename3) + also have "\ = ([(x',x)]\M){x':=.([(c',c)]\N)}" using a by (simp add: substn_rename4) + finally show ?thesis by simp +qed + +lemma subst_rename6: + assumes a: "c'\(c,M)" "x'\(x,N)" + shows "M{c:=(x).N} = ([(c',c)]\M){c':=(x').([(x',x)]\N)}" +proof - + have "M{c:=(x).N} = ([(c',c)]\M){c':=(x).N}" using a by (simp add: substc_rename1) + also have "\ = ([(c',c)]\M){c':=(x').([(x',x)]\N)}" using a by (simp add: substc_rename2) + finally show ?thesis by simp +qed + +lemmas subst_rename = substc_rename1 substc_rename2 substn_rename3 substn_rename4 subst_rename5 subst_rename6 + +lemma better_Cut_substn[simp]: + assumes a: "a\[c].P" "x\(y,P)" + shows "(Cut .M (x).N){y:=.P} = + (if M=Ax y a then Cut .P (x).(N{y:=.P}) else Cut .(M{y:=.P}) (x).(N{y:=.P}))" +proof - + obtain x'::"name" where fs1: "x'\(M,N,c,P,x,y)" by (rule exists_fresh(1), rule fin_supp, blast) + obtain a'::"coname" where fs2: "a'\(M,N,c,P,a)" by (rule exists_fresh(2), rule fin_supp, blast) + have eq1: "(Cut .M (x).N) = (Cut .([(a',a)]\M) (x').([(x',x)]\N))" + using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm) + have eq2: "(M=Ax y a) = (([(a',a)]\M)=Ax y a')" + apply(auto simp add: calc_atm) + apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst]) + apply(simp add: calc_atm) + done + have "(Cut .M (x).N){y:=.P} = (Cut .([(a',a)]\M) (x').([(x',x)]\N)){y:=.P}" + using eq1 by simp + also have "\ = (if ([(a',a)]\M)=Ax y a' then Cut .P (x').(([(x',x)]\N){y:=.P}) + else Cut .(([(a',a)]\M){y:=.P}) (x').(([(x',x)]\N){y:=.P}))" + using fs1 fs2 by (auto simp add: fresh_prod fresh_left calc_atm fresh_atm) + also have "\ =(if M=Ax y a then Cut .P (x).(N{y:=.P}) else Cut .(M{y:=.P}) (x).(N{y:=.P}))" + using fs1 fs2 a + apply - + apply(simp only: eq2[symmetric]) + apply(auto simp add: trm.inject) + apply(simp_all add: alpha fresh_atm fresh_prod subst_fresh) + apply(simp_all add: eqvts perm_fresh_fresh calc_atm) + apply(auto) + apply(rule subst_rename) + apply(simp add: fresh_prod fresh_atm) + apply(simp add: abs_fresh) + apply(simp add: perm_fresh_fresh) + done + finally show ?thesis by simp +qed + +lemma better_Cut_substc[simp]: + assumes a: "a\(c,P)" "x\[y].P" + shows "(Cut .M (x).N){c:=(y).P} = + (if N=Ax x c then Cut .(M{c:=(y).P}) (y).P else Cut .(M{c:=(y).P}) (x).(N{c:=(y).P}))" +proof - + obtain x'::"name" where fs1: "x'\(M,N,c,P,x,y)" by (rule exists_fresh(1), rule fin_supp, blast) + obtain a'::"coname" where fs2: "a'\(M,N,c,P,a)" by (rule exists_fresh(2), rule fin_supp, blast) + have eq1: "(Cut .M (x).N) = (Cut .([(a',a)]\M) (x').([(x',x)]\N))" + using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm) + have eq2: "(N=Ax x c) = (([(x',x)]\N)=Ax x' c)" + apply(auto simp add: calc_atm) + apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst]) + apply(simp add: calc_atm) + done + have "(Cut .M (x).N){c:=(y).P} = (Cut .([(a',a)]\M) (x').([(x',x)]\N)){c:=(y).P}" + using eq1 by simp + also have "\ = (if ([(x',x)]\N)=Ax x' c then Cut .(([(a',a)]\M){c:=(y).P}) (y).P + else Cut .(([(a',a)]\M){c:=(y).P}) (x').(([(x',x)]\N){c:=(y).P}))" + using fs1 fs2 by (simp add: fresh_prod fresh_left calc_atm fresh_atm trm.inject) + also have "\ =(if N=Ax x c then Cut .(M{c:=(y).P}) (y).P else Cut .(M{c:=(y).P}) (x).(N{c:=(y).P}))" + using fs1 fs2 a + apply - + apply(simp only: eq2[symmetric]) + apply(auto simp add: trm.inject) + apply(simp_all add: alpha fresh_atm fresh_prod subst_fresh) + apply(simp_all add: eqvts perm_fresh_fresh calc_atm) + apply(auto) + apply(rule subst_rename) + apply(simp add: fresh_prod fresh_atm) + apply(simp add: abs_fresh) + apply(simp add: perm_fresh_fresh) + done + finally show ?thesis by simp +qed + +lemma better_Cut_substn': + assumes a: "a\[c].P" "y\(N,x)" "M\Ax y a" + shows "(Cut .M (x).N){y:=.P} = Cut .(M{y:=.P}) (x).N" +using a +apply - +apply(generate_fresh "name") +apply(subgoal_tac "Cut .M (x).N = Cut .M (ca).([(ca,x)]\N)") +apply(simp) +apply(subgoal_tac"y\([(ca,x)]\N)") +apply(simp add: forget) +apply(simp add: trm.inject) +apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(simp add: trm.inject) +apply(rule sym) +apply(simp add: alpha fresh_prod fresh_atm) +done + +lemma better_NotR_substc: + assumes a: "d\M" + shows "(NotR (x).M d){d:=(z).P} = fresh_fun (\a'. Cut .NotR (x).M a' (z).P)" +using a +apply - +apply(generate_fresh "name") +apply(subgoal_tac "NotR (x).M d = NotR (c).([(c,x)]\M) d") +apply(auto simp add: fresh_left calc_atm forget) +apply(generate_fresh "coname") +apply(rule_tac f="fresh_fun" in arg_cong) +apply(simp add: expand_fun_eq) +apply(rule allI) +apply(simp add: trm.inject alpha fresh_prod fresh_atm) +apply(perm_simp add: trm.inject alpha fresh_prod fresh_atm fresh_left, auto) +done + +lemma better_NotL_substn: + assumes a: "y\M" + shows "(NotL .M y){y:=.P} = fresh_fun (\x'. Cut .P (x').NotL .M x')" +using a +apply - +apply(generate_fresh "coname") +apply(subgoal_tac "NotL .M y = NotL .([(ca,a)]\M) y") +apply(auto simp add: fresh_left calc_atm forget) +apply(generate_fresh "name") +apply(rule_tac f="fresh_fun" in arg_cong) +apply(simp add: expand_fun_eq) +apply(rule allI) +apply(simp add: trm.inject alpha fresh_prod fresh_atm) +apply(perm_simp add: trm.inject alpha fresh_prod fresh_atm fresh_left, auto) +done + +lemma better_AndL1_substn: + assumes a: "y\[x].M" + shows "(AndL1 (x).M y){y:=.P} = fresh_fun (\z'. Cut .P (z').AndL1 (x).M z')" +using a +apply - +apply(generate_fresh "name") +apply(subgoal_tac "AndL1 (x).M y = AndL1 (ca).([(ca,x)]\M) y") +apply(auto simp add: fresh_left calc_atm forget abs_fresh)[1] +apply(generate_fresh "name") +apply(rule_tac f="fresh_fun" in arg_cong) +apply(simp add: expand_fun_eq) +apply(rule allI) +apply(simp add: trm.inject alpha fresh_prod fresh_atm) +apply(rule forget) +apply(simp add: fresh_left calc_atm) +apply(rule_tac f="fresh_fun" in arg_cong) +apply(simp add: expand_fun_eq) +apply(rule allI) +apply(simp add: trm.inject alpha fresh_prod fresh_atm) +apply(rule forget) +apply(simp add: fresh_left calc_atm) +apply(perm_simp add: trm.inject alpha fresh_left calc_atm fresh_prod fresh_atm) +apply(auto) +done + +lemma better_AndL2_substn: + assumes a: "y\[x].M" + shows "(AndL2 (x).M y){y:=.P} = fresh_fun (\z'. Cut .P (z').AndL2 (x).M z')" +using a +apply - +apply(generate_fresh "name") +apply(subgoal_tac "AndL2 (x).M y = AndL2 (ca).([(ca,x)]\M) y") +apply(auto simp add: fresh_left calc_atm forget abs_fresh)[1] +apply(generate_fresh "name") +apply(rule_tac f="fresh_fun" in arg_cong) +apply(simp add: expand_fun_eq) +apply(rule allI) +apply(simp add: trm.inject alpha fresh_prod fresh_atm) +apply(rule forget) +apply(simp add: fresh_left calc_atm) +apply(rule_tac f="fresh_fun" in arg_cong) +apply(simp add: expand_fun_eq) +apply(rule allI) +apply(simp add: trm.inject alpha fresh_prod fresh_atm) +apply(rule forget) +apply(simp add: fresh_left calc_atm) +apply(perm_simp add: trm.inject alpha fresh_left calc_atm fresh_prod fresh_atm) +apply(auto) +done + +lemma better_AndR_substc: + assumes a: "c\([a].M,[b].N)" + shows "(AndR .M .N c){c:=(z).P} = fresh_fun (\a'. Cut .(AndR .M .N a') (z).P)" +using a +apply - +apply(generate_fresh "coname") +apply(generate_fresh "coname") +apply(subgoal_tac "AndR .M .N c = AndR .([(ca,a)]\M) .([(caa,b)]\N) c") +apply(auto simp add: fresh_left calc_atm forget abs_fresh)[1] +apply(rule trans) +apply(rule substc.simps) +apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(auto simp add: fresh_prod fresh_atm)[1] +apply(simp) +apply(rule_tac f="fresh_fun" in arg_cong) +apply(simp add: expand_fun_eq) +apply(rule allI) +apply(simp add: trm.inject alpha fresh_prod fresh_atm) +apply(rule conjI) +apply(rule forget) +apply(auto simp add: fresh_left calc_atm abs_fresh)[1] +apply(rule forget) +apply(auto simp add: fresh_left calc_atm abs_fresh)[1] +apply(perm_simp add: trm.inject alpha fresh_left calc_atm fresh_prod fresh_atm) +apply(auto) +done + +lemma better_OrL_substn: + assumes a: "x\([y].M,[z].N)" + shows "(OrL (y).M (z).N x){x:=.P} = fresh_fun (\z'. Cut .P (z').OrL (y).M (z).N z')" +using a +apply - +apply(generate_fresh "name") +apply(generate_fresh "name") +apply(subgoal_tac "OrL (y).M (z).N x = OrL (ca).([(ca,y)]\M) (caa).([(caa,z)]\N) x") +apply(auto simp add: fresh_left calc_atm forget abs_fresh)[1] +apply(rule trans) +apply(rule substn.simps) +apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(auto simp add: fresh_prod fresh_atm)[1] +apply(simp) +apply(rule_tac f="fresh_fun" in arg_cong) +apply(simp add: expand_fun_eq) +apply(rule allI) +apply(simp add: trm.inject alpha fresh_prod fresh_atm) +apply(rule conjI) +apply(rule forget) +apply(auto simp add: fresh_left calc_atm abs_fresh)[1] +apply(rule forget) +apply(auto simp add: fresh_left calc_atm abs_fresh)[1] +apply(perm_simp add: trm.inject alpha fresh_left calc_atm fresh_prod fresh_atm) +apply(auto) +done + +lemma better_OrR1_substc: + assumes a: "d\[a].M" + shows "(OrR1 .M d){d:=(z).P} = fresh_fun (\a'. Cut .OrR1 .M a' (z).P)" +using a +apply - +apply(generate_fresh "coname") +apply(subgoal_tac "OrR1 .M d = OrR1 .([(c,a)]\M) d") +apply(auto simp add: fresh_left calc_atm forget abs_fresh)[1] +apply(rule_tac f="fresh_fun" in arg_cong) +apply(simp add: expand_fun_eq) +apply(rule allI) +apply(simp add: trm.inject alpha fresh_prod fresh_atm) +apply(rule forget) +apply(simp add: fresh_left calc_atm) +apply(rule_tac f="fresh_fun" in arg_cong) +apply(simp add: expand_fun_eq) +apply(rule allI) +apply(simp add: trm.inject alpha fresh_prod fresh_atm) +apply(rule forget) +apply(simp add: fresh_left calc_atm) +apply(perm_simp add: trm.inject alpha fresh_left calc_atm fresh_prod fresh_atm) +apply(auto) +done + +lemma better_OrR2_substc: + assumes a: "d\[a].M" + shows "(OrR2 .M d){d:=(z).P} = fresh_fun (\a'. Cut .OrR2 .M a' (z).P)" +using a +apply - +apply(generate_fresh "coname") +apply(subgoal_tac "OrR2 .M d = OrR2 .([(c,a)]\M) d") +apply(auto simp add: fresh_left calc_atm forget abs_fresh)[1] +apply(rule_tac f="fresh_fun" in arg_cong) +apply(simp add: expand_fun_eq) +apply(rule allI) +apply(simp add: trm.inject alpha fresh_prod fresh_atm) +apply(rule forget) +apply(simp add: fresh_left calc_atm) +apply(rule_tac f="fresh_fun" in arg_cong) +apply(simp add: expand_fun_eq) +apply(rule allI) +apply(simp add: trm.inject alpha fresh_prod fresh_atm) +apply(rule forget) +apply(simp add: fresh_left calc_atm) +apply(perm_simp add: trm.inject alpha fresh_left calc_atm fresh_prod fresh_atm) +apply(auto) +done + +lemma better_ImpR_substc: + assumes a: "d\[a].M" + shows "(ImpR (x)..M d){d:=(z).P} = fresh_fun (\a'. Cut .ImpR (x)..M a' (z).P)" +using a +apply - +apply(generate_fresh "coname") +apply(generate_fresh "name") +apply(subgoal_tac "ImpR (x)..M d = ImpR (ca)..([(c,a)]\[(ca,x)]\M) d") +apply(auto simp add: fresh_left calc_atm forget abs_fresh)[1] +apply(rule_tac f="fresh_fun" in arg_cong) +apply(simp add: expand_fun_eq) +apply(rule allI) +apply(simp add: trm.inject alpha fresh_prod fresh_atm abs_perm abs_fresh fresh_left calc_atm) +apply(rule forget) +apply(simp add: fresh_left calc_atm) +apply(rule_tac f="fresh_fun" in arg_cong) +apply(simp add: expand_fun_eq) +apply(rule allI) +apply(simp add: trm.inject alpha fresh_prod fresh_atm abs_perm fresh_left calc_atm abs_fresh) +apply(rule forget) +apply(simp add: fresh_left calc_atm) +apply(rule sym) +apply(perm_simp add: trm.inject alpha fresh_left calc_atm fresh_prod fresh_atm abs_fresh abs_perm) +done + +lemma better_ImpL_substn: + assumes a: "y\(M,[x].N)" + shows "(ImpL .M (x).N y){y:=.P} = fresh_fun (\z'. Cut .P (z').ImpL .M (x).N z')" +using a +apply - +apply(generate_fresh "coname") +apply(generate_fresh "name") +apply(subgoal_tac "ImpL .M (x).N y = ImpL .([(ca,a)]\M) (caa).([(caa,x)]\N) y") +apply(auto simp add: fresh_left calc_atm forget abs_fresh)[1] +apply(rule_tac f="fresh_fun" in arg_cong) +apply(simp add: expand_fun_eq) +apply(rule allI) +apply(simp add: trm.inject alpha fresh_prod fresh_atm abs_perm abs_fresh fresh_left calc_atm) +apply(rule forget) +apply(simp add: fresh_left calc_atm) +apply(auto)[1] +apply(rule sym) +apply(perm_simp add: trm.inject alpha fresh_left calc_atm fresh_prod fresh_atm abs_fresh abs_perm) +done + +lemma freshn_after_substc: + fixes x::"name" + assumes a: "x\M{c:=(y).P}" + shows "x\M" +using a supp_subst8 +apply(simp add: fresh_def) +apply(blast) +done + +lemma freshn_after_substn: + fixes x::"name" + assumes a: "x\M{y:=.P}" "x\y" + shows "x\M" +using a +using a supp_subst5 +apply(simp add: fresh_def) +apply(blast) +done + +lemma freshc_after_substc: + fixes a::"coname" + assumes a: "a\M{c:=(y).P}" "a\c" + shows "a\M" +using a supp_subst7 +apply(simp add: fresh_def) +apply(blast) +done + +lemma freshc_after_substn: + fixes a::"coname" + assumes a: "a\M{y:=.P}" + shows "a\M" +using a supp_subst6 +apply(simp add: fresh_def) +apply(blast) +done + +lemma substn_crename_comm: + assumes a: "c\a" "c\b" + shows "M{x:=.P}[a\c>b] = M[a\c>b]{x:=.(P[a\c>b])}" +using a +apply(nominal_induct M avoiding: x c P a b rule: trm.induct) +apply(auto simp add: subst_fresh rename_fresh trm.inject) +apply(subgoal_tac "\x'::name. x'\(P,x,c)") +apply(erule exE) +apply(subgoal_tac "Cut .P (x).Ax x a = Cut .P (x').Ax x' a") +apply(simp) +apply(rule trans) +apply(rule crename.simps) +apply(simp add: fresh_prod fresh_atm) +apply(simp) +apply(simp add: trm.inject) +apply(simp add: alpha trm.inject calc_atm fresh_atm) +apply(simp add: trm.inject) +apply(simp add: alpha trm.inject calc_atm fresh_atm) +apply(rule exists_fresh') +apply(rule fin_supp) +apply(rule trans) +apply(rule better_crename_Cut) +apply(simp add: fresh_atm) +apply(simp) +apply(simp add: crename_id) +apply(rule trans) +apply(rule better_crename_Cut) +apply(simp add: fresh_atm fresh_prod) +apply(simp add: rename_fresh fresh_atm) +apply(auto simp add: fresh_atm)[1] +apply(rule trans) +apply(rule better_crename_Cut) +apply(simp add: fresh_atm) +apply(auto simp add: fresh_atm)[1] +apply(drule crename_ax) +apply(simp add: fresh_atm) +apply(simp add: fresh_atm) +apply(simp) +apply(subgoal_tac "\x'::name. x'\(trm{x:=.P},P,P[a\c>b],x,trm[a\c>b]{x:=.P[a\c>b]})") +apply(erule exE, simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_NotL) +apply(rule trans) +apply(rule better_crename_Cut) +apply(simp add: fresh_atm fresh_prod) +apply(simp add: rename_fresh fresh_atm) +apply(rule exists_fresh') +apply(rule fin_supp) +apply(subgoal_tac "\x'::name. x'\(trm{x:=.P},P,P[a\c>b],name1,trm[a\c>b]{x:=.P[a\c>b]})") +apply(erule exE, simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndL1) +apply(rule trans) +apply(rule better_crename_Cut) +apply(simp add: fresh_atm fresh_prod) +apply(simp add: rename_fresh fresh_atm) +apply(rule exists_fresh') +apply(rule fin_supp) +apply(subgoal_tac "\x'::name. x'\(trm{x:=.P},P,P[a\c>b],name1,trm[a\c>b]{x:=.P[a\c>b]})") +apply(erule exE, simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndL2) +apply(rule trans) +apply(rule better_crename_Cut) +apply(simp add: fresh_atm fresh_prod) +apply(simp add: rename_fresh fresh_atm) +apply(rule exists_fresh') +apply(rule fin_supp) +apply(subgoal_tac "\x'::name. x'\(trm1{x:=.P},trm2{x:=.P},P,P[a\c>b],name1,name2, + trm1[a\c>b]{x:=.P[a\c>b]},trm2[a\c>b]{x:=.P[a\c>b]})") +apply(erule exE, simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrL) +apply(rule trans) +apply(rule better_crename_Cut) +apply(simp add: fresh_atm fresh_prod) +apply(simp add: rename_fresh subst_fresh fresh_atm) +apply(rule exists_fresh') +apply(rule fin_supp) +apply(subgoal_tac "\x'::name. x'\(trm1{name2:=.P},trm2{name2:=.P},P,P[a\c>b],name1, + trm1[a\c>b]{name2:=.P[a\c>b]},trm2[a\c>b]{name2:=.P[a\c>b]})") +apply(erule exE, simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_ImpL) +apply(rule trans) +apply(rule better_crename_Cut) +apply(simp add: fresh_atm fresh_prod) +apply(simp add: rename_fresh subst_fresh fresh_atm) +apply(rule exists_fresh') +apply(rule fin_supp) +done + +lemma substc_crename_comm: + assumes a: "c\a" "c\b" + shows "M{c:=(x).P}[a\c>b] = M[a\c>b]{c:=(x).(P[a\c>b])}" +using a +apply(nominal_induct M avoiding: x c P a b rule: trm.induct) +apply(auto simp add: subst_fresh rename_fresh trm.inject) +apply(rule trans) +apply(rule better_crename_Cut) +apply(simp add: fresh_atm fresh_prod) +apply(simp add: rename_fresh fresh_atm) +apply(rule trans) +apply(rule better_crename_Cut) +apply(simp add: fresh_atm fresh_prod) +apply(simp add: rename_fresh fresh_atm) +apply(drule crename_ax) +apply(simp add: fresh_atm) +apply(simp add: fresh_atm) +apply(simp) +apply(subgoal_tac "\c'::coname. c'\(a,b,trm{coname:=(x).P},P,P[a\c>b],x,trm[a\c>b]{coname:=(x).P[a\c>b]})") +apply(erule exE, simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_NotR) +apply(rule trans) +apply(rule better_crename_Cut) +apply(simp add: fresh_atm fresh_prod) +apply(simp add: rename_fresh fresh_atm) +apply(rule exists_fresh') +apply(rule fin_supp) +apply(subgoal_tac "\c'::coname. c'\(coname1,coname2,a,b,trm1{coname3:=(x).P},trm2{coname3:=(x).P}, + P,P[a\c>b],x,trm1[a\c>b]{coname3:=(x).P[a\c>b]},trm2[a\c>b]{coname3:=(x).P[a\c>b]})") +apply(erule exE, simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndR) +apply(rule trans) +apply(rule better_crename_Cut) +apply(simp add: fresh_atm fresh_prod) +apply(simp add: rename_fresh subst_fresh fresh_atm) +apply(rule exists_fresh') +apply(rule fin_supp) +apply(subgoal_tac "\c'::coname. c'\(coname1,trm{coname2:=(x).P},P,P[a\c>b],a,b, + trm[a\c>b]{coname2:=(x).P[a\c>b]})") +apply(erule exE, simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrR1) +apply(rule trans) +apply(rule better_crename_Cut) +apply(simp add: fresh_atm fresh_prod) +apply(simp add: rename_fresh fresh_atm) +apply(rule exists_fresh') +apply(rule fin_supp) +apply(subgoal_tac "\c'::coname. c'\(coname1,trm{coname2:=(x).P},P,P[a\c>b],a,b, + trm[a\c>b]{coname2:=(x).P[a\c>b]})") +apply(erule exE, simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrR2) +apply(rule trans) +apply(rule better_crename_Cut) +apply(simp add: fresh_atm fresh_prod) +apply(simp add: rename_fresh fresh_atm) +apply(rule exists_fresh') +apply(rule fin_supp) +apply(subgoal_tac "\c'::coname. c'\(coname1,trm{coname2:=(x).P},P,P[a\c>b],a,b, + trm[a\c>b]{coname2:=(x).P[a\c>b]})") +apply(erule exE, simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_ImpR) +apply(rule trans) +apply(rule better_crename_Cut) +apply(simp add: fresh_atm fresh_prod) +apply(simp add: rename_fresh fresh_atm) +apply(rule exists_fresh') +apply(rule fin_supp) +done + +lemma substn_nrename_comm: + assumes a: "x\y" "x\z" + shows "M{x:=.P}[y\n>z] = M[y\n>z]{x:=.(P[y\n>z])}" +using a +apply(nominal_induct M avoiding: x c P y z rule: trm.induct) +apply(auto simp add: subst_fresh rename_fresh trm.inject) +apply(rule trans) +apply(rule better_nrename_Cut) +apply(simp add: fresh_prod fresh_atm) +apply(simp add: trm.inject) +apply(rule trans) +apply(rule better_nrename_Cut) +apply(simp add: fresh_atm) +apply(simp) +apply(drule nrename_ax) +apply(simp add: fresh_atm) +apply(simp add: fresh_atm) +apply(simp) +apply(subgoal_tac "\x'::name. x'\(y,z,trm{x:=.P},P,P[y\n>z],x,trm[y\n>z]{x:=.P[y\n>z]})") +apply(erule exE, simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_NotL) +apply(rule trans) +apply(rule better_nrename_Cut) +apply(simp add: fresh_atm fresh_prod) +apply(simp add: fresh_atm) +apply(rule exists_fresh') +apply(rule fin_supp) +apply(subgoal_tac "\x'::name. x'\(trm{x:=.P},P,P[y\n>z],name1,trm[y\n>z]{x:=.P[y\n>z]},y,z)") +apply(erule exE, simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndL1) +apply(rule trans) +apply(rule better_nrename_Cut) +apply(simp add: fresh_atm fresh_prod) +apply(simp add: rename_fresh fresh_atm) +apply(rule exists_fresh') +apply(rule fin_supp) +apply(subgoal_tac "\x'::name. x'\(y,z,trm{x:=.P},P,P[y\n>z],name1,trm[y\n>z]{x:=.P[y\n>z]})") +apply(erule exE, simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndL2) +apply(rule trans) +apply(rule better_nrename_Cut) +apply(simp add: fresh_atm fresh_prod) +apply(simp add: rename_fresh fresh_atm) +apply(rule exists_fresh') +apply(rule fin_supp) +apply(subgoal_tac "\x'::name. x'\(trm1{x:=.P},trm2{x:=.P},P,P[y\n>z],name1,name2,y,z, + trm1[y\n>z]{x:=.P[y\n>z]},trm2[y\n>z]{x:=.P[y\n>z]})") +apply(erule exE, simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrL) +apply(rule trans) +apply(rule better_nrename_Cut) +apply(simp add: fresh_atm fresh_prod) +apply(simp add: rename_fresh subst_fresh fresh_atm) +apply(rule exists_fresh') +apply(rule fin_supp) +apply(subgoal_tac "\x'::name. x'\(trm1{name2:=.P},trm2{name2:=.P},P,P[y\n>z],y,z,name1, + trm1[y\n>z]{name2:=.P[y\n>z]},trm2[y\n>z]{name2:=.P[y\n>z]})") +apply(erule exE, simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_ImpL) +apply(rule trans) +apply(rule better_nrename_Cut) +apply(simp add: fresh_atm fresh_prod) +apply(simp add: rename_fresh subst_fresh fresh_atm) +apply(rule exists_fresh') +apply(rule fin_supp) +done + +lemma substc_nrename_comm: + assumes a: "x\y" "x\z" + shows "M{c:=(x).P}[y\n>z] = M[y\n>z]{c:=(x).(P[y\n>z])}" +using a +apply(nominal_induct M avoiding: x c P y z rule: trm.induct) +apply(auto simp add: subst_fresh rename_fresh trm.inject) +apply(rule trans) +apply(rule better_nrename_Cut) +apply(simp add: fresh_atm fresh_prod) +apply(simp add: rename_fresh fresh_atm) +apply(rule trans) +apply(rule better_nrename_Cut) +apply(simp add: fresh_atm fresh_prod) +apply(simp add: rename_fresh fresh_atm) +apply(rule trans) +apply(rule better_nrename_Cut) +apply(simp add: fresh_atm fresh_prod) +apply(simp add: rename_fresh fresh_atm) +apply(rule trans) +apply(rule better_nrename_Cut) +apply(simp add: fresh_atm fresh_prod) +apply(simp add: rename_fresh fresh_atm) +apply(drule nrename_ax) +apply(simp add: fresh_atm) +apply(simp add: fresh_atm) +apply(simp) +apply(rule trans) +apply(rule better_nrename_Cut) +apply(simp add: fresh_atm fresh_prod) +apply(simp add: rename_fresh fresh_atm) +apply(drule nrename_ax) +apply(simp add: fresh_atm) +apply(simp add: fresh_atm) +apply(simp) +apply(subgoal_tac "\c'::coname. c'\(y,z,trm{coname:=(x).P},P,P[y\n>z],x,trm[y\n>z]{coname:=(x).P[y\n>z]})") +apply(erule exE, simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_NotR) +apply(rule trans) +apply(rule better_nrename_Cut) +apply(simp add: fresh_atm fresh_prod) +apply(simp add: rename_fresh fresh_atm) +apply(rule exists_fresh') +apply(rule fin_supp) +apply(subgoal_tac "\c'::coname. c'\(coname1,coname2,y,z,trm1{coname3:=(x).P},trm2{coname3:=(x).P}, + P,P[y\n>z],x,trm1[y\n>z]{coname3:=(x).P[y\n>z]},trm2[y\n>z]{coname3:=(x).P[y\n>z]})") +apply(erule exE, simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndR) +apply(rule trans) +apply(rule better_nrename_Cut) +apply(simp add: fresh_atm fresh_prod) +apply(simp add: rename_fresh subst_fresh fresh_atm) +apply(rule exists_fresh') +apply(rule fin_supp) +apply(subgoal_tac "\c'::coname. c'\(coname1,trm{coname2:=(x).P},P,P[y\n>z],y,z, + trm[y\n>z]{coname2:=(x).P[y\n>z]})") +apply(erule exE, simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrR1) +apply(rule trans) +apply(rule better_nrename_Cut) +apply(simp add: fresh_atm fresh_prod) +apply(simp add: rename_fresh fresh_atm) +apply(rule exists_fresh') +apply(rule fin_supp) +apply(subgoal_tac "\c'::coname. c'\(coname1,trm{coname2:=(x).P},P,P[y\n>z],y,z, + trm[y\n>z]{coname2:=(x).P[y\n>z]})") +apply(erule exE, simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrR2) +apply(rule trans) +apply(rule better_nrename_Cut) +apply(simp add: fresh_atm fresh_prod) +apply(simp add: rename_fresh fresh_atm) +apply(rule exists_fresh') +apply(rule fin_supp) +apply(subgoal_tac "\c'::coname. c'\(coname1,trm{coname2:=(x).P},P,P[y\n>z],y,z, + trm[y\n>z]{coname2:=(x).P[y\n>z]})") +apply(erule exE, simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_ImpR) +apply(rule trans) +apply(rule better_nrename_Cut) +apply(simp add: fresh_atm fresh_prod) +apply(simp add: rename_fresh fresh_atm) +apply(rule exists_fresh') +apply(rule fin_supp) +done + +lemma substn_crename_comm': + assumes a: "a\c" "a\P" + shows "M{x:=.P}[a\c>b] = M[a\c>b]{x:=.P}" +using a +proof - + assume a1: "a\c" + assume a2: "a\P" + obtain c'::"coname" where fs2: "c'\(c,P,a,b)" by (rule exists_fresh(2), rule fin_supp, blast) + have eq: "M{x:=.P} = M{x:=.([(c',c)]\P)}" + using fs2 + apply - + apply(rule subst_rename) + apply(simp) + done + have eq': "M[a\c>b]{x:=.P} = M[a\c>b]{x:=.([(c',c)]\P)}" + using fs2 + apply - + apply(rule subst_rename) + apply(simp) + done + have eq2: "([(c',c)]\P)[a\c>b] = ([(c',c)]\P)" using fs2 a2 a1 + apply - + apply(rule rename_fresh) + apply(simp add: fresh_left calc_atm fresh_prod fresh_atm) + done + have "M{x:=.P}[a\c>b] = M{x:=.([(c',c)]\P)}[a\c>b]" using eq by simp + also have "\ = M[a\c>b]{x:=.(([(c',c)]\P)[a\c>b])}" + using fs2 + apply - + apply(rule substn_crename_comm) + apply(simp_all add: fresh_prod fresh_atm) + done + also have "\ = M[a\c>b]{x:=.(([(c',c)]\P))}" using eq2 by simp + also have "\ = M[a\c>b]{x:=.P}" using eq' by simp + finally show ?thesis by simp +qed + +lemma substc_crename_comm': + assumes a: "c\a" "c\b" "a\P" + shows "M{c:=(x).P}[a\c>b] = M[a\c>b]{c:=(x).P}" +using a +proof - + assume a1: "c\a" + assume a1': "c\b" + assume a2: "a\P" + obtain c'::"coname" where fs2: "c'\(c,M,a,b)" by (rule exists_fresh(2), rule fin_supp, blast) + have eq: "M{c:=(x).P} = ([(c',c)]\M){c':=(x).P}" + using fs2 + apply - + apply(rule subst_rename) + apply(simp) + done + have eq': "([(c',c)]\(M[a\c>b])){c':=(x).P} = M[a\c>b]{c:=(x).P}" + using fs2 + apply - + apply(rule subst_rename[symmetric]) + apply(simp add: rename_fresh) + done + have eq2: "([(c',c)]\M)[a\c>b] = ([(c',c)]\(M[a\c>b]))" using fs2 a2 a1 a1' + apply - + apply(simp add: rename_eqvts) + apply(simp add: fresh_left calc_atm fresh_prod fresh_atm) + done + have "M{c:=(x).P}[a\c>b] = ([(c',c)]\M){c':=(x).P}[a\c>b]" using eq by simp + also have "\ = ([(c',c)]\M)[a\c>b]{c':=(x).P[a\c>b]}" + using fs2 + apply - + apply(rule substc_crename_comm) + apply(simp_all add: fresh_prod fresh_atm) + done + also have "\ = ([(c',c)]\(M[a\c>b])){c':=(x).P[a\c>b]}" using eq2 by simp + also have "\ = ([(c',c)]\(M[a\c>b])){c':=(x).P}" using a2 by (simp add: rename_fresh) + also have "\ = M[a\c>b]{c:=(x).P}" using eq' by simp + finally show ?thesis by simp +qed + +lemma substn_nrename_comm': + assumes a: "x\y" "x\z" "y\P" + shows "M{x:=.P}[y\n>z] = M[y\n>z]{x:=.P}" +using a +proof - + assume a1: "x\y" + assume a1': "x\z" + assume a2: "y\P" + obtain x'::"name" where fs2: "x'\(x,M,y,z)" by (rule exists_fresh(1), rule fin_supp, blast) + have eq: "M{x:=.P} = ([(x',x)]\M){x':=.P}" + using fs2 + apply - + apply(rule subst_rename) + apply(simp) + done + have eq': "([(x',x)]\(M[y\n>z])){x':=.P} = M[y\n>z]{x:=.P}" + using fs2 + apply - + apply(rule subst_rename[symmetric]) + apply(simp add: rename_fresh) + done + have eq2: "([(x',x)]\M)[y\n>z] = ([(x',x)]\(M[y\n>z]))" using fs2 a2 a1 a1' + apply - + apply(simp add: rename_eqvts) + apply(simp add: fresh_left calc_atm fresh_prod fresh_atm) + done + have "M{x:=.P}[y\n>z] = ([(x',x)]\M){x':=.P}[y\n>z]" using eq by simp + also have "\ = ([(x',x)]\M)[y\n>z]{x':=.P[y\n>z]}" + using fs2 + apply - + apply(rule substn_nrename_comm) + apply(simp_all add: fresh_prod fresh_atm) + done + also have "\ = ([(x',x)]\(M[y\n>z])){x':=.P[y\n>z]}" using eq2 by simp + also have "\ = ([(x',x)]\(M[y\n>z])){x':=.P}" using a2 by (simp add: rename_fresh) + also have "\ = M[y\n>z]{x:=.P}" using eq' by simp + finally show ?thesis by simp +qed + +lemma substc_nrename_comm': + assumes a: "x\y" "y\P" + shows "M{c:=(x).P}[y\n>z] = M[y\n>z]{c:=(x).P}" +using a +proof - + assume a1: "x\y" + assume a2: "y\P" + obtain x'::"name" where fs2: "x'\(x,P,y,z)" by (rule exists_fresh(1), rule fin_supp, blast) + have eq: "M{c:=(x).P} = M{c:=(x').([(x',x)]\P)}" + using fs2 + apply - + apply(rule subst_rename) + apply(simp) + done + have eq': "M[y\n>z]{c:=(x).P} = M[y\n>z]{c:=(x').([(x',x)]\P)}" + using fs2 + apply - + apply(rule subst_rename) + apply(simp) + done + have eq2: "([(x',x)]\P)[y\n>z] = ([(x',x)]\P)" using fs2 a2 a1 + apply - + apply(rule rename_fresh) + apply(simp add: fresh_left calc_atm fresh_prod fresh_atm) + done + have "M{c:=(x).P}[y\n>z] = M{c:=(x').([(x',x)]\P)}[y\n>z]" using eq by simp + also have "\ = M[y\n>z]{c:=(x').(([(x',x)]\P)[y\n>z])}" + using fs2 + apply - + apply(rule substc_nrename_comm) + apply(simp_all add: fresh_prod fresh_atm) + done + also have "\ = M[y\n>z]{c:=(x').(([(x',x)]\P))}" using eq2 by simp + also have "\ = M[y\n>z]{c:=(x).P}" using eq' by simp + finally show ?thesis by simp +qed + +lemmas subst_comm = substn_crename_comm substc_crename_comm + substn_nrename_comm substc_nrename_comm +lemmas subst_comm' = substn_crename_comm' substc_crename_comm' + substn_nrename_comm' substc_nrename_comm' text {* typing contexts *} @@ -178,7 +2916,9 @@ validc :: "ctxtc \ bool" where vc1[intro]: "validc []" -| vc2[intro]: "\a\\; validc \\ \ validc ((a,T)#\)" +| vc2[intro]: "\a\\; validc \\ \ validc ((a,T)#\)" + +equivariance validc inductive2 validn :: "ctxtn \ bool" @@ -186,101 +2926,17915 @@ vn1[intro]: "validn []" | vn2[intro]: "\x\\; validn \\ \ validn ((x,T)#\)" +equivariance validn + +lemma fresh_ctxt: + fixes a::"coname" + and x::"name" + and \::"ctxtn" + and \::"ctxtc" + shows "a\\" and "x\\" +proof - + show "a\\" by (induct \) (auto simp add: fresh_list_nil fresh_list_cons fresh_prod fresh_atm fresh_ty) +next + show "x\\" by (induct \) (auto simp add: fresh_list_nil fresh_list_cons fresh_prod fresh_atm fresh_ty) +qed + +text {* cut-reductions *} + +declare abs_perm[eqvt] + +inductive2 + fin :: "trm \ name \ bool" +where + [intro]: "fin (Ax x a) x" +| [intro]: "x\M \ fin (NotL .M x) x" +| [intro]: "y\[x].M \ fin (AndL1 (x).M y) y" +| [intro]: "y\[x].M \ fin (AndL2 (x).M y) y" +| [intro]: "\z\[x].M;z\[y].N\ \ fin (OrL (x).M (y).N z) z" +| [intro]: "\y\M;y\[x].N\ \ fin (ImpL .M (x).N y) y" + +equivariance fin + +lemma fin_Ax_elim: + assumes a: "fin (Ax x a) y" + shows "x=y" +using a +apply(erule_tac fin.cases) +apply(auto simp add: trm.inject) +done + +lemma fin_NotL_elim: + assumes a: "fin (NotL .M x) y" + shows "x=y \ x\M" +using a +apply(erule_tac fin.cases) +apply(auto simp add: trm.inject) +apply(subgoal_tac "y\[aa].Ma") +apply(drule sym) +apply(simp_all add: abs_fresh) +done + +lemma fin_AndL1_elim: + assumes a: "fin (AndL1 (x).M y) z" + shows "z=y \ z\[x].M" +using a +apply(erule_tac fin.cases) +apply(auto simp add: trm.inject) +done + +lemma fin_AndL2_elim: + assumes a: "fin (AndL2 (x).M y) z" + shows "z=y \ z\[x].M" +using a +apply(erule_tac fin.cases) +apply(auto simp add: trm.inject) +done + +lemma fin_OrL_elim: + assumes a: "fin (OrL (x).M (y).N u) z" + shows "z=u \ z\[x].M \ z\[y].N" +using a +apply(erule_tac fin.cases) +apply(auto simp add: trm.inject) +done + +lemma fin_ImpL_elim: + assumes a: "fin (ImpL .M (x).N z) y" + shows "z=y \ z\M \ z\[x].N" +using a +apply(erule_tac fin.cases) +apply(auto simp add: trm.inject) +apply(subgoal_tac "y\[aa].Ma") +apply(drule sym) +apply(simp_all add: abs_fresh) +done + +lemma fin_rest_elims: + shows "fin (Cut .M (x).N) y \ False" + and "fin (NotR (x).M c) y \ False" + and "fin (AndR .M .N c) y \ False" + and "fin (OrR1 .M b) y \ False" + and "fin (OrR2 .M b) y \ False" + and "fin (ImpR (x)..M b) y \ False" +by (erule fin.cases, simp_all add: trm.inject)+ + +lemmas fin_elims = fin_Ax_elim fin_NotL_elim fin_AndL1_elim fin_AndL2_elim fin_OrL_elim + fin_ImpL_elim fin_rest_elims + +lemma fin_rename: + shows "fin M x \ fin ([(x',x)]\M) x'" +by (induct rule: fin.induct) + (auto simp add: calc_atm simp add: fresh_left abs_fresh) + +lemma not_fin_subst1: + assumes a: "\(fin M x)" + shows "\(fin (M{c:=(y).P}) x)" +using a +apply(nominal_induct M avoiding: x c y P rule: trm.induct) +apply(auto) +apply(drule fin_elims, simp) +apply(drule fin_elims, simp) +apply(drule fin_elims, simp) +apply(subgoal_tac "\a'::coname. a'\(trm{coname:=(y).P},P,x)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_NotR) +apply(erule fin.cases, simp_all add: trm.inject) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(drule fin_elims, simp) +apply(drule fin_elims) +apply(auto)[1] +apply(drule freshn_after_substc) +apply(simp add: fin.intros) +apply(subgoal_tac "\a'::coname. a'\(trm1{coname3:=(y).P},trm2{coname3:=(y).P},P,coname1,coname2,coname3,x)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndR) +apply(erule fin.cases, simp_all add: trm.inject) +apply(subgoal_tac "\a'::coname. a'\(trm1{coname3:=(y).P},trm2{coname3:=(y).P},P,coname1,coname2,coname3,x)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndR) +apply(erule fin.cases, simp_all add: trm.inject) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(erule fin.cases, simp_all add: trm.inject) +apply(drule fin_AndL1_elim) +apply(auto simp add: abs_fresh)[1] +apply(drule freshn_after_substc) +apply(subgoal_tac "name2\[name1]. trm") +apply(simp add: fin.intros) +apply(simp add: abs_fresh) +apply(drule fin_AndL2_elim) +apply(auto simp add: abs_fresh)[1] +apply(drule freshn_after_substc) +apply(subgoal_tac "name2\[name1].trm") +apply(simp add: fin.intros) +apply(simp add: abs_fresh) +apply(subgoal_tac "\a'::coname. a'\(trm{coname2:=(y).P},coname1,P,x)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrR1) +apply(erule fin.cases, simp_all add: trm.inject) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(erule fin.cases, simp_all add: trm.inject) +apply(subgoal_tac "\a'::coname. a'\(trm{coname2:=(y).P},coname1,P,x)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrR2) +apply(erule fin.cases, simp_all add: trm.inject) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(erule fin.cases, simp_all add: trm.inject) +apply(drule fin_OrL_elim) +apply(auto simp add: abs_fresh)[1] +apply(drule freshn_after_substc)+ +apply(subgoal_tac "name3\[name1].trm1 \ name3\[name2].trm2") +apply(simp add: fin.intros) +apply(simp add: abs_fresh) +apply(subgoal_tac "\a'::coname. a'\(trm{coname2:=(y).P},coname1,P,x)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_ImpR) +apply(erule fin.cases, simp_all add: trm.inject) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(erule fin.cases, simp_all add: trm.inject) +apply(drule fin_ImpL_elim) +apply(auto simp add: abs_fresh)[1] +apply(drule freshn_after_substc)+ +apply(subgoal_tac "x\[name1].trm2") +apply(simp add: fin.intros) +apply(simp add: abs_fresh) +done + +lemma not_fin_subst2: + assumes a: "\(fin M x)" + shows "\(fin (M{y:=.P}) x)" +using a +apply(nominal_induct M avoiding: x c y P rule: trm.induct) +apply(auto) +apply(erule fin.cases, simp_all add: trm.inject) +apply(erule fin.cases, simp_all add: trm.inject) +apply(erule fin.cases, simp_all add: trm.inject) +apply(erule fin.cases, simp_all add: trm.inject) +apply(subgoal_tac "\a'::name. a'\(trm{y:=.P},P,x)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_NotL) +apply(erule fin.cases, simp_all add: trm.inject) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(drule fin_NotL_elim) +apply(auto)[1] +apply(drule freshn_after_substn) +apply(simp) +apply(simp add: fin.intros) +apply(erule fin.cases, simp_all add: trm.inject) +apply(subgoal_tac "\a'::name. a'\(trm{y:=.P},P,name1,x)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndL1) +apply(erule fin.cases, simp_all add: trm.inject) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(drule fin_AndL1_elim) +apply(auto simp add: abs_fresh)[1] +apply(drule freshn_after_substn) +apply(simp) +apply(subgoal_tac "name2\[name1]. trm") +apply(simp add: fin.intros) +apply(simp add: abs_fresh) +apply(subgoal_tac "\a'::name. a'\(trm{y:=.P},P,name1,x)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndL2) +apply(erule fin.cases, simp_all add: trm.inject) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(drule fin_AndL2_elim) +apply(auto simp add: abs_fresh)[1] +apply(drule freshn_after_substn) +apply(simp) +apply(subgoal_tac "name2\[name1].trm") +apply(simp add: fin.intros) +apply(simp add: abs_fresh) +apply(erule fin.cases, simp_all add: trm.inject) +apply(erule fin.cases, simp_all add: trm.inject) +apply(subgoal_tac "\a'::name. a'\(trm1{y:=.P},trm2{y:=.P},name1,name2,P,x)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrL) +apply(erule fin.cases, simp_all add: trm.inject) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(drule fin_OrL_elim) +apply(auto simp add: abs_fresh)[1] +apply(drule freshn_after_substn) +apply(simp) +apply(drule freshn_after_substn) +apply(simp) +apply(subgoal_tac "name3\[name1].trm1 \ name3\[name2].trm2") +apply(simp add: fin.intros) +apply(simp add: abs_fresh) +apply(erule fin.cases, simp_all add: trm.inject) +apply(subgoal_tac "\a'::name. a'\(trm1{name2:=.P},trm2{name2:=.P},name1,P,x)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_ImpL) +apply(erule fin.cases, simp_all add: trm.inject) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(drule fin_ImpL_elim) +apply(auto simp add: abs_fresh)[1] +apply(drule freshn_after_substn) +apply(simp) +apply(drule freshn_after_substn) +apply(simp) +apply(subgoal_tac "x\[name1].trm2") +apply(simp add: fin.intros) +apply(simp add: abs_fresh) +done + +lemma fin_subst1: + assumes a: "fin M x" "x\y" "x\P" + shows "fin (M{y:=.P}) x" +using a +apply(nominal_induct M avoiding: x y c P rule: trm.induct) +apply(auto dest!: fin_elims simp add: subst_fresh abs_fresh) +apply(rule fin.intros, simp add: subst_fresh abs_fresh) +apply(rule fin.intros, simp add: subst_fresh abs_fresh) +apply(rule fin.intros, simp add: subst_fresh abs_fresh) +apply(rule fin.intros, simp add: subst_fresh abs_fresh) +apply(rule fin.intros, simp add: subst_fresh abs_fresh, simp add: subst_fresh abs_fresh) +apply(rule fin.intros, simp add: subst_fresh abs_fresh, simp add: subst_fresh abs_fresh) +apply(rule fin.intros, simp add: subst_fresh abs_fresh, simp add: subst_fresh abs_fresh) +apply(rule fin.intros, simp add: subst_fresh abs_fresh, simp add: subst_fresh abs_fresh) +apply(rule fin.intros, simp add: subst_fresh abs_fresh, simp add: subst_fresh abs_fresh) +done + +lemma fin_subst2: + assumes a: "fin M y" "x\y" "y\P" "M\Ax y c" + shows "fin (M{c:=(x).P}) y" +using a +apply(nominal_induct M avoiding: x y c P rule: trm.induct) +apply(drule fin_elims) +apply(simp add: trm.inject) +apply(rule fin.intros) +apply(drule fin_elims, simp) +apply(drule fin_elims, simp) +apply(drule fin_elims, simp) +apply(rule fin.intros) +apply(auto)[1] +apply(rule subst_fresh) +apply(simp) +apply(drule fin_elims, simp) +apply(drule fin_elims, simp) +apply(rule fin.intros) +apply(simp add: abs_fresh fresh_atm) +apply(rule subst_fresh) +apply(auto)[1] +apply(drule fin_elims, simp) +apply(rule fin.intros) +apply(simp add: abs_fresh fresh_atm) +apply(rule subst_fresh) +apply(auto)[1] +apply(drule fin_elims, simp) +apply(drule fin_elims, simp) +apply(drule fin_elims, simp) +apply(rule fin.intros) +apply(simp add: abs_fresh fresh_atm) +apply(rule subst_fresh) +apply(auto)[1] +apply(simp add: abs_fresh fresh_atm) +apply(rule subst_fresh) +apply(auto)[1] +apply(drule fin_elims, simp) +apply(drule fin_elims, simp) +apply(rule fin.intros) +apply(simp add: abs_fresh fresh_atm) +apply(rule subst_fresh) +apply(auto)[1] +apply(simp add: abs_fresh fresh_atm) +apply(rule subst_fresh) +apply(auto)[1] +done + +lemma fin_substn_nrename: + assumes a: "fin M x" "x\y" "x\P" + shows "M[x\n>y]{y:=.P} = Cut .P (x).(M{y:=.P})" +using a +apply(nominal_induct M avoiding: x y c P rule: trm.induct) +apply(drule fin_Ax_elim) +apply(simp) +apply(simp add: trm.inject) +apply(simp add: alpha calc_atm fresh_atm) +apply(simp) +apply(drule fin_rest_elims) +apply(simp) +apply(drule fin_rest_elims) +apply(simp) +apply(drule fin_NotL_elim) +apply(simp) +apply(subgoal_tac "\z::name. z\(trm,y,x,P,trm[x\n>y]{y:=.P})") +apply(erule exE, simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_NotL) +apply(simp add: trm.inject alpha fresh_atm calc_atm abs_fresh) +apply(rule conjI) +apply(simp add: nsubst_eqvt calc_atm) +apply(simp add: perm_fresh_fresh) +apply(simp add: nrename_fresh) +apply(rule subst_fresh) +apply(simp) +apply(rule exists_fresh') +apply(rule fin_supp) +apply(drule fin_rest_elims) +apply(simp) +apply(drule fin_AndL1_elim) +apply(simp) +apply(subgoal_tac "\z::name. z\(name2,name1,P,trm[name2\n>y]{y:=.P},y,P,trm)") +apply(erule exE, simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndL1) +apply(simp add: trm.inject alpha fresh_atm calc_atm abs_fresh) +apply(rule conjI) +apply(simp add: nsubst_eqvt calc_atm) +apply(simp add: perm_fresh_fresh) +apply(simp add: nrename_fresh) +apply(rule subst_fresh) +apply(simp) +apply(rule exists_fresh') +apply(rule fin_supp) +apply(drule fin_AndL2_elim) +apply(simp) +apply(subgoal_tac "\z::name. z\(name2,name1,P,trm[name2\n>y]{y:=.P},y,P,trm)") +apply(erule exE, simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndL2) +apply(simp add: trm.inject alpha fresh_atm calc_atm abs_fresh) +apply(rule conjI) +apply(simp add: nsubst_eqvt calc_atm) +apply(simp add: perm_fresh_fresh) +apply(simp add: nrename_fresh) +apply(rule subst_fresh) +apply(simp) +apply(rule exists_fresh') +apply(rule fin_supp) +apply(drule fin_rest_elims) +apply(simp) +apply(drule fin_rest_elims) +apply(simp) +apply(drule fin_OrL_elim) +apply(simp add: abs_fresh) +apply(simp add: subst_fresh rename_fresh) +apply(subgoal_tac "\z::name. z\(name3,name2,name1,P,trm1[name3\n>y]{y:=.P},trm2[name3\n>y]{y:=.P},y,P,trm1,trm2)") +apply(erule exE, simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrL) +apply(simp add: trm.inject alpha fresh_atm calc_atm abs_fresh) +apply(rule conjI) +apply(simp add: nsubst_eqvt calc_atm) +apply(simp add: perm_fresh_fresh) +apply(simp add: nrename_fresh) +apply(simp add: nsubst_eqvt calc_atm) +apply(simp add: perm_fresh_fresh) +apply(simp add: nrename_fresh) +apply(rule exists_fresh') +apply(rule fin_supp) +apply(drule fin_rest_elims) +apply(simp) +apply(drule fin_ImpL_elim) +apply(simp add: abs_fresh) +apply(simp add: subst_fresh rename_fresh) +apply(subgoal_tac "\z::name. z\(name1,x,P,trm1[x\n>y]{y:=.P},trm2[x\n>y]{y:=.P},y,P,trm1,trm2)") +apply(erule exE, simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_ImpL) +apply(simp add: trm.inject alpha fresh_atm calc_atm abs_fresh) +apply(rule conjI) +apply(simp add: nsubst_eqvt calc_atm) +apply(simp add: perm_fresh_fresh) +apply(simp add: nrename_fresh) +apply(simp add: nsubst_eqvt calc_atm) +apply(simp add: perm_fresh_fresh) +apply(simp add: nrename_fresh) +apply(rule exists_fresh') +apply(rule fin_supp) +done + +inductive2 + fic :: "trm \ coname \ bool" +where + [intro]: "fic (Ax x a) a" +| [intro]: "a\M \ fic (NotR (x).M a) a" +| [intro]: "\c\[a].M;c\[b].N\ \ fic (AndR .M .N c) c" +| [intro]: "b\[a].M \ fic (OrR1 .M b) b" +| [intro]: "b\[a].M \ fic (OrR2 .M b) b" +| [intro]: "\b\[a].M\ \ fic (ImpR (x)..M b) b" + +equivariance fic + +lemma fic_Ax_elim: + assumes a: "fic (Ax x a) b" + shows "a=b" +using a +apply(erule_tac fic.cases) +apply(auto simp add: trm.inject) +done + +lemma fic_NotR_elim: + assumes a: "fic (NotR (x).M a) b" + shows "a=b \ b\M" +using a +apply(erule_tac fic.cases) +apply(auto simp add: trm.inject) +apply(subgoal_tac "b\[xa].Ma") +apply(drule sym) +apply(simp_all add: abs_fresh) +done + +lemma fic_OrR1_elim: + assumes a: "fic (OrR1 .M b) c" + shows "b=c \ c\[a].M" +using a +apply(erule_tac fic.cases) +apply(auto simp add: trm.inject) +done + +lemma fic_OrR2_elim: + assumes a: "fic (OrR2 .M b) c" + shows "b=c \ c\[a].M" +using a +apply(erule_tac fic.cases) +apply(auto simp add: trm.inject) +done + +lemma fic_AndR_elim: + assumes a: "fic (AndR .M .N c) d" + shows "c=d \ d\[a].M \ d\[b].N" +using a +apply(erule_tac fic.cases) +apply(auto simp add: trm.inject) +done + +lemma fic_ImpR_elim: + assumes a: "fic (ImpR (x)..M b) c" + shows "b=c \ b\[a].M" +using a +apply(erule_tac fic.cases) +apply(auto simp add: trm.inject) +apply(subgoal_tac "c\[xa].[aa].Ma") +apply(drule sym) +apply(simp_all add: abs_fresh) +done + +lemma fic_rest_elims: + shows "fic (Cut .M (x).N) d \ False" + and "fic (NotL .M x) d \ False" + and "fic (OrL (x).M (y).N z) d \ False" + and "fic (AndL1 (x).M y) d \ False" + and "fic (AndL2 (x).M y) d \ False" + and "fic (ImpL .M (x).N y) d \ False" +by (erule fic.cases, simp_all add: trm.inject)+ + +lemmas fic_elims = fic_Ax_elim fic_NotR_elim fic_OrR1_elim fic_OrR2_elim fic_AndR_elim + fic_ImpR_elim fic_rest_elims + +lemma fic_rename: + shows "fic M a \ fic ([(a',a)]\M) a'" +by (induct rule: fic.induct) + (auto simp add: calc_atm simp add: fresh_left abs_fresh) + +lemma not_fic_subst1: + assumes a: "\(fic M a)" + shows "\(fic (M{y:=.P}) a)" +using a +apply(nominal_induct M avoiding: a c y P rule: trm.induct) +apply(auto) +apply(drule fic_elims, simp) +apply(drule fic_elims, simp) +apply(drule fic_elims, simp) +apply(drule fic_elims) +apply(auto)[1] +apply(drule freshc_after_substn) +apply(simp add: fic.intros) +apply(subgoal_tac "\x'::name. x'\(trm{y:=.P},P,a)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_NotL) +apply(drule fic_elims, simp) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(drule fic_elims, simp) +apply(drule fic_elims) +apply(auto)[1] +apply(simp add: abs_fresh fresh_atm) +apply(drule freshc_after_substn) +apply(drule freshc_after_substn) +apply(simp add: fic.intros abs_fresh) +apply(subgoal_tac "\x'::name. x'\(trm{y:=.P},P,name1,a)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndL1) +apply(drule fic_elims, simp) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(drule fic_elims, simp) +apply(subgoal_tac "\x'::name. x'\(trm{y:=.P},P,name1,a)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndL2) +apply(drule fic_elims, simp) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(drule fic_elims, simp) +apply(drule fic_elims) +apply(auto)[1] +apply(simp add: abs_fresh fresh_atm) +apply(drule freshc_after_substn) +apply(simp add: fic.intros abs_fresh) +apply(drule fic_elims) +apply(auto)[1] +apply(simp add: abs_fresh fresh_atm) +apply(drule freshc_after_substn) +apply(simp add: fic.intros abs_fresh) +apply(subgoal_tac "\x'::name. x'\(trm1{y:=.P},trm2{y:=.P},P,name1,name2,a)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrL) +apply(drule fic_elims, simp) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(drule fic_elims, simp) +apply(drule fic_elims, simp) +apply(auto)[1] +apply(simp add: abs_fresh fresh_atm) +apply(drule freshc_after_substn) +apply(simp add: fic.intros abs_fresh) +apply(subgoal_tac "\x'::name. x'\(trm1{name2:=.P},trm2{name2:=.P},P,name1,name2,a)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_ImpL) +apply(drule fic_elims, simp) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(drule fic_elims, simp) +done + +lemma not_fic_subst2: + assumes a: "\(fic M a)" + shows "\(fic (M{c:=(y).P}) a)" +using a +apply(nominal_induct M avoiding: a c y P rule: trm.induct) +apply(auto) +apply(drule fic_elims, simp) +apply(drule fic_elims, simp) +apply(drule fic_elims, simp) +apply(subgoal_tac "\c'::coname. c'\(trm{coname:=(y).P},P,a)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_NotR) +apply(drule fic_elims, simp) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(drule fic_elims, simp) +apply(erule conjE)+ +apply(drule freshc_after_substc) +apply(simp) +apply(simp add: fic.intros abs_fresh) +apply(drule fic_elims, simp) +apply(subgoal_tac "\c'::coname. c'\(trm1{coname3:=(y).P},trm2{coname3:=(y).P},P,coname1,coname2,a)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndR) +apply(drule fic_elims, simp) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(drule fic_elims, simp) +apply(auto)[1] +apply(simp add: abs_fresh fresh_atm) +apply(drule freshc_after_substc) +apply(simp) +apply(drule freshc_after_substc) +apply(simp) +apply(simp add: fic.intros abs_fresh) +apply(drule fic_elims, simp) +apply(drule fic_elims, simp) +apply(subgoal_tac "\c'::coname. c'\(trm{coname2:=(y).P},P,coname1,a)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrR1) +apply(drule fic_elims, simp) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(drule fic_elims, simp) +apply(auto)[1] +apply(simp add: abs_fresh fresh_atm) +apply(drule freshc_after_substc) +apply(simp) +apply(simp add: fic.intros abs_fresh) +apply(subgoal_tac "\c'::coname. c'\(trm{coname2:=(y).P},P,coname1,a)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrR2) +apply(drule fic_elims, simp) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(drule fic_elims, simp) +apply(auto)[1] +apply(simp add: abs_fresh fresh_atm) +apply(drule freshc_after_substc) +apply(simp) +apply(simp add: fic.intros abs_fresh) +apply(drule fic_elims, simp) +apply(subgoal_tac "\c'::coname. c'\(trm{coname2:=(y).P},P,coname1,a)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_ImpR) +apply(drule fic_elims, simp) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(drule fic_elims, simp) +apply(auto)[1] +apply(simp add: abs_fresh fresh_atm) +apply(drule freshc_after_substc) +apply(simp) +apply(simp add: fic.intros abs_fresh) +apply(drule fic_elims, simp) +done + +lemma fic_subst1: + assumes a: "fic M a" "a\b" "a\P" + shows "fic (M{b:=(x).P}) a" +using a +apply(nominal_induct M avoiding: x b a P rule: trm.induct) +apply(drule fic_elims) +apply(simp add: fic.intros) +apply(drule fic_elims, simp) +apply(drule fic_elims, simp) +apply(rule fic.intros) +apply(auto)[1] +apply(rule subst_fresh) +apply(simp) +apply(drule fic_elims, simp) +apply(drule fic_elims, simp) +apply(rule fic.intros) +apply(simp add: abs_fresh fresh_atm) +apply(rule subst_fresh) +apply(auto)[1] +apply(simp add: abs_fresh fresh_atm) +apply(rule subst_fresh) +apply(auto)[1] +apply(drule fic_elims, simp) +apply(drule fic_elims, simp) +apply(drule fic_elims, simp) +apply(rule fic.intros) +apply(simp add: abs_fresh fresh_atm) +apply(rule subst_fresh) +apply(auto)[1] +apply(drule fic_elims, simp) +apply(rule fic.intros) +apply(simp add: abs_fresh fresh_atm) +apply(rule subst_fresh) +apply(auto)[1] +apply(drule fic_elims, simp) +apply(drule fic_elims, simp) +apply(rule fic.intros) +apply(simp add: abs_fresh fresh_atm) +apply(rule subst_fresh) +apply(auto)[1] +apply(drule fic_elims, simp) +done + +lemma fic_subst2: + assumes a: "fic M a" "c\a" "a\P" "M\Ax x a" + shows "fic (M{x:=.P}) a" +using a +apply(nominal_induct M avoiding: x a c P rule: trm.induct) +apply(drule fic_elims) +apply(simp add: trm.inject) +apply(rule fic.intros) +apply(drule fic_elims, simp) +apply(drule fic_elims, simp) +apply(rule fic.intros) +apply(auto)[1] +apply(rule subst_fresh) +apply(simp) +apply(drule fic_elims, simp) +apply(drule fic_elims, simp) +apply(rule fic.intros) +apply(simp add: abs_fresh fresh_atm) +apply(rule subst_fresh) +apply(auto)[1] +apply(simp add: abs_fresh fresh_atm) +apply(rule subst_fresh) +apply(auto)[1] +apply(drule fic_elims, simp) +apply(drule fic_elims, simp) +apply(drule fic_elims, simp) +apply(rule fic.intros) +apply(simp add: abs_fresh fresh_atm) +apply(rule subst_fresh) +apply(auto)[1] +apply(drule fic_elims, simp) +apply(rule fic.intros) +apply(simp add: abs_fresh fresh_atm) +apply(rule subst_fresh) +apply(auto)[1] +apply(drule fic_elims, simp) +apply(drule fic_elims, simp) +apply(rule fic.intros) +apply(simp add: abs_fresh fresh_atm) +apply(rule subst_fresh) +apply(auto)[1] +apply(drule fic_elims, simp) +done + +lemma fic_substc_crename: + assumes a: "fic M a" "a\b" "a\P" + shows "M[a\c>b]{b:=(y).P} = Cut .(M{b:=(y).P}) (y).P" +using a +apply(nominal_induct M avoiding: a b y P rule: trm.induct) +apply(drule fic_Ax_elim) +apply(simp) +apply(simp add: trm.inject) +apply(simp add: alpha calc_atm fresh_atm trm.inject) +apply(simp) +apply(drule fic_rest_elims) +apply(simp) +apply(drule fic_NotR_elim) +apply(simp) +apply(generate_fresh "coname") +apply(fresh_fun_simp) +apply(simp add: trm.inject alpha fresh_atm fresh_prod fresh_atm calc_atm abs_fresh) +apply(rule conjI) +apply(simp add: csubst_eqvt calc_atm) +apply(simp add: perm_fresh_fresh) +apply(simp add: crename_fresh) +apply(rule subst_fresh) +apply(simp) +apply(drule fic_rest_elims) +apply(simp) +apply(drule fic_AndR_elim) +apply(simp add: abs_fresh fresh_atm subst_fresh rename_fresh) +apply(generate_fresh "coname") +apply(fresh_fun_simp) +apply(simp add: trm.inject alpha fresh_atm calc_atm abs_fresh fresh_prod) +apply(rule conjI) +apply(simp add: csubst_eqvt calc_atm) +apply(simp add: perm_fresh_fresh) +apply(simp add: csubst_eqvt calc_atm) +apply(simp add: perm_fresh_fresh) +apply(simp add: subst_fresh) +apply(drule fic_rest_elims) +apply(simp) +apply(drule fic_rest_elims) +apply(simp) +apply(drule fic_OrR1_elim) +apply(simp) +apply(generate_fresh "coname") +apply(fresh_fun_simp) +apply(simp add: trm.inject alpha fresh_atm calc_atm abs_fresh fresh_prod) +apply(simp add: csubst_eqvt calc_atm) +apply(simp add: perm_fresh_fresh) +apply(simp add: subst_fresh rename_fresh) +apply(drule fic_OrR2_elim) +apply(simp add: abs_fresh fresh_atm) +apply(generate_fresh "coname") +apply(fresh_fun_simp) +apply(simp add: trm.inject alpha fresh_atm calc_atm abs_fresh fresh_prod) +apply(simp add: csubst_eqvt calc_atm) +apply(simp add: perm_fresh_fresh) +apply(simp add: subst_fresh rename_fresh) +apply(drule fic_rest_elims) +apply(simp) +apply(drule fic_ImpR_elim) +apply(simp add: abs_fresh fresh_atm) +apply(auto)[1] +apply(generate_fresh "coname") +apply(fresh_fun_simp) +apply(simp add: trm.inject alpha fresh_atm calc_atm abs_fresh fresh_prod) +apply(simp add: csubst_eqvt calc_atm) +apply(simp add: perm_fresh_fresh) +apply(simp add: subst_fresh rename_fresh) +apply(drule fic_rest_elims) +apply(simp) +done + +inductive2 + l_redu :: "trm \ trm \ bool" ("_ \\<^isub>l _" [100,100] 100) +where + LAxR: "\x\M; a\b; fic M a\ \ Cut .M (x).(Ax x b) \\<^isub>l M[a\c>b]" +| LAxL: "\a\M; x\y; fin M x\ \ Cut .(Ax y a) (x).M \\<^isub>l M[x\n>y]" +| LNot: "\y\(M,N); x\(N,y); a\(M,N,b); b\M; y\x; b\a\ \ + Cut .(NotR (x).M a) (y).(NotL .N y) \\<^isub>l Cut .N (x).M" +| LAnd1: "\b\([a1].M1,[a2].M2,N,a1,a2); y\([x].N,M1,M2,x); x\(M1,M2); a1\(M2,N); a2\(M1,N); a1\a2\ \ + Cut .(AndR .M1 .M2 b) (y).(AndL1 (x).N y) \\<^isub>l Cut .M1 (x).N" +| LAnd2: "\b\([a1].M1,[a2].M2,N,a1,a2); y\([x].N,M1,M2,x); x\(M1,M2); a1\(M2,N); a2\(M1,N); a1\a2\ \ + Cut .(AndR .M1 .M2 b) (y).(AndL2 (x).N y) \\<^isub>l Cut .M2 (x).N" +| LOr1: "\b\([a].M,N1,N2,a); y\([x1].N1,[x2].N2,M,x1,x2); x1\(M,N2); x2\(M,N1); a\(N1,N2); x1\x2\ \ + Cut .(OrR1 .M b) (y).(OrL (x1).N1 (x2).N2 y) \\<^isub>l Cut .M (x1).N1" +| LOr2: "\b\([a].M,N1,N2,a); y\([x1].N1,[x2].N2,M,x1,x2); x1\(M,N2); x2\(M,N1); a\(N1,N2); x1\x2\ \ + Cut .(OrR2 .M b) (y).(OrL (x1).N1 (x2).N2 y) \\<^isub>l Cut .M (x2).N2" +| LImp: "\z\(N,[y].P,[x].M,y,x); b\([a].M,[c].N,P,c,a); x\(N,[y].P,y); + c\(P,[a].M,b,a); a\([c].N,P); y\(N,[x].M)\ \ + Cut .(ImpR (x)..M b) (z).(ImpL .N (y).P z) \\<^isub>l Cut .(Cut .N (x).M) (y).P" + +equivariance l_redu + +lemma l_redu_eqvt': + fixes pi1::"name prm" + and pi2::"coname prm" + shows "(pi1\M) \\<^isub>l (pi1\M') \ M \\<^isub>l M'" + and "(pi2\M) \\<^isub>l (pi2\M') \ M \\<^isub>l M'" +apply - +apply(drule_tac pi="rev pi1" in l_redu.eqvt(1)) +apply(perm_simp) +apply(drule_tac pi="rev pi2" in l_redu.eqvt(2)) +apply(perm_simp) +done + +nominal_inductive l_redu + apply(simp_all add: abs_fresh fresh_atm rename_fresh fresh_prod abs_supp fin_supp) + apply(force)+ + done + +lemma fresh_l_redu: + fixes x::"name" + and a::"coname" + shows "M \\<^isub>l M' \ x\M \ x\M'" + and "M \\<^isub>l M' \ a\M \ a\M'" +apply - +apply(induct rule: l_redu.induct) +apply(auto simp add: abs_fresh rename_fresh) +apply(case_tac "xa=x") +apply(simp add: rename_fresh) +apply(simp add: rename_fresh fresh_atm) +apply(simp add: fresh_prod abs_fresh abs_supp fin_supp)+ +apply(induct rule: l_redu.induct) +apply(auto simp add: abs_fresh rename_fresh) +apply(case_tac "aa=a") +apply(simp add: rename_fresh) +apply(simp add: rename_fresh fresh_atm) +apply(simp add: fresh_prod abs_fresh abs_supp fin_supp)+ +done + +lemma better_LAxR_intro[intro]: + shows "fic M a \ Cut .M (x).(Ax x b) \\<^isub>l M[a\c>b]" +proof - + assume fin: "fic M a" + obtain x'::"name" where fs1: "x'\(M,x)" by (rule exists_fresh(1), rule fin_supp, blast) + obtain a'::"coname" where fs2: "a'\(a,M,b)" by (rule exists_fresh(2), rule fin_supp, blast) + have "Cut .M (x).(Ax x b) = Cut .([(a',a)]\M) (x').(Ax x' b)" + using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm) + also have "\ \\<^isub>l ([(a',a)]\M)[a'\c>b]" using fs1 fs2 fin + by (auto intro: l_redu.intros simp add: fresh_left calc_atm fic_rename) + also have "\ = M[a\c>b]" using fs1 fs2 by (simp add: crename_rename) + finally show ?thesis by simp +qed + +lemma better_LAxL_intro[intro]: + shows "fin M x \ Cut .(Ax y a) (x).M \\<^isub>l M[x\n>y]" +proof - + assume fin: "fin M x" + obtain x'::"name" where fs1: "x'\(y,M,x)" by (rule exists_fresh(1), rule fin_supp, blast) + obtain a'::"coname" where fs2: "a'\(a,M)" by (rule exists_fresh(2), rule fin_supp, blast) + have "Cut .(Ax y a) (x).M = Cut .(Ax y a') (x').([(x',x)]\M)" + using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm) + also have "\ \\<^isub>l ([(x',x)]\M)[x'\n>y]" using fs1 fs2 fin + by (auto intro: l_redu.intros simp add: fresh_left calc_atm fin_rename) + also have "\ = M[x\n>y]" using fs1 fs2 by (simp add: nrename_rename) + finally show ?thesis by simp +qed + +lemma better_LNot_intro[intro]: + shows "\y\N; a\M\ \ Cut .(NotR (x).M a) (y).(NotL .N y) \\<^isub>l Cut .N (x).M" +proof - + assume fs: "y\N" "a\M" + obtain x'::"name" where f1: "x'\(y,N,M,x)" by (rule exists_fresh(1), rule fin_supp, blast) + obtain y'::"name" where f2: "y'\(y,N,M,x,x')" by (rule exists_fresh(1), rule fin_supp, blast) + obtain a'::"coname" where f3: "a'\(a,M,N,b)" by (rule exists_fresh(2), rule fin_supp, blast) + obtain b'::"coname" where f4: "b'\(a,M,N,b,a')" by (rule exists_fresh(2), rule fin_supp, blast) + have "Cut .(NotR (x).M a) (y).(NotL .N y) + = Cut .(NotR (x).([(a',a)]\M) a') (y').(NotL .([(y',y)]\N) y')" + using f1 f2 f3 f4 + by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm abs_fresh) + also have "\ = Cut .(NotR (x).M a') (y').(NotL .N y')" + using f1 f2 f3 f4 fs by (perm_simp) + also have "\ = Cut .(NotR (x').([(x',x)]\M) a') (y').(NotL .([(b',b)]\N) y')" + using f1 f2 f3 f4 + by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm) + also have "\ \\<^isub>l Cut .([(b',b)]\N) (x').([(x',x)]\M)" + using f1 f2 f3 f4 fs + by (auto intro: l_redu.intros simp add: fresh_prod fresh_left calc_atm fresh_atm) + also have "\ = Cut .N (x).M" + using f1 f2 f3 f4 by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm) + finally show ?thesis by simp +qed + +lemma better_LAnd1_intro[intro]: + shows "\a\([b1].M1,[b2].M2); y\[x].N\ + \ Cut .(AndR .M1 .M2 a) (y).(AndL1 (x).N y) \\<^isub>l Cut .M1 (x).N" +proof - + assume fs: "a\([b1].M1,[b2].M2)" "y\[x].N" + obtain x'::"name" where f1: "x'\(y,N,M1,M2,x)" by (rule exists_fresh(1), rule fin_supp, blast) + obtain y'::"name" where f2: "y'\(y,N,M1,M2,x,x')" by (rule exists_fresh(1), rule fin_supp, blast) + obtain a'::"coname" where f3: "a'\(a,M1,M2,N,b1,b2)" by (rule exists_fresh(2), rule fin_supp, blast) + obtain b1'::"coname" where f4:"b1'\(a,M1,M2,N,b1,b2,a')" by (rule exists_fresh(2), rule fin_supp, blast) + obtain b2'::"coname" where f5:"b2'\(a,M1,M2,N,b1,b2,a',b1')" by (rule exists_fresh(2),rule fin_supp, blast) + have "Cut .(AndR .M1 .M2 a) (y).(AndL1 (x).N y) + = Cut .(AndR .M1 .M2 a') (y').(AndL1 (x).N y')" + using f1 f2 f3 f4 fs + apply(rule_tac sym) + apply(perm_simp add: trm.inject alpha calc_atm fresh_prod fresh_left fresh_atm abs_fresh) + apply(auto simp add: perm_fresh_fresh) + done + also have "\ = Cut .(AndR .([(b1',b1)]\M1) .([(b2',b2)]\M2) a') + (y').(AndL1 (x').([(x',x)]\N) y')" + using f1 f2 f3 f4 f5 fs + apply(rule_tac sym) + apply(perm_simp add: trm.inject alpha calc_atm fresh_prod fresh_left fresh_atm abs_fresh) + done + also have "\ \\<^isub>l Cut .([(b1',b1)]\M1) (x').([(x',x)]\N)" + using f1 f2 f3 f4 f5 fs + apply - + apply(rule l_redu.intros) + apply(auto simp add: abs_fresh fresh_prod fresh_left calc_atm fresh_atm) + done + also have "\ = Cut .M1 (x).N" + using f1 f2 f3 f4 f5 fs by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm) + finally show ?thesis by simp +qed + +lemma better_LAnd2_intro[intro]: + shows "\a\([b1].M1,[b2].M2); y\[x].N\ + \ Cut .(AndR .M1 .M2 a) (y).(AndL2 (x).N y) \\<^isub>l Cut .M2 (x).N" +proof - + assume fs: "a\([b1].M1,[b2].M2)" "y\[x].N" + obtain x'::"name" where f1: "x'\(y,N,M1,M2,x)" by (rule exists_fresh(1), rule fin_supp, blast) + obtain y'::"name" where f2: "y'\(y,N,M1,M2,x,x')" by (rule exists_fresh(1), rule fin_supp, blast) + obtain a'::"coname" where f3: "a'\(a,M1,M2,N,b1,b2)" by (rule exists_fresh(2), rule fin_supp, blast) + obtain b1'::"coname" where f4:"b1'\(a,M1,M2,N,b1,b2,a')" by (rule exists_fresh(2), rule fin_supp, blast) + obtain b2'::"coname" where f5:"b2'\(a,M1,M2,N,b1,b2,a',b1')" by (rule exists_fresh(2),rule fin_supp, blast) + have "Cut .(AndR .M1 .M2 a) (y).(AndL2 (x).N y) + = Cut .(AndR .M1 .M2 a') (y').(AndL2 (x).N y')" + using f1 f2 f3 f4 fs + apply(rule_tac sym) + apply(perm_simp add: trm.inject alpha calc_atm fresh_prod fresh_left fresh_atm abs_fresh) + apply(auto simp add: perm_fresh_fresh) + done + also have "\ = Cut .(AndR .([(b1',b1)]\M1) .([(b2',b2)]\M2) a') + (y').(AndL2 (x').([(x',x)]\N) y')" + using f1 f2 f3 f4 f5 fs + apply(rule_tac sym) + apply(perm_simp add: trm.inject alpha calc_atm fresh_prod fresh_left fresh_atm abs_fresh) + done + also have "\ \\<^isub>l Cut .([(b2',b2)]\M2) (x').([(x',x)]\N)" + using f1 f2 f3 f4 f5 fs + apply - + apply(rule l_redu.intros) + apply(auto simp add: abs_fresh fresh_prod fresh_left calc_atm fresh_atm) + done + also have "\ = Cut .M2 (x).N" + using f1 f2 f3 f4 f5 fs by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm) + finally show ?thesis by simp +qed + +lemma better_LOr1_intro[intro]: + shows "\y\([x1].N1,[x2].N2); b\[a].M\ + \ Cut .(OrR1 .M b) (y).(OrL (x1).N1 (x2).N2 y) \\<^isub>l Cut .M (x1).N1" +proof - + assume fs: "y\([x1].N1,[x2].N2)" "b\[a].M" + obtain y'::"name" where f1: "y'\(y,M,N1,N2,x1,x2)" by (rule exists_fresh(1), rule fin_supp, blast) + obtain x1'::"name" where f2: "x1'\(y,M,N1,N2,x1,x2,y')" by (rule exists_fresh(1), rule fin_supp, blast) + obtain x2'::"name" where f3: "x2'\(y,M,N1,N2,x1,x2,y',x1')" by (rule exists_fresh(1), rule fin_supp, blast) + obtain a'::"coname" where f4: "a'\(a,N1,N2,M,b)" by (rule exists_fresh(2), rule fin_supp, blast) + obtain b'::"coname" where f5: "b'\(a,N1,N2,M,b,a')" by (rule exists_fresh(2),rule fin_supp, blast) + have "Cut .(OrR1 .M b) (y).(OrL (x1).N1 (x2).N2 y) + = Cut .(OrR1 .M b') (y').(OrL (x1).N1 (x2).N2 y')" + using f1 f2 f3 f4 f5 fs + apply(rule_tac sym) + apply(perm_simp add: trm.inject alpha calc_atm fresh_prod fresh_left fresh_atm abs_fresh) + apply(auto simp add: perm_fresh_fresh) + done + also have "\ = Cut .(OrR1 .([(a',a)]\M) b') + (y').(OrL (x1').([(x1',x1)]\N1) (x2').([(x2',x2)]\N2) y')" + using f1 f2 f3 f4 f5 fs + apply(rule_tac sym) + apply(perm_simp add: trm.inject alpha calc_atm fresh_prod fresh_left fresh_atm abs_fresh) + done + also have "\ \\<^isub>l Cut .([(a',a)]\M) (x1').([(x1',x1)]\N1)" + using f1 f2 f3 f4 f5 fs + apply - + apply(rule l_redu.intros) + apply(auto simp add: abs_fresh fresh_prod fresh_left calc_atm fresh_atm) + done + also have "\ = Cut .M (x1).N1" + using f1 f2 f3 f4 f5 fs by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm) + finally show ?thesis by simp +qed + +lemma better_LOr2_intro[intro]: + shows "\y\([x1].N1,[x2].N2); b\[a].M\ + \ Cut .(OrR2 .M b) (y).(OrL (x1).N1 (x2).N2 y) \\<^isub>l Cut .M (x2).N2" +proof - + assume fs: "y\([x1].N1,[x2].N2)" "b\[a].M" + obtain y'::"name" where f1: "y'\(y,M,N1,N2,x1,x2)" by (rule exists_fresh(1), rule fin_supp, blast) + obtain x1'::"name" where f2: "x1'\(y,M,N1,N2,x1,x2,y')" by (rule exists_fresh(1), rule fin_supp, blast) + obtain x2'::"name" where f3: "x2'\(y,M,N1,N2,x1,x2,y',x1')" by (rule exists_fresh(1), rule fin_supp, blast) + obtain a'::"coname" where f4: "a'\(a,N1,N2,M,b)" by (rule exists_fresh(2), rule fin_supp, blast) + obtain b'::"coname" where f5: "b'\(a,N1,N2,M,b,a')" by (rule exists_fresh(2),rule fin_supp, blast) + have "Cut .(OrR2 .M b) (y).(OrL (x1).N1 (x2).N2 y) + = Cut .(OrR2 .M b') (y').(OrL (x1).N1 (x2).N2 y')" + using f1 f2 f3 f4 f5 fs + apply(rule_tac sym) + apply(perm_simp add: trm.inject alpha calc_atm fresh_prod fresh_left fresh_atm abs_fresh) + apply(auto simp add: perm_fresh_fresh) + done + also have "\ = Cut .(OrR2 .([(a',a)]\M) b') + (y').(OrL (x1').([(x1',x1)]\N1) (x2').([(x2',x2)]\N2) y')" + using f1 f2 f3 f4 f5 fs + apply(rule_tac sym) + apply(perm_simp add: trm.inject alpha calc_atm fresh_prod fresh_left fresh_atm abs_fresh) + done + also have "\ \\<^isub>l Cut .([(a',a)]\M) (x2').([(x2',x2)]\N2)" + using f1 f2 f3 f4 f5 fs + apply - + apply(rule l_redu.intros) + apply(auto simp add: abs_fresh fresh_prod fresh_left calc_atm fresh_atm) + done + also have "\ = Cut .M (x2).N2" + using f1 f2 f3 f4 f5 fs by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm) + finally show ?thesis by simp +qed + +lemma better_LImp_intro[intro]: + shows "\z\(N,[y].P); b\[a].M; a\N\ + \ Cut .(ImpR (x)..M b) (z).(ImpL .N (y).P z) \\<^isub>l Cut .(Cut .N (x).M) (y).P" +proof - + assume fs: "z\(N,[y].P)" "b\[a].M" "a\N" + obtain y'::"name" where f1: "y'\(y,M,N,P,z,x)" by (rule exists_fresh(1), rule fin_supp, blast) + obtain x'::"name" where f2: "x'\(y,M,N,P,z,x,y')" by (rule exists_fresh(1), rule fin_supp, blast) + obtain z'::"name" where f3: "z'\(y,M,N,P,z,x,y',x')" by (rule exists_fresh(1), rule fin_supp, blast) + obtain a'::"coname" where f4: "a'\(a,N,P,M,b)" by (rule exists_fresh(2), rule fin_supp, blast) + obtain b'::"coname" where f5: "b'\(a,N,P,M,b,c,a')" by (rule exists_fresh(2),rule fin_supp, blast) + obtain c'::"coname" where f6: "c'\(a,N,P,M,b,c,a',b')" by (rule exists_fresh(2),rule fin_supp, blast) + have " Cut .(ImpR (x)..M b) (z).(ImpL .N (y).P z) + = Cut .(ImpR (x)..M b') (z').(ImpL .N (y).P z')" + using f1 f2 f3 f4 f5 fs + apply(rule_tac sym) + apply(perm_simp add: trm.inject alpha calc_atm fresh_prod fresh_left fresh_atm abs_fresh) + apply(auto simp add: perm_fresh_fresh) + done + also have "\ = Cut .(ImpR (x')..([(a',a)]\([(x',x)]\M)) b') + (z').(ImpL .([(c',c)]\N) (y').([(y',y)]\P) z')" + using f1 f2 f3 f4 f5 f6 fs + apply(rule_tac sym) + apply(simp add: trm.inject) + apply(simp add: alpha) + apply(rule conjI) + apply(simp add: trm.inject) + apply(simp add: alpha fresh_prod fresh_atm abs_perm calc_atm fresh_left abs_fresh) + apply(simp add: trm.inject) + apply(simp add: alpha) + apply(rule conjI) + apply(simp add: alpha fresh_prod fresh_atm abs_perm calc_atm fresh_left abs_fresh) + apply(simp add: alpha fresh_prod fresh_atm abs_perm calc_atm fresh_left abs_fresh) + done + also have "\ \\<^isub>l Cut .(Cut .([(c',c)]\N) (x').([(a',a)]\[(x',x)]\M)) (y').([(y',y)]\P)" + using f1 f2 f3 f4 f5 f6 fs + apply - + apply(rule l_redu.intros) + apply(auto simp add: abs_fresh fresh_prod fresh_left calc_atm fresh_atm) + done + also have "\ = Cut .(Cut .N (x).M) (y).P" + using f1 f2 f3 f4 f5 f6 fs + apply(simp add: trm.inject) + apply(rule conjI) + apply(simp add: alpha) + apply(rule disjI2) + apply(simp add: trm.inject) + apply(rule conjI) + apply(simp add: fresh_prod fresh_atm) + apply(rule conjI) + apply(perm_simp add: calc_atm) + apply(auto simp add: fresh_prod fresh_atm)[1] + apply(perm_simp add: alpha) + apply(perm_simp add: alpha) + apply(perm_simp add: alpha) + apply(rule conjI) + apply(perm_simp add: calc_atm) + apply(rule_tac pi="[(a',a)]" in pt_bij4[OF pt_coname_inst, OF at_coname_inst]) + apply(perm_simp add: abs_perm calc_atm) + apply(perm_simp add: alpha fresh_prod fresh_atm) + apply(simp add: abs_fresh) + apply(perm_simp add: alpha fresh_prod fresh_atm) + done + finally show ?thesis by simp +qed + +lemma alpha_coname: + fixes M::"trm" + and a::"coname" + assumes a: "[a].M = [b].N" "c\(a,b,M,N)" + shows "M = [(a,c)]\[(b,c)]\N" +using a +apply(auto simp add: alpha' fresh_prod fresh_atm) +apply(drule sym) +apply(perm_simp) +done + +lemma alpha_name: + fixes M::"trm" + and x::"name" + assumes a: "[x].M = [y].N" "z\(x,y,M,N)" + shows "M = [(x,z)]\[(y,z)]\N" +using a +apply(auto simp add: alpha' fresh_prod fresh_atm) +apply(drule sym) +apply(perm_simp) +done + +lemma alpha_name_coname: + fixes M::"trm" + and x::"name" + and a::"coname" + assumes a: "[x].[b].M = [y].[c].N" "z\(x,y,M,N)" "a\(b,c,M,N)" + shows "M = [(x,z)]\[(b,a)]\[(c,a)]\[(y,z)]\N" +using a +apply(auto simp add: alpha' fresh_prod fresh_atm abs_supp fin_supp abs_fresh abs_perm fresh_left calc_atm) +apply(drule sym) +apply(simp) +apply(perm_simp) +done + +lemma Cut_l_redu_elim: + assumes a: "Cut .M (x).N \\<^isub>l R" + shows "(\b. R = M[a\c>b]) \ (\y. R = N[x\n>y]) \ + (\y M' b N'. M = NotR (y).M' a \ N = NotL .N' x \ R = Cut .N' (y).M' \ fic M a \ fin N x) \ + (\b M1 c M2 y N'. M = AndR .M1 .M2 a \ N = AndL1 (y).N' x \ R = Cut .M1 (y).N' + \ fic M a \ fin N x) \ + (\b M1 c M2 y N'. M = AndR .M1 .M2 a \ N = AndL2 (y).N' x \ R = Cut .M2 (y).N' + \ fic M a \ fin N x) \ + (\b N' z M1 y M2. M = OrR1 .N' a \ N = OrL (z).M1 (y).M2 x \ R = Cut .N' (z).M1 + \ fic M a \ fin N x) \ + (\b N' z M1 y M2. M = OrR2 .N' a \ N = OrL (z).M1 (y).M2 x \ R = Cut .N' (y).M2 + \ fic M a \ fin N x) \ + (\z b M' c N1 y N2. M = ImpR (z)..M' a \ N = ImpL .N1 (y).N2 x \ + R = Cut .(Cut .N1 (z).M') (y).N2 \ b\(c,N1) \ fic M a \ fin N x)" +using a +apply(erule_tac l_redu.cases) +apply(rule disjI1) +(* ax case *) +apply(simp add: trm.inject) +apply(rule_tac x="b" in exI) +apply(erule conjE) +apply(simp add: alpha) +apply(erule disjE) +apply(simp) +apply(simp) +apply(simp add: rename_fresh) +apply(rule disjI2) +apply(rule disjI1) +(* ax case *) +apply(simp add: trm.inject) +apply(rule_tac x="y" in exI) +apply(erule conjE) +apply(thin_tac "[a].M = [aa].Ax y aa") +apply(simp add: alpha) +apply(erule disjE) +apply(simp) +apply(simp) +apply(simp add: rename_fresh) +apply(rule disjI2) +apply(rule disjI2) +apply(rule disjI1) +(* not case *) +apply(simp add: trm.inject) +apply(erule conjE)+ +apply(generate_fresh "coname") +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(auto)[1] +apply(drule_tac c="c" in alpha_coname) +apply(simp add: fresh_prod fresh_atm abs_fresh) +apply(simp add: calc_atm) +apply(rule exI)+ +apply(rule conjI) +apply(rule refl) +apply(generate_fresh "name") +apply(simp add: calc_atm abs_fresh fresh_prod fresh_atm fresh_left) +apply(auto)[1] +apply(drule_tac z="ca" in alpha_name) +apply(simp add: fresh_prod fresh_atm abs_fresh) +apply(simp add: calc_atm) +apply(rule exI)+ +apply(rule conjI) +apply(rule refl) +apply(auto simp add: calc_atm abs_fresh fresh_left)[1] +apply(case_tac "y=x") +apply(perm_simp) +apply(perm_simp) +apply(case_tac "aa=a") +apply(perm_simp) +apply(perm_simp) +(* and1 case *) +apply(rule disjI2) +apply(rule disjI2) +apply(rule disjI2) +apply(rule disjI1) +apply(simp add: trm.inject) +apply(erule conjE)+ +apply(generate_fresh "coname") +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(auto)[1] +apply(drule_tac c="c" in alpha_coname) +apply(simp add: fresh_prod fresh_atm abs_fresh) +apply(simp) +apply(rule exI)+ +apply(rule conjI) +apply(rule exI)+ +apply(rule_tac s="a" and t="[(a,c)]\[(b,c)]\b" in subst) +apply(simp add: calc_atm) +apply(rule refl) +apply(generate_fresh "name") +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(auto)[1] +apply(drule_tac z="ca" in alpha_name) +apply(simp add: fresh_prod fresh_atm abs_fresh) +apply(simp) +apply(rule exI)+ +apply(rule conjI) +apply(rule_tac s="x" and t="[(x,ca)]\[(y,ca)]\y" in subst) +apply(simp add: calc_atm) +apply(rule refl) +apply(auto simp add: fresh_left calc_atm abs_fresh split: if_splits)[1] +apply(generate_fresh "name") +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(auto)[1] +apply(drule_tac z="cb" in alpha_name) +apply(simp add: fresh_prod fresh_atm abs_fresh) +apply(simp) +apply(rule exI)+ +apply(rule conjI) +apply(rule_tac s="x" and t="[(x,cb)]\[(y,cb)]\y" in subst) +apply(simp add: calc_atm) +apply(rule refl) +apply(auto simp add: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] +apply(perm_simp)+ +apply(generate_fresh "name") +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(auto)[1] +apply(drule_tac z="cb" in alpha_name) +apply(simp add: fresh_prod fresh_atm abs_fresh) +apply(simp) +apply(rule exI)+ +apply(rule conjI) +apply(rule_tac s="x" and t="[(x,cb)]\[(y,cb)]\y" in subst) +apply(simp add: calc_atm) +apply(rule refl) +apply(auto simp add: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] +apply(perm_simp)+ +apply(generate_fresh "name") +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(auto)[1] +apply(drule_tac z="cb" in alpha_name) +apply(simp add: fresh_prod fresh_atm abs_fresh) +apply(simp) +apply(rule exI)+ +apply(rule conjI) +apply(rule_tac s="x" and t="[(x,cb)]\[(y,cb)]\y" in subst) +apply(simp add: calc_atm) +apply(rule refl) +apply(auto simp add: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] +apply(perm_simp)+ +(* and2 case *) +apply(rule disjI2) +apply(rule disjI2) +apply(rule disjI2) +apply(rule disjI2) +apply(rule disjI1) +apply(simp add: trm.inject) +apply(erule conjE)+ +apply(generate_fresh "coname") +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(auto)[1] +apply(drule_tac c="c" in alpha_coname) +apply(simp add: fresh_prod fresh_atm abs_fresh) +apply(simp) +apply(rule exI)+ +apply(rule conjI) +apply(rule_tac s="a" and t="[(a,c)]\[(b,c)]\b" in subst) +apply(simp add: calc_atm) +apply(rule refl) +apply(generate_fresh "name") +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(auto)[1] +apply(drule_tac z="ca" in alpha_name) +apply(simp add: fresh_prod fresh_atm abs_fresh) +apply(simp) +apply(rule exI)+ +apply(rule conjI) +apply(rule_tac s="x" and t="[(x,ca)]\[(y,ca)]\y" in subst) +apply(simp add: calc_atm) +apply(rule refl) +apply(auto simp add: fresh_left calc_atm abs_fresh split: if_splits)[1] +apply(generate_fresh "name") +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(auto)[1] +apply(drule_tac z="cb" in alpha_name) +apply(simp add: fresh_prod fresh_atm abs_fresh) +apply(simp) +apply(rule exI)+ +apply(rule conjI) +apply(rule_tac s="x" and t="[(x,cb)]\[(y,cb)]\y" in subst) +apply(simp add: calc_atm) +apply(rule refl) +apply(auto simp add: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] +apply(perm_simp)+ +apply(generate_fresh "name") +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(auto)[1] +apply(drule_tac z="cb" in alpha_name) +apply(simp add: fresh_prod fresh_atm abs_fresh) +apply(simp) +apply(rule exI)+ +apply(rule conjI) +apply(rule_tac s="x" and t="[(x,cb)]\[(y,cb)]\y" in subst) +apply(simp add: calc_atm) +apply(rule refl) +apply(auto simp add: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] +apply(perm_simp)+ +apply(generate_fresh "name") +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(auto)[1] +apply(drule_tac z="cb" in alpha_name) +apply(simp add: fresh_prod fresh_atm abs_fresh) +apply(simp) +apply(rule exI)+ +apply(rule conjI) +apply(rule_tac s="x" and t="[(x,cb)]\[(y,cb)]\y" in subst) +apply(simp add: calc_atm) +apply(rule refl) +apply(auto simp add: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] +apply(perm_simp)+ +(* or1 case *) +apply(rule disjI2) +apply(rule disjI2) +apply(rule disjI2) +apply(rule disjI2) +apply(rule disjI2) +apply(rule disjI1) +apply(simp add: trm.inject) +apply(erule conjE)+ +apply(generate_fresh "coname") +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(auto)[1] +apply(drule_tac c="c" in alpha_coname) +apply(simp add: fresh_prod fresh_atm abs_fresh) +apply(simp) +apply(rule exI)+ +apply(rule conjI) +apply(rule_tac s="a" and t="[(a,c)]\[(b,c)]\b" in subst) +apply(simp add: calc_atm) +apply(rule refl) +apply(generate_fresh "name") +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(auto)[1] +apply(drule_tac z="ca" in alpha_name) +apply(simp add: fresh_prod fresh_atm abs_fresh) +apply(simp) +apply(rule exI)+ +apply(rule conjI) +apply(rule exI)+ +apply(rule_tac s="x" and t="[(x,ca)]\[(y,ca)]\y" in subst) +apply(simp add: calc_atm) +apply(rule refl) +apply(auto simp add: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] +apply(perm_simp)+ +apply(generate_fresh "name") +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(auto)[1] +apply(drule_tac z="cb" in alpha_name) +apply(simp add: fresh_prod fresh_atm abs_fresh) +apply(simp) +apply(rule exI)+ +apply(rule conjI) +apply(rule exI)+ +apply(rule_tac s="x" and t="[(x,cb)]\[(y,cb)]\y" in subst) +apply(simp add: calc_atm) +apply(rule refl) +apply(auto simp add: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] +apply(perm_simp)+ +(* or2 case *) +apply(rule disjI2) +apply(rule disjI2) +apply(rule disjI2) +apply(rule disjI2) +apply(rule disjI2) +apply(rule disjI2) +apply(rule disjI1) +apply(simp add: trm.inject) +apply(erule conjE)+ +apply(generate_fresh "coname") +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(auto)[1] +apply(drule_tac c="c" in alpha_coname) +apply(simp add: fresh_prod fresh_atm abs_fresh) +apply(simp) +apply(rule exI)+ +apply(rule conjI) +apply(rule_tac s="a" and t="[(a,c)]\[(b,c)]\b" in subst) +apply(simp add: calc_atm) +apply(rule refl) +apply(generate_fresh "name") +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(auto)[1] +apply(drule_tac z="ca" in alpha_name) +apply(simp add: fresh_prod fresh_atm abs_fresh) +apply(simp) +apply(rule exI)+ +apply(rule conjI) +apply(rule_tac s="x" and t="[(x,ca)]\[(y,ca)]\y" in subst) +apply(simp add: calc_atm) +apply(rule refl) +apply(auto simp add: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] +apply(perm_simp)+ +apply(generate_fresh "name") +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(auto)[1] +apply(drule_tac z="cb" in alpha_name) +apply(simp add: fresh_prod fresh_atm abs_fresh) +apply(simp) +apply(rule exI)+ +apply(rule conjI) +apply(rule_tac s="x" and t="[(x,cb)]\[(y,cb)]\y" in subst) +apply(simp add: calc_atm) +apply(rule refl) +apply(auto simp add: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] +apply(perm_simp)+ +(* imp-case *) +apply(rule disjI2) +apply(rule disjI2) +apply(rule disjI2) +apply(rule disjI2) +apply(rule disjI2) +apply(rule disjI2) +apply(rule disjI2) +apply(simp add: trm.inject) +apply(erule conjE)+ +apply(generate_fresh "coname") +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(auto)[1] +apply(drule_tac c="ca" in alpha_coname) +apply(simp add: fresh_prod fresh_atm abs_fresh) +apply(simp) +apply(rule exI)+ +apply(rule conjI) +apply(rule_tac s="a" and t="[(a,ca)]\[(b,ca)]\b" in subst) +apply(simp add: calc_atm) +apply(rule refl) +apply(generate_fresh "name") +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(auto)[1] +apply(drule_tac z="cb" in alpha_name) +apply(simp add: fresh_prod fresh_atm abs_fresh) +apply(simp) +apply(rule exI)+ +apply(rule conjI) +apply(rule_tac s="x" and t="[(x,cb)]\[(z,cb)]\z" in subst) +apply(simp add: calc_atm) +apply(rule refl) +apply(auto simp add: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] +apply(perm_simp)+ +apply(generate_fresh "name") +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(auto)[1] +apply(drule_tac z="cc" in alpha_name) +apply(simp add: fresh_prod fresh_atm abs_fresh) +apply(simp) +apply(rule exI)+ +apply(rule conjI) +apply(rule_tac s="x" and t="[(x,cc)]\[(z,cc)]\z" in subst) +apply(simp add: calc_atm) +apply(rule refl) +apply(auto simp add: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] +apply(perm_simp)+ +done + +inductive2 + c_redu :: "trm \ trm \ bool" ("_ \\<^isub>c _" [100,100] 100) +where + left[intro]: "\\fic M a; a\N; x\M\ \ Cut .M (x).N \\<^isub>c M{a:=(x).N}" +| right[intro]: "\\fin N x; a\N; x\M\ \ Cut .M (x).N \\<^isub>c N{x:=.M}" + +equivariance c_redu + +nominal_inductive c_redu + by (simp_all add: abs_fresh subst_fresh) + +lemma better_left[intro]: + shows "\fic M a \ Cut .M (x).N \\<^isub>c M{a:=(x).N}" +proof - + assume not_fic: "\fic M a" + obtain x'::"name" where fs1: "x'\(N,M,x)" by (rule exists_fresh(1), rule fin_supp, blast) + obtain a'::"coname" where fs2: "a'\(a,M,N)" by (rule exists_fresh(2), rule fin_supp, blast) + have "Cut .M (x).N = Cut .([(a',a)]\M) (x').([(x',x)]\N)" + using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm) + also have "\ \\<^isub>c ([(a',a)]\M){a':=(x').([(x',x)]\N)}" using fs1 fs2 not_fic + apply - + apply(rule left) + apply(clarify) + apply(drule_tac a'="a" in fic_rename) + apply(simp add: perm_swap) + apply(simp add: fresh_left calc_atm)+ + done + also have "\ = M{a:=(x).N}" using fs1 fs2 + by (simp add: subst_rename[symmetric] fresh_atm fresh_prod fresh_left calc_atm) + finally show ?thesis by simp +qed + +lemma better_right[intro]: + shows "\fin N x \ Cut .M (x).N \\<^isub>c N{x:=.M}" +proof - + assume not_fin: "\fin N x" + obtain x'::"name" where fs1: "x'\(N,M,x)" by (rule exists_fresh(1), rule fin_supp, blast) + obtain a'::"coname" where fs2: "a'\(a,M,N)" by (rule exists_fresh(2), rule fin_supp, blast) + have "Cut .M (x).N = Cut .([(a',a)]\M) (x').([(x',x)]\N)" + using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm) + also have "\ \\<^isub>c ([(x',x)]\N){x':=.([(a',a)]\M)}" using fs1 fs2 not_fin + apply - + apply(rule right) + apply(clarify) + apply(drule_tac x'="x" in fin_rename) + apply(simp add: perm_swap) + apply(simp add: fresh_left calc_atm)+ + done + also have "\ = N{x:=.M}" using fs1 fs2 + by (simp add: subst_rename[symmetric] fresh_atm fresh_prod fresh_left calc_atm) + finally show ?thesis by simp +qed + +lemma fresh_c_redu: + fixes x::"name" + and c::"coname" + shows "M \\<^isub>c M' \ x\M \ x\M'" + and "M \\<^isub>c M' \ c\M \ c\M'" +apply - +apply(induct rule: c_redu.induct) +apply(auto simp add: abs_fresh rename_fresh subst_fresh) +apply(induct rule: c_redu.induct) +apply(auto simp add: abs_fresh rename_fresh subst_fresh) +done + +inductive2 + a_redu :: "trm \ trm \ bool" ("_ \\<^isub>a _" [100,100] 100) +where + al_redu[intro]: "M\\<^isub>l M' \ M \\<^isub>a M'" +| ac_redu[intro]: "M\\<^isub>c M' \ M \\<^isub>a M'" +| a_Cut_l: "\a\N; x\M; M\\<^isub>a M'\ \ Cut .M (x).N \\<^isub>a Cut .M' (x).N" +| a_Cut_r: "\a\N; x\M; N\\<^isub>a N'\ \ Cut .M (x).N \\<^isub>a Cut .M (x).N'" +| a_NotL[intro]: "M\\<^isub>a M' \ NotL .M x \\<^isub>a NotL .M' x" +| a_NotR[intro]: "M\\<^isub>a M' \ NotR (x).M a \\<^isub>a NotR (x).M' a" +| a_AndR_l: "\a\(N,c); b\(M,c); b\a; M\\<^isub>a M'\ \ AndR .M .N c \\<^isub>a AndR .M' .N c" +| a_AndR_r: "\a\(N,c); b\(M,c); b\a; N\\<^isub>a N'\ \ AndR .M .N c \\<^isub>a AndR .M .N' c" +| a_AndL1: "\x\y; M\\<^isub>a M'\ \ AndL1 (x).M y \\<^isub>a AndL1 (x).M' y" +| a_AndL2: "\x\y; M\\<^isub>a M'\ \ AndL2 (x).M y \\<^isub>a AndL2 (x).M' y" +| a_OrL_l: "\x\(N,z); y\(M,z); y\x; M\\<^isub>a M'\ \ OrL (x).M (y).N z \\<^isub>a OrL (x).M' (y).N z" +| a_OrL_r: "\x\(N,z); y\(M,z); y\x; N\\<^isub>a N'\ \ OrL (x).M (y).N z \\<^isub>a OrL (x).M (y).N' z" +| a_OrR1: "\a\b; M\\<^isub>a M'\ \ OrR1 .M b \\<^isub>a OrR1 .M' b" +| a_OrR2: "\a\b; M\\<^isub>a M'\ \ OrR2 .M b \\<^isub>a OrR2 .M' b" +| a_ImpL_l: "\a\N; x\(M,y); M\\<^isub>a M'\ \ ImpL .M (x).N y \\<^isub>a ImpL .M' (x).N y" +| a_ImpL_r: "\a\N; x\(M,y); N\\<^isub>a N'\ \ ImpL .M (x).N y \\<^isub>a ImpL .M (x).N' y" +| a_ImpR: "\a\b; M\\<^isub>a M'\ \ ImpR (x)..M b \\<^isub>a ImpR (x)..M' b" + +lemma fresh_a_redu: + fixes x::"name" + and c::"coname" + shows "M \\<^isub>a M' \ x\M \ x\M'" + and "M \\<^isub>a M' \ c\M \ c\M'" +apply - +apply(induct rule: a_redu.induct) +apply(simp add: fresh_l_redu) +apply(simp add: fresh_c_redu) +apply(auto simp add: abs_fresh abs_supp fin_supp) +apply(induct rule: a_redu.induct) +apply(simp add: fresh_l_redu) +apply(simp add: fresh_c_redu) +apply(auto simp add: abs_fresh abs_supp fin_supp) +done + +equivariance a_redu + +nominal_inductive a_redu + by (simp_all add: abs_fresh fresh_atm fresh_prod abs_supp fin_supp fresh_a_redu) + +lemma better_CutL_intro[intro]: + shows "M\\<^isub>a M' \ Cut .M (x).N \\<^isub>a Cut .M' (x).N" +proof - + assume red: "M\\<^isub>a M'" + obtain x'::"name" where fs1: "x'\(M,N,x)" by (rule exists_fresh(1), rule fin_supp, blast) + obtain a'::"coname" where fs2: "a'\(M,N,a)" by (rule exists_fresh(2), rule fin_supp, blast) + have "Cut .M (x).N = Cut .([(a',a)]\M) (x').([(x',x)]\N)" + using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm) + also have "\ \\<^isub>a Cut .([(a',a)]\M') (x').([(x',x)]\N)" using fs1 fs2 red + by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt) + also have "\ = Cut .M' (x).N" + using fs1 fs2 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu) + finally show ?thesis by simp +qed + +lemma better_CutR_intro[intro]: + shows "N\\<^isub>a N' \ Cut .M (x).N \\<^isub>a Cut .M (x).N'" +proof - + assume red: "N\\<^isub>a N'" + obtain x'::"name" where fs1: "x'\(M,N,x)" by (rule exists_fresh(1), rule fin_supp, blast) + obtain a'::"coname" where fs2: "a'\(M,N,a)" by (rule exists_fresh(2), rule fin_supp, blast) + have "Cut .M (x).N = Cut .([(a',a)]\M) (x').([(x',x)]\N)" + using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm) + also have "\ \\<^isub>a Cut .([(a',a)]\M) (x').([(x',x)]\N')" using fs1 fs2 red + by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt) + also have "\ = Cut .M (x).N'" + using fs1 fs2 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu) + finally show ?thesis by simp +qed + +lemma better_AndRL_intro[intro]: + shows "M\\<^isub>a M' \ AndR .M .N c \\<^isub>a AndR .M' .N c" +proof - + assume red: "M\\<^isub>a M'" + obtain b'::"coname" where fs1: "b'\(M,N,a,b,c)" by (rule exists_fresh(2), rule fin_supp, blast) + obtain a'::"coname" where fs2: "a'\(M,N,a,b,c,b')" by (rule exists_fresh(2), rule fin_supp, blast) + have "AndR .M .N c = AndR .([(a',a)]\M) .([(b',b)]\N) c" + using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm) + also have "\ \\<^isub>a AndR .([(a',a)]\M') .([(b',b)]\N) c" using fs1 fs2 red + by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt fresh_atm fresh_prod) + also have "\ = AndR .M' .N c" + using fs1 fs2 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu) + finally show ?thesis by simp +qed + +lemma better_AndRR_intro[intro]: + shows "N\\<^isub>a N' \ AndR .M .N c \\<^isub>a AndR .M .N' c" +proof - + assume red: "N\\<^isub>a N'" + obtain b'::"coname" where fs1: "b'\(M,N,a,b,c)" by (rule exists_fresh(2), rule fin_supp, blast) + obtain a'::"coname" where fs2: "a'\(M,N,a,b,c,b')" by (rule exists_fresh(2), rule fin_supp, blast) + have "AndR .M .N c = AndR .([(a',a)]\M) .([(b',b)]\N) c" + using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm) + also have "\ \\<^isub>a AndR .([(a',a)]\M) .([(b',b)]\N') c" using fs1 fs2 red + by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt fresh_atm fresh_prod) + also have "\ = AndR .M .N' c" + using fs1 fs2 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu) + finally show ?thesis by simp +qed + +lemma better_AndL1_intro[intro]: + shows "M\\<^isub>a M' \ AndL1 (x).M y \\<^isub>a AndL1 (x).M' y" +proof - + assume red: "M\\<^isub>a M'" + obtain x'::"name" where fs1: "x'\(M,y,x)" by (rule exists_fresh(1), rule fin_supp, blast) + have "AndL1 (x).M y = AndL1 (x').([(x',x)]\M) y" + using fs1 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm) + also have "\ \\<^isub>a AndL1 (x').([(x',x)]\M') y" using fs1 red + by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt fresh_atm fresh_prod) + also have "\ = AndL1 (x).M' y" + using fs1 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu) + finally show ?thesis by simp +qed + +lemma better_AndL2_intro[intro]: + shows "M\\<^isub>a M' \ AndL2 (x).M y \\<^isub>a AndL2 (x).M' y" +proof - + assume red: "M\\<^isub>a M'" + obtain x'::"name" where fs1: "x'\(M,y,x)" by (rule exists_fresh(1), rule fin_supp, blast) + have "AndL2 (x).M y = AndL2 (x').([(x',x)]\M) y" + using fs1 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm) + also have "\ \\<^isub>a AndL2 (x').([(x',x)]\M') y" using fs1 red + by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt fresh_atm fresh_prod) + also have "\ = AndL2 (x).M' y" + using fs1 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu) + finally show ?thesis by simp +qed + +lemma better_OrLL_intro[intro]: + shows "M\\<^isub>a M' \ OrL (x).M (y).N z \\<^isub>a OrL (x).M' (y).N z" +proof - + assume red: "M\\<^isub>a M'" + obtain x'::"name" where fs1: "x'\(M,N,x,y,z)" by (rule exists_fresh(1), rule fin_supp, blast) + obtain y'::"name" where fs2: "y'\(M,N,x,y,z,x')" by (rule exists_fresh(1), rule fin_supp, blast) + have "OrL (x).M (y).N z = OrL (x').([(x',x)]\M) (y').([(y',y)]\N) z" + using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm) + also have "\ \\<^isub>a OrL (x').([(x',x)]\M') (y').([(y',y)]\N) z" using fs1 fs2 red + by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt fresh_atm fresh_prod) + also have "\ = OrL (x).M' (y).N z" + using fs1 fs2 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu) + finally show ?thesis by simp +qed + +lemma better_OrLR_intro[intro]: + shows "N\\<^isub>a N' \ OrL (x).M (y).N z \\<^isub>a OrL (x).M (y).N' z" +proof - + assume red: "N\\<^isub>a N'" + obtain x'::"name" where fs1: "x'\(M,N,x,y,z)" by (rule exists_fresh(1), rule fin_supp, blast) + obtain y'::"name" where fs2: "y'\(M,N,x,y,z,x')" by (rule exists_fresh(1), rule fin_supp, blast) + have "OrL (x).M (y).N z = OrL (x').([(x',x)]\M) (y').([(y',y)]\N) z" + using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm) + also have "\ \\<^isub>a OrL (x').([(x',x)]\M) (y').([(y',y)]\N') z" using fs1 fs2 red + by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt fresh_atm fresh_prod) + also have "\ = OrL (x).M (y).N' z" + using fs1 fs2 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu) + finally show ?thesis by simp +qed + +lemma better_OrR1_intro[intro]: + shows "M\\<^isub>a M' \ OrR1 .M b \\<^isub>a OrR1 .M' b" +proof - + assume red: "M\\<^isub>a M'" + obtain a'::"coname" where fs1: "a'\(M,b,a)" by (rule exists_fresh(2), rule fin_supp, blast) + have "OrR1 .M b = OrR1 .([(a',a)]\M) b" + using fs1 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm) + also have "\ \\<^isub>a OrR1 .([(a',a)]\M') b" using fs1 red + by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt fresh_atm fresh_prod) + also have "\ = OrR1 .M' b" + using fs1 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu) + finally show ?thesis by simp +qed + +lemma better_OrR2_intro[intro]: + shows "M\\<^isub>a M' \ OrR2 .M b \\<^isub>a OrR2 .M' b" +proof - + assume red: "M\\<^isub>a M'" + obtain a'::"coname" where fs1: "a'\(M,b,a)" by (rule exists_fresh(2), rule fin_supp, blast) + have "OrR2 .M b = OrR2 .([(a',a)]\M) b" + using fs1 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm) + also have "\ \\<^isub>a OrR2 .([(a',a)]\M') b" using fs1 red + by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt fresh_atm fresh_prod) + also have "\ = OrR2 .M' b" + using fs1 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu) + finally show ?thesis by simp +qed + +lemma better_ImpLL_intro[intro]: + shows "M\\<^isub>a M' \ ImpL .M (x).N y \\<^isub>a ImpL .M' (x).N y" +proof - + assume red: "M\\<^isub>a M'" + obtain x'::"name" where fs1: "x'\(M,N,x,y)" by (rule exists_fresh(1), rule fin_supp, blast) + obtain a'::"coname" where fs2: "a'\(M,N,a)" by (rule exists_fresh(2), rule fin_supp, blast) + have "ImpL .M (x).N y = ImpL .([(a',a)]\M) (x').([(x',x)]\N) y" + using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm) + also have "\ \\<^isub>a ImpL .([(a',a)]\M') (x').([(x',x)]\N) y" using fs1 fs2 red + by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt fresh_atm fresh_prod) + also have "\ = ImpL .M' (x).N y" + using fs1 fs2 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu) + finally show ?thesis by simp +qed + +lemma better_ImpLR_intro[intro]: + shows "N\\<^isub>a N' \ ImpL .M (x).N y \\<^isub>a ImpL .M (x).N' y" +proof - + assume red: "N\\<^isub>a N'" + obtain x'::"name" where fs1: "x'\(M,N,x,y)" by (rule exists_fresh(1), rule fin_supp, blast) + obtain a'::"coname" where fs2: "a'\(M,N,a)" by (rule exists_fresh(2), rule fin_supp, blast) + have "ImpL .M (x).N y = ImpL .([(a',a)]\M) (x').([(x',x)]\N) y" + using fs1 fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm) + also have "\ \\<^isub>a ImpL .([(a',a)]\M) (x').([(x',x)]\N') y" using fs1 fs2 red + by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt fresh_atm fresh_prod) + also have "\ = ImpL .M (x).N' y" + using fs1 fs2 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu) + finally show ?thesis by simp +qed + +lemma better_ImpR_intro[intro]: + shows "M\\<^isub>a M' \ ImpR (x)..M b \\<^isub>a ImpR (x)..M' b" +proof - + assume red: "M\\<^isub>a M'" + obtain a'::"coname" where fs2: "a'\(M,a,b)" by (rule exists_fresh(2), rule fin_supp, blast) + have "ImpR (x)..M b = ImpR (x)..([(a',a)]\M) b" + using fs2 by (rule_tac sym, auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm) + also have "\ \\<^isub>a ImpR (x)..([(a',a)]\M') b" using fs2 red + by (auto intro: a_redu.intros simp add: fresh_left calc_atm a_redu.eqvt fresh_atm fresh_prod) + also have "\ = ImpR (x)..M' b" + using fs2 red by (auto simp add: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu) + finally show ?thesis by simp +qed + +text {* axioms do not reduce *} + +lemma ax_do_not_l_reduce: + shows "Ax x a \\<^isub>l M \ False" +by (erule_tac l_redu.cases) (simp_all add: trm.inject) + +lemma ax_do_not_c_reduce: + shows "Ax x a \\<^isub>c M \ False" +by (erule_tac c_redu.cases) (simp_all add: trm.inject) + +lemma ax_do_not_a_reduce: + shows "Ax x a \\<^isub>a M \ False" +apply(erule_tac a_redu.cases) +apply(auto simp add: trm.inject) +apply(drule ax_do_not_l_reduce) +apply(simp) +apply(drule ax_do_not_c_reduce) +apply(simp) +done + +lemma a_redu_NotL_elim: + assumes a: "NotL .M x \\<^isub>a R" + shows "\M'. R = NotL .M' x \ M\\<^isub>aM'" +using a +apply(erule_tac a_redu.cases, simp_all add: trm.inject) +apply(erule_tac l_redu.cases, simp_all add: trm.inject) +apply(erule_tac c_redu.cases, simp_all add: trm.inject) +apply(auto) +apply(rotate_tac 1) +apply(erule_tac a_redu.cases, simp_all add: trm.inject) +apply(erule_tac l_redu.cases, simp_all add: trm.inject) +apply(erule_tac c_redu.cases, simp_all add: trm.inject) +apply(auto simp add: alpha a_redu.eqvt) +apply(rule_tac x="([(a,aa)]\M'a)" in exI) +apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu) +apply(simp add: perm_swap) +apply(rule_tac x="([(a,aaa)]\M'a)" in exI) +apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu) +apply(simp add: perm_swap) +done + +lemma a_redu_NotR_elim: + assumes a: "NotR (x).M a \\<^isub>a R" + shows "\M'. R = NotR (x).M' a \ M\\<^isub>aM'" +using a +apply(erule_tac a_redu.cases, simp_all add: trm.inject) +apply(erule_tac l_redu.cases, simp_all add: trm.inject) +apply(erule_tac c_redu.cases, simp_all add: trm.inject) +apply(auto) +apply(rotate_tac 1) +apply(erule_tac a_redu.cases, simp_all add: trm.inject) +apply(erule_tac l_redu.cases, simp_all add: trm.inject) +apply(erule_tac c_redu.cases, simp_all add: trm.inject) +apply(auto simp add: alpha a_redu.eqvt) +apply(rule_tac x="([(x,xa)]\M'a)" in exI) +apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu) +apply(simp add: perm_swap) +apply(rule_tac x="([(x,xaa)]\M'a)" in exI) +apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu) +apply(simp add: perm_swap) +done + +lemma a_redu_AndR_elim: + assumes a: "AndR .M .N c\\<^isub>a R" + shows "(\M'. R = AndR .M' .N c \ M\\<^isub>aM') \ (\N'. R = AndR .M .N' c \ N\\<^isub>aN')" +using a +apply(erule_tac a_redu.cases, simp_all add: trm.inject) +apply(erule_tac l_redu.cases, simp_all add: trm.inject) +apply(erule_tac c_redu.cases, simp_all add: trm.inject) +apply(rotate_tac 6) +apply(erule_tac a_redu.cases, simp_all add: trm.inject) +apply(erule_tac l_redu.cases, simp_all add: trm.inject) +apply(erule_tac c_redu.cases, simp_all add: trm.inject) +apply(rule disjI1) +apply(auto simp add: alpha a_redu.eqvt)[1] +apply(rule_tac x="([(a,aa)]\M'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(a,aa)]\M'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(a,aa)]\M'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(a,aa)]\M'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(a,aaa)]\M'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(a,aaa)]\M'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(a,aaa)]\M'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(a,aaa)]\M'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule disjI2) +apply(auto simp add: alpha a_redu.eqvt)[1] +apply(rule_tac x="([(b,ba)]\N')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(b,baa)]\N')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(b,ba)]\N')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(b,baa)]\N')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(b,ba)]\N')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(b,baa)]\N')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(b,ba)]\N')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(b,baa)]\N')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rotate_tac 6) +apply(erule_tac a_redu.cases, simp_all add: trm.inject) +apply(erule_tac l_redu.cases, simp_all add: trm.inject) +apply(erule_tac c_redu.cases, simp_all add: trm.inject) +apply(rule disjI1) +apply(auto simp add: alpha a_redu.eqvt)[1] +apply(rule_tac x="([(a,aa)]\M')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(a,aa)]\M')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(a,aa)]\M')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(a,aa)]\M')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(a,aaa)]\M')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(a,aaa)]\M')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(a,aaa)]\M')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(a,aaa)]\M')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule disjI2) +apply(auto simp add: alpha a_redu.eqvt)[1] +apply(rule_tac x="([(b,ba)]\N'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(b,ba)]\N'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(b,ba)]\N'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(b,ba)]\N'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(b,baa)]\N'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(b,baa)]\N'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(b,baa)]\N'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(b,baa)]\N'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +done + +lemma a_redu_AndL1_elim: + assumes a: "AndL1 (x).M y \\<^isub>a R" + shows "\M'. R = AndL1 (x).M' y \ M\\<^isub>aM'" +using a +apply(erule_tac a_redu.cases, simp_all add: trm.inject) +apply(erule_tac l_redu.cases, simp_all add: trm.inject) +apply(erule_tac c_redu.cases, simp_all add: trm.inject) +apply(auto) +apply(rotate_tac 2) +apply(erule_tac a_redu.cases, simp_all add: trm.inject) +apply(erule_tac l_redu.cases, simp_all add: trm.inject) +apply(erule_tac c_redu.cases, simp_all add: trm.inject) +apply(auto simp add: alpha a_redu.eqvt) +apply(rule_tac x="([(x,xa)]\M'a)" in exI) +apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu) +apply(simp add: perm_swap) +apply(rule_tac x="([(x,xaa)]\M'a)" in exI) +apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu) +apply(simp add: perm_swap) +done + +lemma a_redu_AndL2_elim: + assumes a: "AndL2 (x).M y \\<^isub>a R" + shows "\M'. R = AndL2 (x).M' y \ M\\<^isub>aM'" +using a +apply(erule_tac a_redu.cases, simp_all add: trm.inject) +apply(erule_tac l_redu.cases, simp_all add: trm.inject) +apply(erule_tac c_redu.cases, simp_all add: trm.inject) +apply(auto) +apply(rotate_tac 2) +apply(erule_tac a_redu.cases, simp_all add: trm.inject) +apply(erule_tac l_redu.cases, simp_all add: trm.inject) +apply(erule_tac c_redu.cases, simp_all add: trm.inject) +apply(auto simp add: alpha a_redu.eqvt) +apply(rule_tac x="([(x,xa)]\M'a)" in exI) +apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu) +apply(simp add: perm_swap) +apply(rule_tac x="([(x,xaa)]\M'a)" in exI) +apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu) +apply(simp add: perm_swap) +done + +lemma a_redu_OrL_elim: + assumes a: "OrL (x).M (y).N z\\<^isub>a R" + shows "(\M'. R = OrL (x).M' (y).N z \ M\\<^isub>aM') \ (\N'. R = OrL (x).M (y).N' z \ N\\<^isub>aN')" +using a +apply(erule_tac a_redu.cases, simp_all add: trm.inject) +apply(erule_tac l_redu.cases, simp_all add: trm.inject) +apply(erule_tac c_redu.cases, simp_all add: trm.inject) +apply(rotate_tac 6) +apply(erule_tac a_redu.cases, simp_all add: trm.inject) +apply(erule_tac l_redu.cases, simp_all add: trm.inject) +apply(erule_tac c_redu.cases, simp_all add: trm.inject) +apply(rule disjI1) +apply(auto simp add: alpha a_redu.eqvt)[1] +apply(rule_tac x="([(x,xa)]\M'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(x,xa)]\M'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(x,xa)]\M'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(x,xa)]\M'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(x,xaa)]\M'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(x,xaa)]\M'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(x,xaa)]\M'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(x,xaa)]\M'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule disjI2) +apply(auto simp add: alpha a_redu.eqvt)[1] +apply(rule_tac x="([(y,ya)]\N')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(y,yaa)]\N')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(y,ya)]\N')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(y,yaa)]\N')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(y,ya)]\N')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(y,yaa)]\N')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(y,ya)]\N')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(y,yaa)]\N')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rotate_tac 6) +apply(erule_tac a_redu.cases, simp_all add: trm.inject) +apply(erule_tac l_redu.cases, simp_all add: trm.inject) +apply(erule_tac c_redu.cases, simp_all add: trm.inject) +apply(rule disjI1) +apply(auto simp add: alpha a_redu.eqvt)[1] +apply(rule_tac x="([(x,xa)]\M')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(x,xa)]\M')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(x,xa)]\M')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(x,xa)]\M')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(x,xaa)]\M')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(x,xaa)]\M')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(x,xaa)]\M')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(x,xaa)]\M')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule disjI2) +apply(auto simp add: alpha a_redu.eqvt)[1] +apply(rule_tac x="([(y,ya)]\N'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(y,ya)]\N'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(y,ya)]\N'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(y,ya)]\N'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(y,yaa)]\N'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(y,yaa)]\N'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(y,yaa)]\N'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(y,yaa)]\N'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +done + +lemma a_redu_OrR1_elim: + assumes a: "OrR1 .M b \\<^isub>a R" + shows "\M'. R = OrR1 .M' b \ M\\<^isub>aM'" +using a +apply(erule_tac a_redu.cases, simp_all add: trm.inject) +apply(erule_tac l_redu.cases, simp_all add: trm.inject) +apply(erule_tac c_redu.cases, simp_all add: trm.inject) +apply(auto) +apply(rotate_tac 2) +apply(erule_tac a_redu.cases, simp_all add: trm.inject) +apply(erule_tac l_redu.cases, simp_all add: trm.inject) +apply(erule_tac c_redu.cases, simp_all add: trm.inject) +apply(auto simp add: alpha a_redu.eqvt) +apply(rule_tac x="([(a,aa)]\M'a)" in exI) +apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu) +apply(simp add: perm_swap) +apply(rule_tac x="([(a,aaa)]\M'a)" in exI) +apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu) +apply(simp add: perm_swap) +done + +lemma a_redu_OrR2_elim: + assumes a: "OrR2 .M b \\<^isub>a R" + shows "\M'. R = OrR2 .M' b \ M\\<^isub>aM'" +using a +apply(erule_tac a_redu.cases, simp_all add: trm.inject) +apply(erule_tac l_redu.cases, simp_all add: trm.inject) +apply(erule_tac c_redu.cases, simp_all add: trm.inject) +apply(auto) +apply(rotate_tac 2) +apply(erule_tac a_redu.cases, simp_all add: trm.inject) +apply(erule_tac l_redu.cases, simp_all add: trm.inject) +apply(erule_tac c_redu.cases, simp_all add: trm.inject) +apply(auto simp add: alpha a_redu.eqvt) +apply(rule_tac x="([(a,aa)]\M'a)" in exI) +apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu) +apply(simp add: perm_swap) +apply(rule_tac x="([(a,aaa)]\M'a)" in exI) +apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu) +apply(simp add: perm_swap) +done + +lemma a_redu_ImpL_elim: + assumes a: "ImpL .M (y).N z\\<^isub>a R" + shows "(\M'. R = ImpL .M' (y).N z \ M\\<^isub>aM') \ (\N'. R = ImpL .M (y).N' z \ N\\<^isub>aN')" +using a +apply(erule_tac a_redu.cases, simp_all add: trm.inject) +apply(erule_tac l_redu.cases, simp_all add: trm.inject) +apply(erule_tac c_redu.cases, simp_all add: trm.inject) +apply(rotate_tac 5) +apply(erule_tac a_redu.cases, simp_all add: trm.inject) +apply(erule_tac l_redu.cases, simp_all add: trm.inject) +apply(erule_tac c_redu.cases, simp_all add: trm.inject) +apply(rule disjI1) +apply(auto simp add: alpha a_redu.eqvt)[1] +apply(rule_tac x="([(a,aa)]\M'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(a,aa)]\M'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(a,aa)]\M'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(a,aa)]\M'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(a,aaa)]\M'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(a,aaa)]\M'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(a,aaa)]\M'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(a,aaa)]\M'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule disjI2) +apply(auto simp add: alpha a_redu.eqvt)[1] +apply(rule_tac x="([(y,xa)]\N')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(y,xa)]\N')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(y,xa)]\N')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(y,xa)]\N')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(y,xa)]\N')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(y,xa)]\N')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(y,xa)]\N')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(y,xa)]\N')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rotate_tac 5) +apply(erule_tac a_redu.cases, simp_all add: trm.inject) +apply(erule_tac l_redu.cases, simp_all add: trm.inject) +apply(erule_tac c_redu.cases, simp_all add: trm.inject) +apply(rule disjI1) +apply(auto simp add: alpha a_redu.eqvt)[1] +apply(rule_tac x="([(a,aa)]\M')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(a,aa)]\M')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(a,aa)]\M')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(a,aa)]\M')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(a,aaa)]\M')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(a,aaa)]\M')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(a,aaa)]\M')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(a,aaa)]\M')" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule disjI2) +apply(auto simp add: alpha a_redu.eqvt)[1] +apply(rule_tac x="([(y,xa)]\N'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(y,xa)]\N'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(y,xa)]\N'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(y,xa)]\N'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(y,xa)]\N'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(y,xa)]\N'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(y,xa)]\N'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +apply(rule_tac x="([(y,xa)]\N'a)" in exI) +apply(auto simp add: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] +done + +lemma a_redu_ImpR_elim: + assumes a: "ImpR (x)..M b \\<^isub>a R" + shows "\M'. R = ImpR (x)..M' b \ M\\<^isub>aM'" +using a +apply(erule_tac a_redu.cases, simp_all add: trm.inject) +apply(erule_tac l_redu.cases, simp_all add: trm.inject) +apply(erule_tac c_redu.cases, simp_all add: trm.inject) +apply(auto) +apply(rotate_tac 2) +apply(erule_tac a_redu.cases, simp_all add: trm.inject) +apply(erule_tac l_redu.cases, simp_all add: trm.inject) +apply(erule_tac c_redu.cases, simp_all add: trm.inject) +apply(auto simp add: alpha a_redu.eqvt abs_perm abs_fresh) +apply(rule_tac x="([(a,aa)]\M'a)" in exI) +apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap) +apply(rule_tac x="([(a,aaa)]\M'a)" in exI) +apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap) +apply(rule_tac x="([(a,aa)]\M'a)" in exI) +apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap) +apply(rule_tac x="([(a,aaa)]\M'a)" in exI) +apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap) +apply(rule_tac x="([(x,xa)]\M'a)" in exI) +apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap) +apply(rule_tac x="([(x,xa)]\M'a)" in exI) +apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap) +apply(rule_tac x="([(a,aa)]\[(x,xa)]\M'a)" in exI) +apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap) +apply(rule sym) +apply(rule trans) +apply(rule perm_compose) +apply(simp add: calc_atm perm_swap) +apply(rule_tac x="([(a,aaa)]\[(x,xa)]\M'a)" in exI) +apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap) +apply(rule sym) +apply(rule trans) +apply(rule perm_compose) +apply(simp add: calc_atm perm_swap) +apply(rule_tac x="([(x,xaa)]\M'a)" in exI) +apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap) +apply(rule_tac x="([(x,xaa)]\M'a)" in exI) +apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap) +apply(rule_tac x="([(a,aa)]\[(x,xaa)]\M'a)" in exI) +apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap) +apply(rule sym) +apply(rule trans) +apply(rule perm_compose) +apply(simp add: calc_atm perm_swap) +apply(rule_tac x="([(a,aaa)]\[(x,xaa)]\M'a)" in exI) +apply(auto simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap) +apply(rule sym) +apply(rule trans) +apply(rule perm_compose) +apply(simp add: calc_atm perm_swap) +done + +text {* Transitive Closure*} + +abbreviation + a_star_redu :: "trm \ trm \ bool" ("_ \\<^isub>a* _" [100,100] 100) +where + "M \\<^isub>a* M' \ (a_redu)^** M M'" + +lemma a_starI: + assumes a: "M \\<^isub>a M'" + shows "M \\<^isub>a* M'" +using a by blast + +lemma a_star_refl: + shows "M \\<^isub>a* M" + by blast + +lemma a_starE: + assumes a: "M \\<^isub>a* M'" + shows "M = M' \ (\N. M \\<^isub>a N \ N \\<^isub>a* M')" +using a +by (induct) (auto) + +lemma a_star_refl: + shows "M \\<^isub>a* M" + by blast + +lemma a_star_trans[trans]: + assumes a1: "M1\\<^isub>a* M2" + and a2: "M2\\<^isub>a* M3" + shows "M1 \\<^isub>a* M3" +using a2 a1 +by (induct) (auto) + +text {* congruence rules for \\<^isub>a* *} + +lemma ax_do_not_a_star_reduce: + shows "Ax x a \\<^isub>a* M \ M = Ax x a" +apply(induct set: rtrancl) +apply(auto) +apply(drule ax_do_not_a_reduce) +apply(simp) +done + + +lemma a_star_CutL: + "M \\<^isub>a* M' \ Cut .M (x).N \\<^isub>a* Cut .M' (x).N" +by (induct set: rtrancl) (blast intro: rtrancl.rtrancl_into_rtrancl)+ + +lemma a_star_CutR: + "N \\<^isub>a* N'\ Cut .M (x).N \\<^isub>a* Cut .M (x).N'" +by (induct set: rtrancl) (blast intro: rtrancl.rtrancl_into_rtrancl)+ + +lemma a_star_Cut: + "\M \\<^isub>a* M'; N \\<^isub>a* N'\ \ Cut .M (x).N \\<^isub>a* Cut .M' (x).N'" +by (blast intro!: a_star_CutL a_star_CutR intro: rtrancl_trans') + +lemma a_star_NotR: + "M \\<^isub>a* M' \ NotR (x).M a \\<^isub>a* NotR (x).M' a" +by (induct set: rtrancl) (blast intro: rtrancl.rtrancl_into_rtrancl)+ + +lemma a_star_NotL: + "M \\<^isub>a* M' \ NotL .M x \\<^isub>a* NotL .M' x" +by (induct set: rtrancl) (blast intro: rtrancl.rtrancl_into_rtrancl)+ + +lemma a_star_AndRL: + "M \\<^isub>a* M'\ AndR .M .N c \\<^isub>a* AndR .M' .N c" +by (induct set: rtrancl) (blast intro: rtrancl.rtrancl_into_rtrancl)+ + +lemma a_star_AndRR: + "N \\<^isub>a* N'\ AndR .M .N c \\<^isub>a* AndR .M .N' c" +by (induct set: rtrancl) (blast intro: rtrancl.rtrancl_into_rtrancl)+ + +lemma a_star_AndR: + "\M \\<^isub>a* M'; N \\<^isub>a* N'\ \ AndR .M .N c \\<^isub>a* AndR .M' .N' c" +by (blast intro!: a_star_AndRL a_star_AndRR intro: rtrancl_trans') + +lemma a_star_AndL1: + "M \\<^isub>a* M' \ AndL1 (x).M y \\<^isub>a* AndL1 (x).M' y" +by (induct set: rtrancl) (blast intro: rtrancl.rtrancl_into_rtrancl)+ + +lemma a_star_AndL2: + "M \\<^isub>a* M' \ AndL2 (x).M y \\<^isub>a* AndL2 (x).M' y" +by (induct set: rtrancl) (blast intro: rtrancl.rtrancl_into_rtrancl)+ + +lemma a_star_OrLL: + "M \\<^isub>a* M'\ OrL (x).M (y).N z \\<^isub>a* OrL (x).M' (y).N z" +by (induct set: rtrancl) (blast intro: rtrancl.rtrancl_into_rtrancl)+ + +lemma a_star_OrLR: + "N \\<^isub>a* N'\ OrL (x).M (y).N z \\<^isub>a* OrL (x).M (y).N' z" +by (induct set: rtrancl) (blast intro: rtrancl.rtrancl_into_rtrancl)+ + +lemma a_star_OrL: + "\M \\<^isub>a* M'; N \\<^isub>a* N'\ \ OrL (x).M (y).N z \\<^isub>a* OrL (x).M' (y).N' z" +by (blast intro!: a_star_OrLL a_star_OrLR intro: rtrancl_trans') + +lemma a_star_OrR1: + "M \\<^isub>a* M' \ OrR1 .M b \\<^isub>a* OrR1 .M' b" +by (induct set: rtrancl) (blast intro: rtrancl.rtrancl_into_rtrancl)+ + +lemma a_star_OrR2: + "M \\<^isub>a* M' \ OrR2 .M b \\<^isub>a* OrR2 .M' b" +by (induct set: rtrancl) (blast intro: rtrancl.rtrancl_into_rtrancl)+ + +lemma a_star_ImpLL: + "M \\<^isub>a* M'\ ImpL .M (y).N z \\<^isub>a* ImpL .M' (y).N z" +by (induct set: rtrancl) (blast intro: rtrancl.rtrancl_into_rtrancl)+ + +lemma a_star_ImpLR: + "N \\<^isub>a* N'\ ImpL .M (y).N z \\<^isub>a* ImpL .M (y).N' z" +by (induct set: rtrancl) (blast intro: rtrancl.rtrancl_into_rtrancl)+ + +lemma a_star_ImpL: + "\M \\<^isub>a* M'; N \\<^isub>a* N'\ \ ImpL .M (y).N z \\<^isub>a* ImpL .M' (y).N' z" +by (blast intro!: a_star_ImpLL a_star_ImpLR intro: rtrancl_trans') + +lemma a_star_ImpR: + "M \\<^isub>a* M' \ ImpR (x)..M b \\<^isub>a* ImpR (x)..M' b" +by (induct set: rtrancl) (blast intro: rtrancl.rtrancl_into_rtrancl)+ + +lemmas a_star_congs = a_star_Cut a_star_NotR a_star_NotL a_star_AndR a_star_AndL1 a_star_AndL2 + a_star_OrL a_star_OrR1 a_star_OrR2 a_star_ImpL a_star_ImpR + +lemma a_star_redu_NotL_elim: + assumes a: "NotL .M x \\<^isub>a* R" + shows "\M'. R = NotL .M' x \ M \\<^isub>a* M'" +using a +apply(induct set: rtrancl) +apply(auto) +apply(drule a_redu_NotL_elim) +apply(auto) +done + +lemma a_star_redu_NotR_elim: + assumes a: "NotR (x).M a \\<^isub>a* R" + shows "\M'. R = NotR (x).M' a \ M \\<^isub>a* M'" +using a +apply(induct set: rtrancl) +apply(auto) +apply(drule a_redu_NotR_elim) +apply(auto) +done + +lemma a_star_redu_AndR_elim: + assumes a: "AndR .M .N c\\<^isub>a* R" + shows "(\M' N'. R = AndR .M' .N' c \ M \\<^isub>a* M' \ N \\<^isub>a* N')" +using a +apply(induct set: rtrancl) +apply(auto) +apply(drule a_redu_AndR_elim) +apply(auto simp add: alpha trm.inject) +done + +lemma a_star_redu_AndL1_elim: + assumes a: "AndL1 (x).M y \\<^isub>a* R" + shows "\M'. R = AndL1 (x).M' y \ M \\<^isub>a* M'" +using a +apply(induct set: rtrancl) +apply(auto) +apply(drule a_redu_AndL1_elim) +apply(auto simp add: alpha trm.inject) +done + +lemma a_star_redu_AndL2_elim: + assumes a: "AndL2 (x).M y \\<^isub>a* R" + shows "\M'. R = AndL2 (x).M' y \ M \\<^isub>a* M'" +using a +apply(induct set: rtrancl) +apply(auto) +apply(drule a_redu_AndL2_elim) +apply(auto simp add: alpha trm.inject) +done + +lemma a_star_redu_OrL_elim: + assumes a: "OrL (x).M (y).N z \\<^isub>a* R" + shows "(\M' N'. R = OrL (x).M' (y).N' z \ M \\<^isub>a* M' \ N \\<^isub>a* N')" +using a +apply(induct set: rtrancl) +apply(auto) +apply(drule a_redu_OrL_elim) +apply(auto simp add: alpha trm.inject) +done + +lemma a_star_redu_OrR1_elim: + assumes a: "OrR1 .M y \\<^isub>a* R" + shows "\M'. R = OrR1 .M' y \ M \\<^isub>a* M'" +using a +apply(induct set: rtrancl) +apply(auto) +apply(drule a_redu_OrR1_elim) +apply(auto simp add: alpha trm.inject) +done + +lemma a_star_redu_OrR2_elim: + assumes a: "OrR2 .M y \\<^isub>a* R" + shows "\M'. R = OrR2 .M' y \ M \\<^isub>a* M'" +using a +apply(induct set: rtrancl) +apply(auto) +apply(drule a_redu_OrR2_elim) +apply(auto simp add: alpha trm.inject) +done + +lemma a_star_redu_ImpR_elim: + assumes a: "ImpR (x)..M y \\<^isub>a* R" + shows "\M'. R = ImpR (x)..M' y \ M \\<^isub>a* M'" +using a +apply(induct set: rtrancl) +apply(auto) +apply(drule a_redu_ImpR_elim) +apply(auto simp add: alpha trm.inject) +done + +lemma a_star_redu_ImpL_elim: + assumes a: "ImpL .M (y).N z \\<^isub>a* R" + shows "(\M' N'. R = ImpL .M' (y).N' z \ M \\<^isub>a* M' \ N \\<^isub>a* N')" +using a +apply(induct set: rtrancl) +apply(auto) +apply(drule a_redu_ImpL_elim) +apply(auto simp add: alpha trm.inject) +done + +text {* Substitution *} + +lemma subst_not_fin1: + shows "\fin(M{x:=.P}) x" +apply(nominal_induct M avoiding: x c P rule: trm.induct) +apply(auto) +apply(drule fin_elims, simp) +apply(drule fin_elims, simp) +apply(erule fin.cases, simp_all add: trm.inject) +apply(erule fin.cases, simp_all add: trm.inject) +apply(erule fin.cases, simp_all add: trm.inject) +apply(subgoal_tac "\x'::name. x'\(trm{x:=.P},P)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_NotL) +apply(erule fin.cases, simp_all add: trm.inject) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(erule fin.cases, simp_all add: trm.inject) +apply(erule fin.cases, simp_all add: trm.inject) +apply(subgoal_tac "\x'::name. x'\(trm{x:=.P},P,name1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndL1) +apply(erule fin.cases, simp_all add: trm.inject) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(erule fin.cases, simp_all add: trm.inject) +apply(subgoal_tac "\x'::name. x'\(trm{x:=.P},P,name1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndL2) +apply(erule fin.cases, simp_all add: trm.inject) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(erule fin.cases, simp_all add: trm.inject) +apply(erule fin.cases, simp_all add: trm.inject) +apply(erule fin.cases, simp_all add: trm.inject) +apply(subgoal_tac "\x'::name. x'\(trm1{x:=.P},P,name1,trm2{x:=.P},name2)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrL) +apply(erule fin.cases, simp_all add: trm.inject) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(erule fin.cases, simp_all add: trm.inject) +apply(erule fin.cases, simp_all add: trm.inject) +apply(subgoal_tac "\x'::name. x'\(trm1{name2:=.P},P,name1,trm2{name2:=.P})") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_ImpL) +apply(erule fin.cases, simp_all add: trm.inject) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(erule fin.cases, simp_all add: trm.inject) +done + +lemma subst_not_fin2: + assumes a: "\fin M y" + shows "\fin(M{c:=(x).P}) y" +using a +apply(nominal_induct M avoiding: x c P y rule: trm.induct) +apply(auto) +apply(drule fin_elims, simp) +apply(drule fin_elims, simp) +apply(drule fin_elims, simp) +apply(subgoal_tac "\c'::coname. c'\(trm{coname:=(x).P},P)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_NotR) +apply(drule fin_elims, simp) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(drule fin_elims, simp) +apply(drule fin_elims, simp) +apply(auto)[1] +apply(drule freshn_after_substc) +apply(simp add: fin.intros) +apply(subgoal_tac "\c'::coname. c'\(trm1{coname3:=(x).P},P,coname1,trm2{coname3:=(x).P},coname2)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndR) +apply(drule fin_elims, simp) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(drule fin_elims, simp) +apply(drule fin_elims, simp) +apply(auto)[1] +apply(simp add: abs_fresh fresh_atm) +apply(drule freshn_after_substc) +apply(simp add: fin.intros abs_fresh) +apply(drule fin_elims, simp) +apply(auto)[1] +apply(simp add: abs_fresh fresh_atm) +apply(drule freshn_after_substc) +apply(simp add: fin.intros abs_fresh) +apply(subgoal_tac "\c'::coname. c'\(trm{coname2:=(x).P},P,coname1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrR1) +apply(drule fin_elims, simp) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(drule fin_elims, simp) +apply(subgoal_tac "\c'::coname. c'\(trm{coname2:=(x).P},P,coname1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrR2) +apply(drule fin_elims, simp) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(drule fin_elims, simp) +apply(drule fin_elims, simp) +apply(auto)[1] +apply(simp add: abs_fresh fresh_atm) +apply(drule freshn_after_substc) +apply(drule freshn_after_substc) +apply(simp add: fin.intros abs_fresh) +apply(subgoal_tac "\c'::coname. c'\(trm{coname2:=(x).P},P,coname1,coname2)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_ImpR) +apply(drule fin_elims, simp) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(drule fin_elims, simp) +apply(drule fin_elims, simp) +apply(auto)[1] +apply(simp add: abs_fresh fresh_atm) +apply(drule freshn_after_substc) +apply(drule freshn_after_substc) +apply(simp add: fin.intros abs_fresh) +done + +lemma subst_not_fic1: + shows "\fic (M{a:=(x).P}) a" +apply(nominal_induct M avoiding: a x P rule: trm.induct) +apply(auto) +apply(erule fic.cases, simp_all add: trm.inject) +apply(erule fic.cases, simp_all add: trm.inject) +apply(erule fic.cases, simp_all add: trm.inject) +apply(erule fic.cases, simp_all add: trm.inject) +apply(subgoal_tac "\a'::coname. a'\(trm{coname:=(x).P},P)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_NotR) +apply(erule fic.cases, simp_all add: trm.inject) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(erule fic.cases, simp_all add: trm.inject) +apply(erule fic.cases, simp_all add: trm.inject) +apply(subgoal_tac "\a'::coname. a'\(trm1{coname3:=(x).P},P,trm2{coname3:=(x).P},coname1,coname2)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndR) +apply(erule fic.cases, simp_all add: trm.inject) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(erule fic.cases, simp_all add: trm.inject) +apply(erule fic.cases, simp_all add: trm.inject) +apply(erule fic.cases, simp_all add: trm.inject) +apply(subgoal_tac "\a'::coname. a'\(trm{coname2:=(x).P},P,coname1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrR1) +apply(erule fic.cases, simp_all add: trm.inject) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(erule fic.cases, simp_all add: trm.inject) +apply(subgoal_tac "\a'::coname. a'\(trm{coname2:=(x).P},P,coname1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrR2) +apply(erule fic.cases, simp_all add: trm.inject) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(erule fic.cases, simp_all add: trm.inject) +apply(erule fic.cases, simp_all add: trm.inject) +apply(subgoal_tac "\a'::coname. a'\(trm{coname2:=(x).P},P,coname1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_ImpR) +apply(erule fic.cases, simp_all add: trm.inject) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(erule fic.cases, simp_all add: trm.inject) +apply(erule fic.cases, simp_all add: trm.inject) +done + +lemma subst_not_fic2: + assumes a: "\fic M a" + shows "\fic(M{x:=.P}) a" +using a +apply(nominal_induct M avoiding: x a P b rule: trm.induct) +apply(auto) +apply(drule fic_elims, simp) +apply(drule fic_elims, simp) +apply(drule fic_elims, simp) +apply(drule fic_elims, simp) +apply(auto)[1] +apply(drule freshc_after_substn) +apply(simp add: fic.intros) +apply(subgoal_tac "\x'::name. x'\(trm{x:=.P},P)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_NotL) +apply(drule fic_elims, simp) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(drule fic_elims, simp) +apply(drule fic_elims, simp) +apply(auto)[1] +apply(simp add: abs_fresh fresh_atm) +apply(drule freshc_after_substn) +apply(drule freshc_after_substn) +apply(simp add: fic.intros abs_fresh) +apply(subgoal_tac "\x'::name. x'\(trm{x:=.P},P,name1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndL1) +apply(drule fic_elims, simp) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(drule fic_elims, simp) +apply(subgoal_tac "\x'::name. x'\(trm{x:=.P},P,name1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndL2) +apply(drule fic_elims, simp) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(drule fic_elims, simp) +apply(drule fic_elims, simp) +apply(auto)[1] +apply(simp add: abs_fresh fresh_atm) +apply(drule freshc_after_substn) +apply(simp add: fic.intros abs_fresh) +apply(drule fic_elims, simp) +apply(auto)[1] +apply(simp add: abs_fresh fresh_atm) +apply(drule freshc_after_substn) +apply(simp add: fic.intros abs_fresh) +apply(subgoal_tac "\x'::name. x'\(trm1{x:=.P},P,name1,trm2{x:=.P},name2)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrL) +apply(drule fic_elims, simp) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(drule fic_elims, simp) +apply(drule fic_elims, simp) +apply(auto)[1] +apply(simp add: abs_fresh fresh_atm) +apply(drule freshc_after_substn) +apply(simp add: fic.intros abs_fresh) +apply(subgoal_tac "\x'::name. x'\(trm1{name2:=.P},trm2{name2:=.P},P,name1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_ImpL) +apply(drule fic_elims, simp) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(drule fic_elims, simp) +done + +text {* Reductions *} + +lemma fin_l_reduce: + assumes a: "fin M x" + and b: "M \\<^isub>l M'" + shows "fin M' x" +using b a +apply(induct) +apply(erule fin.cases) +apply(simp_all add: trm.inject) +apply(rotate_tac 3) +apply(erule fin.cases) +apply(simp_all add: trm.inject) +apply(erule fin.cases, simp_all add: trm.inject)+ +done + +lemma fin_c_reduce: + assumes a: "fin M x" + and b: "M \\<^isub>c M'" + shows "fin M' x" +using b a +apply(induct) +apply(erule fin.cases, simp_all add: trm.inject)+ +done + +lemma fin_a_reduce: + assumes a: "fin M x" + and b: "M \\<^isub>a M'" + shows "fin M' x" +using a b +apply(induct) +apply(drule ax_do_not_a_reduce) +apply(simp) +apply(drule a_redu_NotL_elim) +apply(auto) +apply(rule fin.intros) +apply(simp add: fresh_a_redu) +apply(drule a_redu_AndL1_elim) +apply(auto) +apply(rule fin.intros) +apply(force simp add: abs_fresh fresh_a_redu) +apply(drule a_redu_AndL2_elim) +apply(auto) +apply(rule fin.intros) +apply(force simp add: abs_fresh fresh_a_redu) +apply(drule a_redu_OrL_elim) +apply(auto) +apply(rule fin.intros) +apply(force simp add: abs_fresh fresh_a_redu) +apply(force simp add: abs_fresh fresh_a_redu) +apply(rule fin.intros) +apply(force simp add: abs_fresh fresh_a_redu) +apply(force simp add: abs_fresh fresh_a_redu) +apply(drule a_redu_ImpL_elim) +apply(auto) +apply(rule fin.intros) +apply(force simp add: abs_fresh fresh_a_redu) +apply(force simp add: abs_fresh fresh_a_redu) +apply(rule fin.intros) +apply(force simp add: abs_fresh fresh_a_redu) +apply(force simp add: abs_fresh fresh_a_redu) +done + +lemma fin_a_star_reduce: + assumes a: "fin M x" + and b: "M \\<^isub>a* M'" + shows "fin M' x" +using b a +apply(induct set: rtrancl) +apply(auto simp add: fin_a_reduce) +done + +lemma fic_l_reduce: + assumes a: "fic M x" + and b: "M \\<^isub>l M'" + shows "fic M' x" +using b a +apply(induct) +apply(erule fic.cases) +apply(simp_all add: trm.inject) +apply(rotate_tac 3) +apply(erule fic.cases) +apply(simp_all add: trm.inject) +apply(erule fic.cases, simp_all add: trm.inject)+ +done + +lemma fic_c_reduce: + assumes a: "fic M x" + and b: "M \\<^isub>c M'" + shows "fic M' x" +using b a +apply(induct) +apply(erule fic.cases, simp_all add: trm.inject)+ +done + +lemma fic_a_reduce: + assumes a: "fic M x" + and b: "M \\<^isub>a M'" + shows "fic M' x" +using a b +apply(induct) +apply(drule ax_do_not_a_reduce) +apply(simp) +apply(drule a_redu_NotR_elim) +apply(auto) +apply(rule fic.intros) +apply(simp add: fresh_a_redu) +apply(drule a_redu_AndR_elim) +apply(auto) +apply(rule fic.intros) +apply(force simp add: abs_fresh fresh_a_redu) +apply(force simp add: abs_fresh fresh_a_redu) +apply(rule fic.intros) +apply(force simp add: abs_fresh fresh_a_redu) +apply(force simp add: abs_fresh fresh_a_redu) +apply(drule a_redu_OrR1_elim) +apply(auto) +apply(rule fic.intros) +apply(force simp add: abs_fresh fresh_a_redu) +apply(drule a_redu_OrR2_elim) +apply(auto) +apply(rule fic.intros) +apply(force simp add: abs_fresh fresh_a_redu) +apply(drule a_redu_ImpR_elim) +apply(auto) +apply(rule fic.intros) +apply(force simp add: abs_fresh fresh_a_redu) +done + +lemma fic_a_star_reduce: + assumes a: "fic M x" + and b: "M \\<^isub>a* M'" + shows "fic M' x" +using b a +apply(induct set: rtrancl) +apply(auto simp add: fic_a_reduce) +done + +text {* substitution properties *} + +lemma subst_with_ax1: + shows "M{x:=.Ax y a} \\<^isub>a* M[x\n>y]" +proof(nominal_induct M avoiding: x a y rule: trm.induct) + case (Ax z b x a y) + show "(Ax z b){x:=.Ax y a} \\<^isub>a* (Ax z b)[x\n>y]" + proof (cases "z=x") + case True + assume eq: "z=x" + have "(Ax z b){x:=.Ax y a} = Cut .Ax y a (x).Ax x b" using eq by simp + also have "\ \\<^isub>a* (Ax x b)[x\n>y]" by blast + finally show "Ax z b{x:=.Ax y a} \\<^isub>a* (Ax z b)[x\n>y]" using eq by simp + next + case False + assume neq: "z\x" + then show "(Ax z b){x:=.Ax y a} \\<^isub>a* (Ax z b)[x\n>y]" using neq by simp + qed +next + case (Cut b M z N x a y) + have fs: "b\x" "b\a" "b\y" "b\N" "z\x" "z\a" "z\y" "z\M" by fact + have ih1: "M{x:=.Ax y a} \\<^isub>a* M[x\n>y]" by fact + have ih2: "N{x:=.Ax y a} \\<^isub>a* N[x\n>y]" by fact + show "(Cut .M (z).N){x:=.Ax y a} \\<^isub>a* (Cut .M (z).N)[x\n>y]" + proof (cases "M = Ax x b") + case True + assume eq: "M = Ax x b" + have "(Cut .M (z).N){x:=.Ax y a} = Cut .Ax y a (z).(N{x:=.Ax y a})" using fs eq by simp + also have "\ \\<^isub>a* Cut .Ax y a (z).(N[x\n>y])" using ih2 a_star_congs by blast + also have "\ = Cut .(M[x\n>y]) (z).(N[x\n>y])" using eq + by (simp add: trm.inject alpha calc_atm fresh_atm) + finally show "(Cut .M (z).N){x:=.Ax y a} \\<^isub>a* (Cut .M (z).N)[x\n>y]" using fs by simp + next + case False + assume neq: "M \ Ax x b" + have "(Cut .M (z).N){x:=.Ax y a} = Cut .(M{x:=.Ax y a}) (z).(N{x:=.Ax y a})" + using fs neq by simp + also have "\ \\<^isub>a* Cut .(M[x\n>y]) (z).(N[x\n>y])" using ih1 ih2 a_star_congs by blast + finally show "(Cut .M (z).N){x:=.Ax y a} \\<^isub>a* (Cut .M (z).N)[x\n>y]" using fs by simp + qed +next + case (NotR z M b x a y) + have fs: "z\x" "z\a" "z\y" "z\b" by fact + have ih: "M{x:=.Ax y a} \\<^isub>a* M[x\n>y]" by fact + have "(NotR (z).M b){x:=.Ax y a} = NotR (z).(M{x:=.Ax y a}) b" using fs by simp + also have "\ \\<^isub>a* NotR (z).(M[x\n>y]) b" using ih by (auto intro: a_star_congs) + finally show "(NotR (z).M b){x:=.Ax y a} \\<^isub>a* (NotR (z).M b)[x\n>y]" using fs by simp +next + case (NotL b M z x a y) + have fs: "b\x" "b\a" "b\y" "b\z" by fact + have ih: "M{x:=.Ax y a} \\<^isub>a* M[x\n>y]" by fact + show "(NotL .M z){x:=.Ax y a} \\<^isub>a* (NotL .M z)[x\n>y]" + proof(cases "z=x") + case True + assume eq: "z=x" + obtain x'::"name" where new: "x'\(Ax y a,M{x:=.Ax y a})" by (rule exists_fresh(1)[OF fs_name1]) + have "(NotL .M z){x:=.Ax y a} = + fresh_fun (\x'. Cut .Ax y a (x').NotL .(M{x:=.Ax y a}) x')" + using eq fs by simp + also have "\ = Cut .Ax y a (x').NotL .(M{x:=.Ax y a}) x'" + using new by (simp add: fresh_fun_simp_NotL fresh_prod) + also have "\ \\<^isub>a* (NotL .(M{x:=.Ax y a}) x')[x'\n>y]" + using new + apply(rule_tac a_starI) + apply(rule al_redu) + apply(rule better_LAxL_intro) + apply(auto) + done + also have "\ = NotL .(M{x:=.Ax y a}) y" using new by (simp add: nrename_fresh) + also have "\ \\<^isub>a* NotL .(M[x\n>y]) y" using ih by (auto intro: a_star_congs) + also have "\ = (NotL .M z)[x\n>y]" using eq by simp + finally show "(NotL .M z){x:=.Ax y a} \\<^isub>a* (NotL .M z)[x\n>y]" by simp + next + case False + assume neq: "z\x" + have "(NotL .M z){x:=.Ax y a} = NotL .(M{x:=.Ax y a}) z" using fs neq by simp + also have "\ \\<^isub>a* NotL .(M[x\n>y]) z" using ih by (auto intro: a_star_congs) + finally show "(NotL .M z){x:=.Ax y a} \\<^isub>a* (NotL .M z)[x\n>y]" using neq by simp + qed +next + case (AndR c M d N e x a y) + have fs: "c\x" "c\a" "c\y" "d\x" "d\a" "d\y" "d\c" "c\N" "c\e" "d\M" "d\e" by fact + have ih1: "M{x:=.Ax y a} \\<^isub>a* M[x\n>y]" by fact + have ih2: "N{x:=.Ax y a} \\<^isub>a* N[x\n>y]" by fact + have "(AndR .M .N e){x:=.Ax y a} = AndR .(M{x:=.Ax y a}) .(N{x:=.Ax y a}) e" + using fs by simp + also have "\ \\<^isub>a* AndR .(M[x\n>y]) .(N[x\n>y]) e" using ih1 ih2 by (auto intro: a_star_congs) + finally show "(AndR .M .N e){x:=.Ax y a} \\<^isub>a* (AndR .M .N e)[x\n>y]" + using fs by simp +next + case (AndL1 u M v x a y) + have fs: "u\x" "u\a" "u\y" "u\v" by fact + have ih: "M{x:=.Ax y a} \\<^isub>a* M[x\n>y]" by fact + show "(AndL1 (u).M v){x:=.Ax y a} \\<^isub>a* (AndL1 (u).M v)[x\n>y]" + proof(cases "v=x") + case True + assume eq: "v=x" + obtain v'::"name" where new: "v'\(Ax y a,M{x:=.Ax y a},u)" by (rule exists_fresh(1)[OF fs_name1]) + have "(AndL1 (u).M v){x:=.Ax y a} = + fresh_fun (\v'. Cut .Ax y a (v').AndL1 (u).(M{x:=.Ax y a}) v')" + using eq fs by simp + also have "\ = Cut .Ax y a (v').AndL1 (u).(M{x:=.Ax y a}) v'" + using new by (simp add: fresh_fun_simp_AndL1 fresh_prod) + also have "\ \\<^isub>a* (AndL1 (u).(M{x:=.Ax y a}) v')[v'\n>y]" + using new + apply(rule_tac a_starI) + apply(rule a_redu.intros) + apply(rule better_LAxL_intro) + apply(rule fin.intros) + apply(simp add: abs_fresh) + done + also have "\ = AndL1 (u).(M{x:=.Ax y a}) y" using fs new + by (auto simp add: fresh_prod fresh_atm nrename_fresh) + also have "\ \\<^isub>a* AndL1 (u).(M[x\n>y]) y" using ih by (auto intro: a_star_congs) + also have "\ = (AndL1 (u).M v)[x\n>y]" using eq fs by simp + finally show "(AndL1 (u).M v){x:=.Ax y a} \\<^isub>a* (AndL1 (u).M v)[x\n>y]" by simp + next + case False + assume neq: "v\x" + have "(AndL1 (u).M v){x:=.Ax y a} = AndL1 (u).(M{x:=.Ax y a}) v" using fs neq by simp + also have "\ \\<^isub>a* AndL1 (u).(M[x\n>y]) v" using ih by (auto intro: a_star_congs) + finally show "(AndL1 (u).M v){x:=.Ax y a} \\<^isub>a* (AndL1 (u).M v)[x\n>y]" using fs neq by simp + qed +next + case (AndL2 u M v x a y) + have fs: "u\x" "u\a" "u\y" "u\v" by fact + have ih: "M{x:=.Ax y a} \\<^isub>a* M[x\n>y]" by fact + show "(AndL2 (u).M v){x:=.Ax y a} \\<^isub>a* (AndL2 (u).M v)[x\n>y]" + proof(cases "v=x") + case True + assume eq: "v=x" + obtain v'::"name" where new: "v'\(Ax y a,M{x:=.Ax y a},u)" by (rule exists_fresh(1)[OF fs_name1]) + have "(AndL2 (u).M v){x:=.Ax y a} = + fresh_fun (\v'. Cut .Ax y a (v').AndL2 (u).(M{x:=.Ax y a}) v')" + using eq fs by simp + also have "\ = Cut .Ax y a (v').AndL2 (u).(M{x:=.Ax y a}) v'" + using new by (simp add: fresh_fun_simp_AndL2 fresh_prod) + also have "\ \\<^isub>a* (AndL2 (u).(M{x:=.Ax y a}) v')[v'\n>y]" + using new + apply(rule_tac a_starI) + apply(rule a_redu.intros) + apply(rule better_LAxL_intro) + apply(rule fin.intros) + apply(simp add: abs_fresh) + done + also have "\ = AndL2 (u).(M{x:=.Ax y a}) y" using fs new + by (auto simp add: fresh_prod fresh_atm nrename_fresh) + also have "\ \\<^isub>a* AndL2 (u).(M[x\n>y]) y" using ih by (auto intro: a_star_congs) + also have "\ = (AndL2 (u).M v)[x\n>y]" using eq fs by simp + finally show "(AndL2 (u).M v){x:=.Ax y a} \\<^isub>a* (AndL2 (u).M v)[x\n>y]" by simp + next + case False + assume neq: "v\x" + have "(AndL2 (u).M v){x:=.Ax y a} = AndL2 (u).(M{x:=.Ax y a}) v" using fs neq by simp + also have "\ \\<^isub>a* AndL2 (u).(M[x\n>y]) v" using ih by (auto intro: a_star_congs) + finally show "(AndL2 (u).M v){x:=.Ax y a} \\<^isub>a* (AndL2 (u).M v)[x\n>y]" using fs neq by simp + qed +next + case (OrR1 c M d x a y) + have fs: "c\x" "c\a" "c\y" "c\d" by fact + have ih: "M{x:=.Ax y a} \\<^isub>a* M[x\n>y]" by fact + have "(OrR1 .M d){x:=.Ax y a} = OrR1 .(M{x:=.Ax y a}) d" using fs by (simp add: fresh_atm) + also have "\ \\<^isub>a* OrR1 .(M[x\n>y]) d" using ih by (auto intro: a_star_congs) + finally show "(OrR1 .M d){x:=.Ax y a} \\<^isub>a* (OrR1 .M d)[x\n>y]" using fs by simp +next + case (OrR2 c M d x a y) + have fs: "c\x" "c\a" "c\y" "c\d" by fact + have ih: "M{x:=.Ax y a} \\<^isub>a* M[x\n>y]" by fact + have "(OrR2 .M d){x:=.Ax y a} = OrR2 .(M{x:=.Ax y a}) d" using fs by (simp add: fresh_atm) + also have "\ \\<^isub>a* OrR2 .(M[x\n>y]) d" using ih by (auto intro: a_star_congs) + finally show "(OrR2 .M d){x:=.Ax y a} \\<^isub>a* (OrR2 .M d)[x\n>y]" using fs by simp +next + case (OrL u M v N z x a y) + have fs: "u\x" "u\a" "u\y" "v\x" "v\a" "v\y" "v\u" "u\N" "u\z" "v\M" "v\z" by fact + have ih1: "M{x:=.Ax y a} \\<^isub>a* M[x\n>y]" by fact + have ih2: "N{x:=.Ax y a} \\<^isub>a* N[x\n>y]" by fact + show "(OrL (u).M (v).N z){x:=.Ax y a} \\<^isub>a* (OrL (u).M (v).N z)[x\n>y]" + proof(cases "z=x") + case True + assume eq: "z=x" + obtain z'::"name" where new: "z'\(Ax y a,M{x:=.Ax y a},N{x:=.Ax y a},u,v,y,a)" + by (rule exists_fresh(1)[OF fs_name1]) + have "(OrL (u).M (v).N z){x:=.Ax y a} = + fresh_fun (\z'. Cut .Ax y a (z').OrL (u).(M{x:=.Ax y a}) (v).(N{x:=.Ax y a}) z')" + using eq fs by simp + also have "\ = Cut .Ax y a (z').OrL (u).(M{x:=.Ax y a}) (v).(N{x:=.Ax y a}) z'" + using new by (simp add: fresh_fun_simp_OrL) + also have "\ \\<^isub>a* (OrL (u).(M{x:=.Ax y a}) (v).(N{x:=.Ax y a}) z')[z'\n>y]" + using new + apply(rule_tac a_starI) + apply(rule a_redu.intros) + apply(rule better_LAxL_intro) + apply(rule fin.intros) + apply(simp_all add: abs_fresh) + done + also have "\ = OrL (u).(M{x:=.Ax y a}) (v).(N{x:=.Ax y a}) y" using fs new + by (auto simp add: fresh_prod fresh_atm nrename_fresh subst_fresh) + also have "\ \\<^isub>a* OrL (u).(M[x\n>y]) (v).(N[x\n>y]) y" + using ih1 ih2 by (auto intro: a_star_congs) + also have "\ = (OrL (u).M (v).N z)[x\n>y]" using eq fs by simp + finally show "(OrL (u).M (v).N z){x:=.Ax y a} \\<^isub>a* (OrL (u).M (v).N z)[x\n>y]" by simp + next + case False + assume neq: "z\x" + have "(OrL (u).M (v).N z){x:=.Ax y a} = OrL (u).(M{x:=.Ax y a}) (v).(N{x:=.Ax y a}) z" + using fs neq by simp + also have "\ \\<^isub>a* OrL (u).(M[x\n>y]) (v).(N[x\n>y]) z" + using ih1 ih2 by (auto intro: a_star_congs) + finally show "(OrL (u).M (v).N z){x:=.Ax y a} \\<^isub>a* (OrL (u).M (v).N z)[x\n>y]" using fs neq by simp + qed +next + case (ImpR z c M d x a y) + have fs: "z\x" "z\a" "z\y" "c\x" "c\a" "c\y" "z\d" "c\d" by fact + have ih: "M{x:=.Ax y a} \\<^isub>a* M[x\n>y]" by fact + have "(ImpR (z)..M d){x:=.Ax y a} = ImpR (z)..(M{x:=.Ax y a}) d" using fs by simp + also have "\ \\<^isub>a* ImpR (z)..(M[x\n>y]) d" using ih by (auto intro: a_star_congs) + finally show "(ImpR (z)..M d){x:=.Ax y a} \\<^isub>a* (ImpR (z)..M d)[x\n>y]" using fs by simp +next + case (ImpL c M u N v x a y) + have fs: "c\x" "c\a" "c\y" "u\x" "u\a" "u\y" "c\N" "c\v" "u\M" "u\v" by fact + have ih1: "M{x:=.Ax y a} \\<^isub>a* M[x\n>y]" by fact + have ih2: "N{x:=.Ax y a} \\<^isub>a* N[x\n>y]" by fact + show "(ImpL .M (u).N v){x:=.Ax y a} \\<^isub>a* (ImpL .M (u).N v)[x\n>y]" + proof(cases "v=x") + case True + assume eq: "v=x" + obtain v'::"name" where new: "v'\(Ax y a,M{x:=.Ax y a},N{x:=.Ax y a},y,a,u)" + by (rule exists_fresh(1)[OF fs_name1]) + have "(ImpL .M (u).N v){x:=.Ax y a} = + fresh_fun (\v'. Cut .Ax y a (v').ImpL .(M{x:=.Ax y a}) (u).(N{x:=.Ax y a}) v')" + using eq fs by simp + also have "\ = Cut .Ax y a (v').ImpL .(M{x:=.Ax y a}) (u).(N{x:=.Ax y a}) v'" + using new by (simp add: fresh_fun_simp_ImpL) + also have "\ \\<^isub>a* (ImpL .(M{x:=.Ax y a}) (u).(N{x:=.Ax y a}) v')[v'\n>y]" + using new + apply(rule_tac a_starI) + apply(rule a_redu.intros) + apply(rule better_LAxL_intro) + apply(rule fin.intros) + apply(simp_all add: abs_fresh) + done + also have "\ = ImpL .(M{x:=.Ax y a}) (u).(N{x:=.Ax y a}) y" using fs new + by (auto simp add: fresh_prod subst_fresh fresh_atm trm.inject alpha rename_fresh) + also have "\ \\<^isub>a* ImpL .(M[x\n>y]) (u).(N[x\n>y]) y" + using ih1 ih2 by (auto intro: a_star_congs) + also have "\ = (ImpL .M (u).N v)[x\n>y]" using eq fs by simp + finally show "(ImpL .M (u).N v){x:=.Ax y a} \\<^isub>a* (ImpL .M (u).N v)[x\n>y]" using fs by simp + next + case False + assume neq: "v\x" + have "(ImpL .M (u).N v){x:=.Ax y a} = ImpL .(M{x:=.Ax y a}) (u).(N{x:=.Ax y a}) v" + using fs neq by simp + also have "\ \\<^isub>a* ImpL .(M[x\n>y]) (u).(N[x\n>y]) v" + using ih1 ih2 by (auto intro: a_star_congs) + finally show "(ImpL .M (u).N v){x:=.Ax y a} \\<^isub>a* (ImpL .M (u).N v)[x\n>y]" + using fs neq by simp + qed +qed + +lemma subst_with_ax2: + shows "M{b:=(x).Ax x a} \\<^isub>a* M[b\c>a]" +proof(nominal_induct M avoiding: b a x rule: trm.induct) + case (Ax z c b a x) + show "(Ax z c){b:=(x).Ax x a} \\<^isub>a* (Ax z c)[b\c>a]" + proof (cases "c=b") + case True + assume eq: "c=b" + have "(Ax z c){b:=(x).Ax x a} = Cut .Ax z c (x).Ax x a" using eq by simp + also have "\ \\<^isub>a* (Ax z c)[b\c>a]" using eq by blast + finally show "(Ax z c){b:=(x).Ax x a} \\<^isub>a* (Ax z c)[b\c>a]" by simp + next + case False + assume neq: "c\b" + then show "(Ax z c){b:=(x).Ax x a} \\<^isub>a* (Ax z c)[b\c>a]" by simp + qed +next + case (Cut c M z N b a x) + have fs: "c\b" "c\a" "c\x" "c\N" "z\b" "z\a" "z\x" "z\M" by fact + have ih1: "M{b:=(x).Ax x a} \\<^isub>a* M[b\c>a]" by fact + have ih2: "N{b:=(x).Ax x a} \\<^isub>a* N[b\c>a]" by fact + show "(Cut .M (z).N){b:=(x).Ax x a} \\<^isub>a* (Cut .M (z).N)[b\c>a]" + proof (cases "N = Ax z b") + case True + assume eq: "N = Ax z b" + have "(Cut .M (z).N){b:=(x).Ax x a} = Cut .(M{b:=(x).Ax x a}) (x).Ax x a" using eq fs by simp + also have "\ \\<^isub>a* Cut .(M[b\c>a]) (x).Ax x a" using ih1 a_star_congs by blast + also have "\ = Cut .(M[b\c>a]) (z).(N[b\c>a])" using eq fs + by (simp add: trm.inject alpha calc_atm fresh_atm) + finally show "(Cut .M (z).N){b:=(x).Ax x a} \\<^isub>a* (Cut .M (z).N)[b\c>a]" using fs by simp + next + case False + assume neq: "N \ Ax z b" + have "(Cut .M (z).N){b:=(x).Ax x a} = Cut .(M{b:=(x).Ax x a}) (z).(N{b:=(x).Ax x a})" + using fs neq by simp + also have "\ \\<^isub>a* Cut .(M[b\c>a]) (z).(N[b\c>a])" using ih1 ih2 a_star_congs by blast + finally show "(Cut .M (z).N){b:=(x).Ax x a} \\<^isub>a* (Cut .M (z).N)[b\c>a]" using fs by simp + qed +next + case (NotR z M c b a x) + have fs: "z\b" "z\a" "z\x" "z\c" by fact + have ih: "M{b:=(x).Ax x a} \\<^isub>a* M[b\c>a]" by fact + show "(NotR (z).M c){b:=(x).Ax x a} \\<^isub>a* (NotR (z).M c)[b\c>a]" + proof (cases "c=b") + case True + assume eq: "c=b" + obtain a'::"coname" where new: "a'\(Ax x a,M{b:=(x).Ax x a})" by (rule exists_fresh(2)[OF fs_coname1]) + have "(NotR (z).M c){b:=(x).Ax x a} = + fresh_fun (\a'. Cut .NotR (z).M{b:=(x).Ax x a} a' (x).Ax x a)" + using eq fs by simp + also have "\ = Cut .NotR (z).M{b:=(x).Ax x a} a' (x).Ax x a" + using new by (simp add: fresh_fun_simp_NotR fresh_prod) + also have "\ \\<^isub>a* (NotR (z).(M{b:=(x).Ax x a}) a')[a'\c>a]" + using new + apply(rule_tac a_starI) + apply(rule a_redu.intros) + apply(rule better_LAxR_intro) + apply(rule fic.intros) + apply(simp) + done + also have "\ = NotR (z).(M{b:=(x).Ax x a}) a" using new by (simp add: crename_fresh) + also have "\ \\<^isub>a* NotR (z).(M[b\c>a]) a" using ih by (auto intro: a_star_congs) + also have "\ = (NotR (z).M c)[b\c>a]" using eq by simp + finally show "(NotR (z).M c){b:=(x).Ax x a} \\<^isub>a* (NotR (z).M c)[b\c>a]" by simp + next + case False + assume neq: "c\b" + have "(NotR (z).M c){b:=(x).Ax x a} = NotR (z).(M{b:=(x).Ax x a}) c" using fs neq by simp + also have "\ \\<^isub>a* NotR (z).(M[b\c>a]) c" using ih by (auto intro: a_star_congs) + finally show "(NotR (z).M c){b:=(x).Ax x a} \\<^isub>a* (NotR (z).M c)[b\c>a]" using neq by simp + qed +next + case (NotL c M z b a x) + have fs: "c\b" "c\a" "c\x" "c\z" by fact + have ih: "M{b:=(x).Ax x a} \\<^isub>a* M[b\c>a]" by fact + have "(NotL .M z){b:=(x).Ax x a} = NotL .(M{b:=(x).Ax x a}) z" using fs by simp + also have "\ \\<^isub>a* NotL .(M[b\c>a]) z" using ih by (auto intro: a_star_congs) + finally show "(NotL .M z){b:=(x).Ax x a} \\<^isub>a* (NotL .M z)[b\c>a]" using fs by simp +next + case (AndR c M d N e b a x) + have fs: "c\b" "c\a" "c\x" "d\b" "d\a" "d\x" "d\c" "c\N" "c\e" "d\M" "d\e" by fact + have ih1: "M{b:=(x).Ax x a} \\<^isub>a* M[b\c>a]" by fact + have ih2: "N{b:=(x).Ax x a} \\<^isub>a* N[b\c>a]" by fact + show "(AndR .M .N e){b:=(x).Ax x a} \\<^isub>a* (AndR .M .N e)[b\c>a]" + proof(cases "e=b") + case True + assume eq: "e=b" + obtain e'::"coname" where new: "e'\(Ax x a,M{b:=(x).Ax x a},N{b:=(x).Ax x a},c,d)" + by (rule exists_fresh(2)[OF fs_coname1]) + have "(AndR .M .N e){b:=(x).Ax x a} = + fresh_fun (\e'. Cut .AndR .(M{b:=(x).Ax x a}) .(N{b:=(x).Ax x a}) e' (x).Ax x a)" + using eq fs by simp + also have "\ = Cut .AndR .(M{b:=(x).Ax x a}) .(N{b:=(x).Ax x a}) e' (x).Ax x a" + using new by (simp add: fresh_fun_simp_AndR fresh_prod) + also have "\ \\<^isub>a* (AndR .(M{b:=(x).Ax x a}) .(N{b:=(x).Ax x a}) e')[e'\c>a]" + using new + apply(rule_tac a_starI) + apply(rule a_redu.intros) + apply(rule better_LAxR_intro) + apply(rule fic.intros) + apply(simp_all add: abs_fresh) + done + also have "\ = AndR .(M{b:=(x).Ax x a}) .(N{b:=(x).Ax x a}) a" using fs new + by (auto simp add: fresh_prod fresh_atm subst_fresh crename_fresh) + also have "\ \\<^isub>a* AndR .(M[b\c>a]) .(N[b\c>a]) a" using ih1 ih2 by (auto intro: a_star_congs) + also have "\ = (AndR .M .N e)[b\c>a]" using eq fs by simp + finally show "(AndR .M .N e){b:=(x).Ax x a} \\<^isub>a* (AndR .M .N e)[b\c>a]" by simp + next + case False + assume neq: "e\b" + have "(AndR .M .N e){b:=(x).Ax x a} = AndR .(M{b:=(x).Ax x a}) .(N{b:=(x).Ax x a}) e" + using fs neq by simp + also have "\ \\<^isub>a* AndR .(M[b\c>a]) .(N[b\c>a]) e" using ih1 ih2 by (auto intro: a_star_congs) + finally show "(AndR .M .N e){b:=(x).Ax x a} \\<^isub>a* (AndR .M .N e)[b\c>a]" + using fs neq by simp + qed +next + case (AndL1 u M v b a x) + have fs: "u\b" "u\a" "u\x" "u\v" by fact + have ih: "M{b:=(x).Ax x a} \\<^isub>a* M[b\c>a]" by fact + have "(AndL1 (u).M v){b:=(x).Ax x a} = AndL1 (u).(M{b:=(x).Ax x a}) v" using fs by simp + also have "\ \\<^isub>a* AndL1 (u).(M[b\c>a]) v" using ih by (auto intro: a_star_congs) + finally show "(AndL1 (u).M v){b:=(x).Ax x a} \\<^isub>a* (AndL1 (u).M v)[b\c>a]" using fs by simp +next + case (AndL2 u M v b a x) + have fs: "u\b" "u\a" "u\x" "u\v" by fact + have ih: "M{b:=(x).Ax x a} \\<^isub>a* M[b\c>a]" by fact + have "(AndL2 (u).M v){b:=(x).Ax x a} = AndL2 (u).(M{b:=(x).Ax x a}) v" using fs by simp + also have "\ \\<^isub>a* AndL2 (u).(M[b\c>a]) v" using ih by (auto intro: a_star_congs) + finally show "(AndL2 (u).M v){b:=(x).Ax x a} \\<^isub>a* (AndL2 (u).M v)[b\c>a]" using fs by simp +next + case (OrR1 c M d b a x) + have fs: "c\b" "c\a" "c\x" "c\d" by fact + have ih: "M{b:=(x).Ax x a} \\<^isub>a* M[b\c>a]" by fact + show "(OrR1 .M d){b:=(x).Ax x a} \\<^isub>a* (OrR1 .M d)[b\c>a]" + proof(cases "d=b") + case True + assume eq: "d=b" + obtain a'::"coname" where new: "a'\(Ax x a,M{b:=(x).Ax x a},c,x,a)" + by (rule exists_fresh(2)[OF fs_coname1]) + have "(OrR1 .M d){b:=(x).Ax x a} = + fresh_fun (\a'. Cut .OrR1 .M{b:=(x).Ax x a} a' (x).Ax x a)" using fs eq by (simp) + also have "\ = Cut .OrR1 .M{b:=(x).Ax x a} a' (x).Ax x a" + using new by (simp add: fresh_fun_simp_OrR1) + also have "\ \\<^isub>a* (OrR1 .M{b:=(x).Ax x a} a')[a'\c>a]" + using new + apply(rule_tac a_starI) + apply(rule a_redu.intros) + apply(rule better_LAxR_intro) + apply(rule fic.intros) + apply(simp_all add: abs_fresh) + done + also have "\ = OrR1 .M{b:=(x).Ax x a} a" using fs new + by (auto simp add: fresh_prod fresh_atm crename_fresh subst_fresh) + also have "\ \\<^isub>a* OrR1 .(M[b\c>a]) a" using ih by (auto intro: a_star_congs) + also have "\ = (OrR1 .M d)[b\c>a]" using eq fs by simp + finally show "(OrR1 .M d){b:=(x).Ax x a} \\<^isub>a* (OrR1 .M d)[b\c>a]" by simp + next + case False + assume neq: "d\b" + have "(OrR1 .M d){b:=(x).Ax x a} = OrR1 .(M{b:=(x).Ax x a}) d" using fs neq by (simp) + also have "\ \\<^isub>a* OrR1 .(M[b\c>a]) d" using ih by (auto intro: a_star_congs) + finally show "(OrR1 .M d){b:=(x).Ax x a} \\<^isub>a* (OrR1 .M d)[b\c>a]" using fs neq by simp + qed +next + case (OrR2 c M d b a x) + have fs: "c\b" "c\a" "c\x" "c\d" by fact + have ih: "M{b:=(x).Ax x a} \\<^isub>a* M[b\c>a]" by fact + show "(OrR2 .M d){b:=(x).Ax x a} \\<^isub>a* (OrR2 .M d)[b\c>a]" + proof(cases "d=b") + case True + assume eq: "d=b" + obtain a'::"coname" where new: "a'\(Ax x a,M{b:=(x).Ax x a},c,x,a)" + by (rule exists_fresh(2)[OF fs_coname1]) + have "(OrR2 .M d){b:=(x).Ax x a} = + fresh_fun (\a'. Cut .OrR2 .M{b:=(x).Ax x a} a' (x).Ax x a)" using fs eq by (simp) + also have "\ = Cut .OrR2 .M{b:=(x).Ax x a} a' (x).Ax x a" + using new by (simp add: fresh_fun_simp_OrR2) + also have "\ \\<^isub>a* (OrR2 .M{b:=(x).Ax x a} a')[a'\c>a]" + using new + apply(rule_tac a_starI) + apply(rule a_redu.intros) + apply(rule better_LAxR_intro) + apply(rule fic.intros) + apply(simp_all add: abs_fresh) + done + also have "\ = OrR2 .M{b:=(x).Ax x a} a" using fs new + by (auto simp add: fresh_prod fresh_atm crename_fresh subst_fresh) + also have "\ \\<^isub>a* OrR2 .(M[b\c>a]) a" using ih by (auto intro: a_star_congs) + also have "\ = (OrR2 .M d)[b\c>a]" using eq fs by simp + finally show "(OrR2 .M d){b:=(x).Ax x a} \\<^isub>a* (OrR2 .M d)[b\c>a]" by simp + next + case False + assume neq: "d\b" + have "(OrR2 .M d){b:=(x).Ax x a} = OrR2 .(M{b:=(x).Ax x a}) d" using fs neq by (simp) + also have "\ \\<^isub>a* OrR2 .(M[b\c>a]) d" using ih by (auto intro: a_star_congs) + finally show "(OrR2 .M d){b:=(x).Ax x a} \\<^isub>a* (OrR2 .M d)[b\c>a]" using fs neq by simp + qed +next + case (OrL u M v N z b a x) + have fs: "u\b" "u\a" "u\x" "v\b" "v\a" "v\x" "v\u" "u\N" "u\z" "v\M" "v\z" by fact + have ih1: "M{b:=(x).Ax x a} \\<^isub>a* M[b\c>a]" by fact + have ih2: "N{b:=(x).Ax x a} \\<^isub>a* N[b\c>a]" by fact + have "(OrL (u).M (v).N z){b:=(x).Ax x a} = OrL (u).(M{b:=(x).Ax x a}) (v).(N{b:=(x).Ax x a}) z" + using fs by simp + also have "\ \\<^isub>a* OrL (u).(M[b\c>a]) (v).(N[b\c>a]) z" using ih1 ih2 by (auto intro: a_star_congs) + finally show "(OrL (u).M (v).N z){b:=(x).Ax x a} \\<^isub>a* (OrL (u).M (v).N z)[b\c>a]" using fs by simp +next + case (ImpR z c M d b a x) + have fs: "z\b" "z\a" "z\x" "c\b" "c\a" "c\x" "z\d" "c\d" by fact + have ih: "M{b:=(x).Ax x a} \\<^isub>a* M[b\c>a]" by fact + show "(ImpR (z)..M d){b:=(x).Ax x a} \\<^isub>a* (ImpR (z)..M d)[b\c>a]" + proof(cases "b=d") + case True + assume eq: "b=d" + obtain a'::"coname" where new: "a'\(Ax x a,M{b:=(x).Ax x a},x,a,c)" + by (rule exists_fresh(2)[OF fs_coname1]) + have "(ImpR (z)..M d){b:=(x).Ax x a} = + fresh_fun (\a'. Cut .ImpR z..M{b:=(x).Ax x a} a' (x).Ax x a)" using fs eq by simp + also have "\ = Cut .ImpR z..M{b:=(x).Ax x a} a' (x).Ax x a" + using new by (simp add: fresh_fun_simp_ImpR) + also have "\ \\<^isub>a* (ImpR z..M{b:=(x).Ax x a} a')[a'\c>a]" + using new + apply(rule_tac a_starI) + apply(rule a_redu.intros) + apply(rule better_LAxR_intro) + apply(rule fic.intros) + apply(simp_all add: abs_fresh) + done + also have "\ = ImpR z..M{b:=(x).Ax x a} a" using fs new + by (auto simp add: fresh_prod crename_fresh subst_fresh fresh_atm) + also have "\ \\<^isub>a* ImpR z..(M[b\c>a]) a" using ih by (auto intro: a_star_congs) + also have "\ = (ImpR z..M b)[b\c>a]" using eq fs by simp + finally show "(ImpR (z)..M d){b:=(x).Ax x a} \\<^isub>a* (ImpR (z)..M d)[b\c>a]" using eq by simp + next + case False + assume neq: "b\d" + have "(ImpR (z)..M d){b:=(x).Ax x a} = ImpR (z)..(M{b:=(x).Ax x a}) d" using fs neq by simp + also have "\ \\<^isub>a* ImpR (z)..(M[b\c>a]) d" using ih by (auto intro: a_star_congs) + finally show "(ImpR (z)..M d){b:=(x).Ax x a} \\<^isub>a* (ImpR (z)..M d)[b\c>a]" using neq fs by simp + qed +next + case (ImpL c M u N v b a x) + have fs: "c\b" "c\a" "c\x" "u\b" "u\a" "u\x" "c\N" "c\v" "u\M" "u\v" by fact + have ih1: "M{b:=(x).Ax x a} \\<^isub>a* M[b\c>a]" by fact + have ih2: "N{b:=(x).Ax x a} \\<^isub>a* N[b\c>a]" by fact + have "(ImpL .M (u).N v){b:=(x).Ax x a} = ImpL .(M{b:=(x).Ax x a}) (u).(N{b:=(x).Ax x a}) v" + using fs by simp + also have "\ \\<^isub>a* ImpL .(M[b\c>a]) (u).(N[b\c>a]) v" + using ih1 ih2 by (auto intro: a_star_congs) + finally show "(ImpL .M (u).N v){b:=(x).Ax x a} \\<^isub>a* (ImpL .M (u).N v)[b\c>a]" + using fs by simp +qed + +text {* substitution lemmas *} + +lemma not_Ax1: + shows "\(b\M) \ M{b:=(y).Q} \ Ax x a" +apply(nominal_induct M avoiding: b y Q x a rule: trm.induct) +apply(auto simp add: fresh_atm abs_fresh abs_supp fin_supp) +apply(subgoal_tac "\x'::coname. x'\(trm{coname:=(y).Q},Q)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_NotR abs_fresh fresh_atm) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(subgoal_tac "\x'::coname. x'\(trm{coname:=(y).Q},Q)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_NotR abs_fresh fresh_atm) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(subgoal_tac "\x'::coname. x'\(trm1{coname3:=(y).Q},Q,trm2{coname3:=(y).Q},coname1,coname2)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndR abs_fresh fresh_atm) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(subgoal_tac "\x'::coname. x'\(trm1{coname3:=(y).Q},Q,trm2{coname3:=(y).Q},coname1,coname2)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndR abs_fresh fresh_atm) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(subgoal_tac "\x'::coname. x'\(trm1{coname3:=(y).Q},Q,trm2{coname3:=(y).Q},coname1,coname2)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndR abs_fresh fresh_atm) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(subgoal_tac "\x'::coname. x'\(trm{coname2:=(y).Q},Q,coname1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrR1 abs_fresh fresh_atm) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(subgoal_tac "\x'::coname. x'\(trm{coname2:=(y).Q},Q,coname1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrR1 abs_fresh fresh_atm) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(subgoal_tac "\x'::coname. x'\(trm{coname2:=(y).Q},Q,coname1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrR2 abs_fresh fresh_atm) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(subgoal_tac "\x'::coname. x'\(trm{coname2:=(y).Q},Q,coname1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrR2 abs_fresh fresh_atm) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(subgoal_tac "\x'::coname. x'\(trm{coname2:=(y).Q},Q,coname1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_ImpR abs_fresh abs_supp fin_supp fresh_atm) +apply(rule exists_fresh'(2)[OF fs_coname1]) +apply(subgoal_tac "\x'::coname. x'\(trm{coname2:=(y).Q},Q,coname1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_ImpR abs_fresh abs_supp fin_supp fresh_atm) +apply(rule exists_fresh'(2)[OF fs_coname1]) +done + +lemma not_Ax2: + shows "\(x\M) \ M{x:=.Q} \ Ax y a" +apply(nominal_induct M avoiding: b y Q x a rule: trm.induct) +apply(auto simp add: fresh_atm abs_fresh abs_supp fin_supp) +apply(subgoal_tac "\x'::name. x'\(trm{x:=.Q},Q)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_NotL abs_fresh fresh_atm) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(subgoal_tac "\x'::name. x'\(trm{x:=.Q},Q)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_NotL abs_fresh fresh_atm) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(subgoal_tac "\x'::name. x'\(trm{x:=.Q},Q,name1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndL1 abs_fresh fresh_atm) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(subgoal_tac "\x'::name. x'\(trm{x:=.Q},Q,name1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndL1 abs_fresh fresh_atm) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(subgoal_tac "\x'::name. x'\(trm{x:=.Q},Q,name1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndL2 abs_fresh fresh_atm) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(subgoal_tac "\x'::name. x'\(trm{x:=.Q},Q,name1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_AndL2 abs_fresh fresh_atm) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(subgoal_tac "\x'::name. x'\(trm1{x:=.Q},Q,trm2{x:=.Q},name1,name2)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrL abs_fresh fresh_atm) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(subgoal_tac "\x'::name. x'\(trm1{x:=.Q},Q,trm2{x:=.Q},name1,name2)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrL abs_fresh fresh_atm) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(subgoal_tac "\x'::name. x'\(trm1{x:=.Q},Q,trm2{x:=.Q},name1,name2)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_OrL abs_fresh fresh_atm) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(subgoal_tac "\x'::name. x'\(trm1{name2:=.Q},Q,trm2{name2:=.Q},name1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_ImpL abs_fresh fresh_atm) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(subgoal_tac "\x'::name. x'\(trm1{name2:=.Q},Q,trm2{name2:=.Q},name1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_ImpL abs_fresh fresh_atm) +apply(rule exists_fresh'(1)[OF fs_name1]) +apply(subgoal_tac "\x'::name. x'\(trm1{name2:=.Q},Q,trm2{name2:=.Q},name1)") +apply(erule exE) +apply(simp add: fresh_prod) +apply(erule conjE)+ +apply(simp add: fresh_fun_simp_ImpL abs_fresh fresh_atm) +apply(rule exists_fresh'(1)[OF fs_name1]) +done + +lemma interesting_subst1: + assumes a: "x\y" "x\P" "y\P" + shows "N{y:=.P}{x:=.P} = N{x:=.Ax y c}{y:=.P}" +using a +proof(nominal_induct N avoiding: x y c P rule: trm.induct) + case Ax + then show ?case + by (auto simp add: abs_fresh fresh_atm forget trm.inject) +next + case (Cut d M u M' x' y' c P) + from prems show ?case + apply(simp) + apply(auto) + apply(rule trans) + apply(rule better_Cut_substn) + apply(simp add: abs_fresh) + apply(simp) + apply(simp) + apply(auto) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substn) + apply(simp add: abs_fresh) + apply(simp) + apply(simp) + apply(simp add: trm.inject alpha forget) + apply(rule trans) + apply(rule better_Cut_substn) + apply(simp add: abs_fresh) + apply(simp) + apply(simp) + apply(auto) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substn) + apply(simp add: abs_fresh) + apply(simp) + apply(simp) + apply(rule impI) + apply(simp add: trm.inject alpha forget) + apply(simp add: trm.inject alpha forget) + apply(rule trans) + apply(rule better_Cut_substn) + apply(simp add: abs_fresh) + apply(simp) + apply(simp) + apply(auto) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substn) + apply(simp add: abs_fresh) + apply(simp) + apply(simp) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substn) + apply(simp add: abs_fresh) + apply(simp) + apply(simp) + apply(simp add: trm.inject alpha forget) + apply(rule trans) + apply(rule better_Cut_substn) + apply(simp add: abs_fresh) + apply(simp) + apply(simp) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substn) + apply(simp add: abs_fresh) + apply(simp) + apply(simp) + apply(rule trans) + apply(rule better_Cut_substn) + apply(simp add: abs_fresh) + apply(simp) + apply(simp) + apply(auto) + apply(case_tac "y'\M") + apply(simp add: forget) + apply(simp add: not_Ax2) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substn) + apply(simp add: abs_fresh) + apply(simp) + apply(simp) + apply(auto) + apply(case_tac "x'\M") + apply(simp add: forget) + apply(simp add: not_Ax2) + done +next + case NotR + then show ?case + by (auto simp add: abs_fresh fresh_atm forget) +next + case (NotL d M u) + then show ?case + apply (auto simp add: abs_fresh fresh_atm forget) + apply(subgoal_tac "\x'::name. x'\(P,M{y:=.P},M{x:=.Ax y c}{y:=.P},y,x)") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_NotL) + apply(rule trans) + apply(rule better_Cut_substn) + apply(simp add: abs_fresh) + apply(simp) + apply(simp) + apply(auto simp add: fresh_atm) + apply(simp add: trm.inject alpha forget) + apply(rule exists_fresh'(1)[OF fs_name1]) + apply(subgoal_tac "\x'::name. x'\(P,M{x:=.Ax y c},M{x:=.Ax y c}{y:=.P},Ax y c,y,x)") + apply(erule exE, simp only: fresh_prod) + apply(erule conjE)+ + apply(simp only: fresh_fun_simp_NotL) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substn) + apply(simp add: abs_fresh) + apply(simp) + apply(simp) + apply(simp add: trm.inject alpha forget subst_fresh) + apply(rule trans) + apply(rule substn.simps) + apply(simp add: abs_fresh fresh_prod fresh_atm) + apply(simp add: fresh_atm) + apply(rule exists_fresh'(1)[OF fs_name1]) + done +next + case (AndR d1 M d2 M' d3) + then show ?case + by (auto simp add: abs_fresh fresh_atm forget trm.inject subst_fresh) +next + case (AndL1 u M d) + then show ?case + apply(auto simp add: abs_fresh fresh_atm forget trm.inject subst_fresh) + apply(subgoal_tac "\x'::name. x'\(P,M{y:=.P},M{x:=.Ax y c}{y:=.P},u,y,x)") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_AndL1) + apply(rule trans) + apply(rule better_Cut_substn) + apply(simp add: abs_fresh) + apply(simp) + apply(simp) + apply(auto simp add: fresh_atm) + apply(simp add: trm.inject alpha forget) + apply(rule exists_fresh'(1)[OF fs_name1]) + apply(subgoal_tac "\x'::name. x'\(P,Ax y c,M{x:=.Ax y c},M{x:=.Ax y c}{y:=.P},u,y,x)") + apply(erule exE, simp only: fresh_prod) + apply(erule conjE)+ + apply(simp only: fresh_fun_simp_AndL1) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substn) + apply(simp add: abs_fresh) + apply(simp) + apply(simp) + apply(auto simp add: fresh_atm) + apply(rule exists_fresh'(1)[OF fs_name1]) + done +next + case (AndL2 u M d) + then show ?case + apply(auto simp add: abs_fresh fresh_atm forget trm.inject subst_fresh) + apply(subgoal_tac "\x'::name. x'\(P,M{y:=.P},M{x:=.Ax y c}{y:=.P},u,y,x)") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_AndL2) + apply(rule trans) + apply(rule better_Cut_substn) + apply(simp add: abs_fresh) + apply(simp) + apply(simp) + apply(auto simp add: fresh_atm) + apply(simp add: trm.inject alpha forget) + apply(rule exists_fresh'(1)[OF fs_name1]) + apply(subgoal_tac "\x'::name. x'\(P,Ax y c,M{x:=.Ax y c},M{x:=.Ax y c}{y:=.P},u,y,x)") + apply(erule exE, simp only: fresh_prod) + apply(erule conjE)+ + apply(simp only: fresh_fun_simp_AndL2) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substn) + apply(simp add: abs_fresh) + apply(simp) + apply(simp) + apply(auto simp add: fresh_atm) + apply(rule exists_fresh'(1)[OF fs_name1]) + done +next + case OrR1 + then show ?case + by (auto simp add: abs_fresh fresh_atm forget trm.inject subst_fresh) +next + case OrR2 + then show ?case + by (auto simp add: abs_fresh fresh_atm forget trm.inject subst_fresh) +next + case (OrL x1 M x2 M' x3) + then show ?case + apply(auto simp add: abs_fresh fresh_atm forget trm.inject subst_fresh) + apply(subgoal_tac "\x'::name. x'\(P,M{y:=.P},M{x:=.Ax y c}{y:=.P}, + M'{y:=.P},M'{x:=.Ax y c}{y:=.P},x1,x2,x3,y,x)") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_OrL) + apply(rule trans) + apply(rule better_Cut_substn) + apply(simp add: abs_fresh) + apply(simp) + apply(simp) + apply(auto simp add: fresh_atm) + apply(simp add: trm.inject alpha forget) + apply(rule trans) + apply(rule substn.simps) + apply(simp add: abs_fresh subst_fresh fresh_atm) + apply(simp add: abs_fresh subst_fresh fresh_atm) + apply(force) + apply(simp) + apply(rule exists_fresh'(1)[OF fs_name1]) + apply(subgoal_tac "\x'::name. x'\(P,Ax y c,M{x:=.Ax y c},M{x:=.Ax y c}{y:=.P}, + M'{x:=.Ax y c},M'{x:=.Ax y c}{y:=.P},x1,x2,x3,y,x)") + apply(erule exE, simp only: fresh_prod) + apply(erule conjE)+ + apply(simp only: fresh_fun_simp_OrL) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substn) + apply(simp add: abs_fresh) + apply(simp) + apply(simp) + apply(simp add: trm.inject alpha) + apply(rule trans) + apply(rule substn.simps) + apply(simp add: abs_fresh subst_fresh fresh_atm) + apply(simp add: abs_fresh subst_fresh fresh_atm) + apply(force) + apply(simp) + apply(auto simp add: fresh_atm) + apply(rule exists_fresh'(1)[OF fs_name1]) + done +next + case ImpR + then show ?case + by (auto simp add: abs_fresh fresh_atm forget trm.inject subst_fresh) +next + case (ImpL a M x1 M' x2) + then show ?case + apply(auto simp add: abs_fresh fresh_atm forget trm.inject subst_fresh) + apply(subgoal_tac "\x'::name. x'\(P,M{x2:=.P},M{x:=.Ax x2 c}{x2:=.P}, + M'{x2:=.P},M'{x:=.Ax x2 c}{x2:=.P},x1,y,x)") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_ImpL) + apply(rule trans) + apply(rule better_Cut_substn) + apply(simp add: abs_fresh) + apply(simp) + apply(simp) + apply(auto simp add: fresh_atm) + apply(simp add: trm.inject alpha forget) + apply(rule trans) + apply(rule substn.simps) + apply(simp add: abs_fresh subst_fresh fresh_atm) + apply(simp add: abs_fresh subst_fresh fresh_atm) + apply(force) + apply(rule exists_fresh'(1)[OF fs_name1]) + apply(subgoal_tac "\x'::name. x'\(P,Ax y c,M{x2:=.Ax y c},M{x2:=.Ax y c}{y:=.P}, + M'{x2:=.Ax y c},M'{x2:=.Ax y c}{y:=.P},x1,x2,x3,y,x)") + apply(erule exE, simp only: fresh_prod) + apply(erule conjE)+ + apply(simp only: fresh_fun_simp_ImpL) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substn) + apply(simp add: abs_fresh) + apply(simp) + apply(simp) + apply(simp add: trm.inject alpha) + apply(rule trans) + apply(rule substn.simps) + apply(simp add: abs_fresh subst_fresh fresh_atm) + apply(simp add: abs_fresh subst_fresh fresh_atm) + apply(simp) + apply(auto simp add: fresh_atm) + apply(rule exists_fresh'(1)[OF fs_name1]) + done +qed + +lemma interesting_subst1': + assumes a: "x\y" "x\P" "y\P" + shows "N{y:=.P}{x:=.P} = N{x:=.Ax y a}{y:=.P}" +proof - + show ?thesis + proof (cases "c=a") + case True then show ?thesis using a by (simp add: interesting_subst1) + next + case False then show ?thesis using a + apply - + apply(subgoal_tac "N{x:=.Ax y a} = N{x:=.([(c,a)]\Ax y a)}") + apply(simp add: interesting_subst1 calc_atm) + apply(rule subst_rename) + apply(simp add: fresh_prod fresh_atm) + done + qed +qed + +lemma interesting_subst2: + assumes a: "a\b" "a\P" "b\P" + shows "N{a:=(y).P}{b:=(y).P} = N{b:=(y).Ax y a}{a:=(y).P}" +using a +proof(nominal_induct N avoiding: a b y P rule: trm.induct) + case Ax + then show ?case + by (auto simp add: abs_fresh fresh_atm forget trm.inject) +next + case (Cut d M u M' x' y' c P) + from prems show ?case + apply(simp) + apply(auto simp add: trm.inject) + apply(rule trans) + apply(rule better_Cut_substc) + apply(simp) + apply(simp add: abs_fresh) + apply(simp add: forget) + apply(auto) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substc) + apply(simp add: abs_fresh) + apply(simp add: abs_fresh) + apply(simp) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substc) + apply(simp add: abs_fresh) + apply(simp add: abs_fresh) + apply(simp) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substc) + apply(simp add: abs_fresh) + apply(simp add: abs_fresh) + apply(simp) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substc) + apply(simp add: abs_fresh) + apply(simp add: abs_fresh) + apply(simp) + apply(rule trans) + apply(rule better_Cut_substc) + apply(simp add: abs_fresh) + apply(simp add: abs_fresh) + apply(auto)[1] + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substc) + apply(simp add: abs_fresh) + apply(simp add: abs_fresh) + apply(simp) + apply(rule impI) + apply(simp add: fresh_atm trm.inject alpha forget) + apply(case_tac "x'\M'") + apply(simp add: forget) + apply(simp add: not_Ax1) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substc) + apply(simp add: abs_fresh) + apply(simp add: abs_fresh) + apply(auto) + apply(case_tac "y'\M'") + apply(simp add: forget) + apply(simp add: not_Ax1) + done +next + case NotL + then show ?case + by (auto simp add: abs_fresh fresh_atm forget) +next + case (NotR u M d) + then show ?case + apply (auto simp add: abs_fresh fresh_atm forget) + apply(subgoal_tac "\a'::coname. a'\(b,P,M{d:=(y).P},M{b:=(y).Ax y d}{d:=(y).P},u,y)") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_NotR) + apply(rule trans) + apply(rule better_Cut_substc) + apply(simp add: abs_fresh) + apply(simp add: abs_fresh) + apply(simp) + apply(auto simp add: fresh_atm) + apply(simp add: trm.inject alpha forget) + apply(rule exists_fresh'(2)[OF fs_coname1]) + apply(subgoal_tac "\a'::coname. a'\(P,M{d:=(y).Ax y a},M{d:=(y).Ax y a}{a:=(y).P},Ax y a,y,d)") + apply(erule exE, simp only: fresh_prod) + apply(erule conjE)+ + apply(simp only: fresh_fun_simp_NotR) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substc) + apply(simp add: abs_fresh) + apply(simp add: abs_fresh) + apply(simp) + apply(simp add: trm.inject alpha forget subst_fresh) + apply(rule trans) + apply(rule substc.simps) + apply(simp add: abs_fresh fresh_prod fresh_atm) + apply(simp add: fresh_atm) + apply(rule exists_fresh'(2)[OF fs_coname1]) + done +next + case (AndR d1 M d2 M' d3) + then show ?case + apply(auto simp add: abs_fresh fresh_atm forget trm.inject subst_fresh) + apply(subgoal_tac "\a'::coname. a'\(P,M{d3:=(y).P},M{b:=(y).Ax y d3}{d3:=(y).P}, + M'{d3:=(y).P},M'{b:=(y).Ax y d3}{d3:=(y).P},d1,d2,d3,b,y)") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_AndR) + apply(rule trans) + apply(rule better_Cut_substc) + apply(simp add: abs_fresh fresh_atm) + apply(simp add: abs_fresh fresh_atm) + apply(simp) + apply(auto simp add: fresh_atm) + apply(simp add: trm.inject alpha forget) + apply(rule trans) + apply(rule substc.simps) + apply(simp add: abs_fresh subst_fresh fresh_atm) + apply(simp add: abs_fresh subst_fresh fresh_atm) + apply(force) + apply(simp) + apply(rule exists_fresh'(2)[OF fs_coname1]) + apply(subgoal_tac "\a'::coname. a'\(P,Ax y a,M{d3:=(y).Ax y a},M{d3:=(y).Ax y a}{a:=(y).P}, + M'{d3:=(y).Ax y a},M'{d3:=(y).Ax y a}{a:=(y).P},d1,d2,d3,y,b)") + apply(erule exE, simp only: fresh_prod) + apply(erule conjE)+ + apply(simp only: fresh_fun_simp_AndR) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substc) + apply(simp add: abs_fresh) + apply(simp add: abs_fresh) + apply(simp) + apply(simp add: trm.inject alpha) + apply(rule trans) + apply(rule substc.simps) + apply(simp add: abs_fresh subst_fresh fresh_atm) + apply(simp add: abs_fresh subst_fresh fresh_atm) + apply(force) + apply(simp) + apply(auto simp add: fresh_atm) + apply(rule exists_fresh'(2)[OF fs_coname1]) + done +next + case (AndL1 u M d) + then show ?case + by (auto simp add: abs_fresh fresh_atm forget trm.inject subst_fresh) +next + case (AndL2 u M d) + then show ?case + by (auto simp add: abs_fresh fresh_atm forget trm.inject subst_fresh) +next + case (OrR1 d M e) + then show ?case + apply (auto simp add: abs_fresh fresh_atm forget) + apply(subgoal_tac "\a'::coname. a'\(b,P,M{e:=(y).P},M{b:=(y).Ax y e}{e:=(y).P},d,e)") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_OrR1) + apply(rule trans) + apply(rule better_Cut_substc) + apply(simp add: abs_fresh) + apply(simp add: abs_fresh) + apply(simp) + apply(auto simp add: fresh_atm) + apply(simp add: trm.inject alpha forget) + apply(rule exists_fresh'(2)[OF fs_coname1]) + apply(subgoal_tac "\a'::coname. a'\(b,P,Ax y a,M{e:=(y).Ax y a},M{e:=(y).Ax y a}{a:=(y).P},d,e)") + apply(erule exE, simp only: fresh_prod) + apply(erule conjE)+ + apply(simp only: fresh_fun_simp_OrR1) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substc) + apply(simp add: abs_fresh) + apply(simp add: abs_fresh) + apply(simp) + apply(simp add: trm.inject alpha forget subst_fresh) + apply(rule trans) + apply(rule substc.simps) + apply(simp add: abs_fresh fresh_prod fresh_atm) + apply(simp add: fresh_atm) + apply(rule exists_fresh'(2)[OF fs_coname1]) + done +next + case (OrR2 d M e) + then show ?case + apply (auto simp add: abs_fresh fresh_atm forget) + apply(subgoal_tac "\a'::coname. a'\(b,P,M{e:=(y).P},M{b:=(y).Ax y e}{e:=(y).P},d,e)") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_OrR2) + apply(rule trans) + apply(rule better_Cut_substc) + apply(simp add: abs_fresh) + apply(simp add: abs_fresh) + apply(simp) + apply(auto simp add: fresh_atm) + apply(simp add: trm.inject alpha forget) + apply(rule exists_fresh'(2)[OF fs_coname1]) + apply(subgoal_tac "\a'::coname. a'\(b,P,Ax y a,M{e:=(y).Ax y a},M{e:=(y).Ax y a}{a:=(y).P},d,e)") + apply(erule exE, simp only: fresh_prod) + apply(erule conjE)+ + apply(simp only: fresh_fun_simp_OrR2) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substc) + apply(simp add: abs_fresh) + apply(simp add: abs_fresh) + apply(simp) + apply(simp add: trm.inject alpha forget subst_fresh) + apply(rule trans) + apply(rule substc.simps) + apply(simp add: abs_fresh fresh_prod fresh_atm) + apply(simp add: fresh_atm) + apply(rule exists_fresh'(2)[OF fs_coname1]) + done +next + case (OrL x1 M x2 M' x3) + then show ?case + by(auto simp add: abs_fresh fresh_atm forget trm.inject subst_fresh) +next + case ImpL + then show ?case + by (auto simp add: abs_fresh fresh_atm forget trm.inject subst_fresh) +next + case (ImpR u e M d) + then show ?case + apply(auto simp add: abs_fresh fresh_atm forget trm.inject subst_fresh) + apply(subgoal_tac "\a'::coname. a'\(b,e,d,P,M{d:=(y).P},M{b:=(y).Ax y d}{d:=(y).P})") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_ImpR) + apply(rule trans) + apply(rule better_Cut_substc) + apply(simp add: abs_fresh) + apply(simp add: abs_fresh) + apply(simp) + apply(auto simp add: fresh_atm) + apply(simp add: trm.inject alpha forget) + apply(rule exists_fresh'(2)[OF fs_coname1]) + apply(subgoal_tac "\a'::coname. a'\(e,d,P,Ax y a,M{d:=(y).Ax y a},M{d:=(y).Ax y a}{a:=(y).P})") + apply(erule exE, simp only: fresh_prod) + apply(erule conjE)+ + apply(simp only: fresh_fun_simp_ImpR) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substc) + apply(simp add: abs_fresh) + apply(simp add: abs_fresh) + apply(simp) + apply(simp add: trm.inject alpha) + apply(rule trans) + apply(rule substc.simps) + apply(simp add: abs_fresh subst_fresh fresh_atm) + apply(simp add: abs_fresh subst_fresh fresh_atm) + apply(simp) + apply(auto simp add: fresh_atm) + apply(rule exists_fresh'(2)[OF fs_coname1]) + done +qed + +lemma interesting_subst2': + assumes a: "a\b" "a\P" "b\P" + shows "N{a:=(y).P}{b:=(y).P} = N{b:=(z).Ax z a}{a:=(y).P}" +proof - + show ?thesis + proof (cases "z=y") + case True then show ?thesis using a by (simp add: interesting_subst2) + next + case False then show ?thesis using a + apply - + apply(subgoal_tac "N{b:=(z).Ax z a} = N{b:=(y).([(y,z)]\Ax z a)}") + apply(simp add: interesting_subst2 calc_atm) + apply(rule subst_rename) + apply(simp add: fresh_prod fresh_atm) + done + qed +qed + +lemma subst_subst1: + assumes a: "a\(Q,b)" "x\(y,P,Q)" "b\Q" "y\P" + shows "M{x:=.P}{b:=(y).Q} = M{b:=(y).Q}{x:=.(P{b:=(y).Q})}" +using a +proof(nominal_induct M avoiding: x a P b y Q rule: trm.induct) + case (Ax z c) + have fs: "a\(Q,b)" "x\(y,P,Q)" "b\Q" "y\P" by fact + { assume asm: "z=x \ c=b" + have "(Ax x b){x:=.P}{b:=(y).Q} = (Cut .P (x).Ax x b){b:=(y).Q}" using fs by simp + also have "\ = Cut .(P{b:=(y).Q}) (y).Q" + using fs by (simp_all add: fresh_prod fresh_atm) + also have "\ = Cut .(P{b:=(y).Q}) (y).(Q{x:=.(P{b:=(y).Q})})" using fs by (simp add: forget) + also have "\ = (Cut .Ax x b (y).Q){x:=.(P{b:=(y).Q})}" + using fs asm by (auto simp add: fresh_prod fresh_atm subst_fresh) + also have "\ = (Ax x b){b:=(y).Q}{x:=.(P{b:=(y).Q})}" using fs by simp + finally have "(Ax z c){x:=.P}{b:=(y).Q} = (Ax z c){b:=(y).Q}{x:=.(P{b:=(y).Q})}" + using asm by simp + } + moreover + { assume asm: "z\x \ c=b" + have "(Ax z c){x:=.P}{b:=(y).Q} = (Ax z c){b:=(y).Q}" using asm by simp + also have "\ = Cut .Ax z c (y).Q" using fs asm by simp + also have "\ = Cut .(Ax z c{x:=.(P{b:=(y).Q})}) (y).(Q{x:=.(P{b:=(y).Q})})" + using fs asm by (simp add: forget) + also have "\ = (Cut .Ax z c (y).Q){x:=.(P{b:=(y).Q})}" using asm fs + by (auto simp add: trm.inject subst_fresh fresh_prod fresh_atm abs_fresh) + also have "\ = (Ax z c){b:=(y).Q}{x:=.(P{b:=(y).Q})}" using asm fs by simp + finally have "(Ax z c){x:=.P}{b:=(y).Q} = (Ax z c){b:=(y).Q}{x:=.(P{b:=(y).Q})}" by simp + } + moreover + { assume asm: "z=x \ c\b" + have "(Ax z c){x:=.P}{b:=(y).Q} = (Cut .P (x).Ax z c){b:=(y).Q}" using fs asm by simp + also have "\ = Cut .(P{b:=(y).Q}) (x).Ax z c" using fs asm by (auto simp add: trm.inject abs_fresh) + also have "\ = (Ax z c){x:=.(P{b:=(y).Q})}" using fs asm by simp + also have "\ = (Ax z c){b:=(y).Q}{x:=.(P{b:=(y).Q})}" using asm by auto + finally have "(Ax z c){x:=.P}{b:=(y).Q} = (Ax z c){b:=(y).Q}{x:=.(P{b:=(y).Q})}" by simp + } + moreover + { assume asm: "z\x \ c\b" + have "(Ax z c){x:=.P}{b:=(y).Q} = (Ax z c){b:=(y).Q}{x:=.(P{b:=(y).Q})}" using asm by auto + } + ultimately show ?case by blast +next + case (Cut c M z N) + { assume asm: "M = Ax x c \ N = Ax z b" + have "(Cut .M (z).N){x:=.P}{b:=(y).Q} = (Cut .P (z).(N{x:=.P})){b:=(y).Q}" + using asm prems by simp + also have "\ = (Cut .P (z).N){b:=(y).Q}" using asm prems by (simp add: fresh_atm) + also have "\ = (Cut .(P{b:=(y).Q}) (y).Q)" using asm prems by (auto simp add: fresh_prod fresh_atm) + finally have eq1: "(Cut .M (z).N){x:=.P}{b:=(y).Q} = (Cut .(P{b:=(y).Q}) (y).Q)" by simp + have "(Cut .M (z).N){b:=(y).Q}{x:=.(P{b:=(y).Q})} = (Cut .M (y).Q){x:=.(P{b:=(y).Q})}" + using prems asm by (simp add: fresh_atm) + also have "\ = Cut .(P{b:=(y).Q}) (y).(Q{x:=.(P{b:=(y).Q})})" using asm prems + by (auto simp add: fresh_prod fresh_atm subst_fresh) + also have "\ = Cut .(P{b:=(y).Q}) (y).Q" using asm prems by (simp add: forget) + finally have eq2: "(Cut .M (z).N){b:=(y).Q}{x:=.(P{b:=(y).Q})} = Cut .(P{b:=(y).Q}) (y).Q" + by simp + have "(Cut .M (z).N){x:=.P}{b:=(y).Q} = (Cut .M (z).N){b:=(y).Q}{x:=.(P{b:=(y).Q})}" + using eq1 eq2 by simp + } + moreover + { assume asm: "M \ Ax x c \ N = Ax z b" + have neq: "M{b:=(y).Q} \ Ax x c" + proof (cases "b\M") + case True then show ?thesis using asm by (simp add: forget) + next + case False then show ?thesis by (simp add: not_Ax1) + qed + have "(Cut .M (z).N){x:=.P}{b:=(y).Q} = (Cut .(M{x:=.P}) (z).(N{x:=.P})){b:=(y).Q}" + using asm prems by simp + also have "\ = (Cut .(M{x:=.P}) (z).N){b:=(y).Q}" using asm prems by (simp add: fresh_atm) + also have "\ = Cut .(M{x:=.P}{b:=(y).Q}) (y).Q" using asm prems by (simp add: abs_fresh) + also have "\ = Cut .(M{b:=(y).Q}{x:=.P{b:=(y).Q}}) (y).Q" using asm prems by simp + finally + have eq1: "(Cut .M (z).N){x:=.P}{b:=(y).Q} = Cut .(M{b:=(y).Q}{x:=.P{b:=(y).Q}}) (y).Q" + by simp + have "(Cut .M (z).N){b:=(y).Q}{x:=.(P{b:=(y).Q})} = + (Cut .(M{b:=(y).Q}) (y).Q){x:=.(P{b:=(y).Q})}" using asm prems by simp + also have "\ = Cut .(M{b:=(y).Q}{x:=.(P{b:=(y).Q})}) (y).(Q{x:=.(P{b:=(y).Q})})" + using asm prems neq by (auto simp add: fresh_prod fresh_atm subst_fresh abs_fresh) + also have "\ = Cut .(M{b:=(y).Q}{x:=.(P{b:=(y).Q})}) (y).Q" using asm prems by (simp add: forget) + finally have eq2: "(Cut .M (z).N){b:=(y).Q}{x:=.(P{b:=(y).Q})} + = Cut .(M{b:=(y).Q}{x:=.(P{b:=(y).Q})}) (y).Q" by simp + have "(Cut .M (z).N){x:=.P}{b:=(y).Q} = (Cut .M (z).N){b:=(y).Q}{x:=.(P{b:=(y).Q})}" + using eq1 eq2 by simp + } + moreover + { assume asm: "M = Ax x c \ N \ Ax z b" + have neq: "N{x:=.P} \ Ax z b" + proof (cases "x\N") + case True then show ?thesis using asm by (simp add: forget) + next + case False then show ?thesis by (simp add: not_Ax2) + qed + have "(Cut .M (z).N){x:=.P}{b:=(y).Q} = (Cut .P (z).(N{x:=.P})){b:=(y).Q}" + using asm prems by simp + also have "\ = Cut .(P{b:=(y).Q}) (z).(N{x:=.P}{b:=(y).Q})" using asm prems neq + by (simp add: abs_fresh) + also have "\ = Cut .(P{b:=(y).Q}) (z).(N{b:=(y).Q}{x:=.(P{b:=(y).Q})})" using asm prems by simp + finally have eq1: "(Cut .M (z).N){x:=.P}{b:=(y).Q} + = Cut .(P{b:=(y).Q}) (z).(N{b:=(y).Q}{x:=.(P{b:=(y).Q})})" by simp + have "(Cut .M (z).N){b:=(y).Q}{x:=.(P{b:=(y).Q})} + = (Cut .(M{b:=(y).Q}) (z).(N{b:=(y).Q})){x:=.(P{b:=(y).Q})}" + using asm prems by auto + also have "\ = (Cut .M (z).(N{b:=(y).Q})){x:=.(P{b:=(y).Q})}" + using asm prems by (auto simp add: fresh_atm) + also have "\ = Cut .(P{b:=(y).Q}) (z).(N{b:=(y).Q}{x:=.(P{b:=(y).Q})})" + using asm prems by (simp add: fresh_prod fresh_atm subst_fresh) + finally + have eq2: "(Cut .M (z).N){b:=(y).Q}{x:=.(P{b:=(y).Q})} + = Cut .(P{b:=(y).Q}) (z).(N{b:=(y).Q}{x:=.(P{b:=(y).Q})})" by simp + have "(Cut .M (z).N){x:=.P}{b:=(y).Q} = (Cut .M (z).N){b:=(y).Q}{x:=.(P{b:=(y).Q})}" + using eq1 eq2 by simp + } + moreover + { assume asm: "M \ Ax x c \ N \ Ax z b" + have neq1: "N{x:=.P} \ Ax z b" + proof (cases "x\N") + case True then show ?thesis using asm by (simp add: forget) + next + case False then show ?thesis by (simp add: not_Ax2) + qed + have neq2: "M{b:=(y).Q} \ Ax x c" + proof (cases "b\M") + case True then show ?thesis using asm by (simp add: forget) + next + case False then show ?thesis by (simp add: not_Ax1) + qed + have "(Cut .M (z).N){x:=.P}{b:=(y).Q} = (Cut .(M{x:=.P}) (z).(N{x:=.P})){b:=(y).Q}" + using asm prems by simp + also have "\ = Cut .(M{x:=.P}{b:=(y).Q}) (z).(N{x:=.P}{b:=(y).Q})" using asm prems neq1 + by (simp add: abs_fresh) + also have "\ = Cut .(M{b:=(y).Q}{x:=.(P{b:=(y).Q})}) (z).(N{b:=(y).Q}{x:=.(P{b:=(y).Q})})" + using asm prems by simp + finally have eq1: "(Cut .M (z).N){x:=.P}{b:=(y).Q} + = Cut .(M{b:=(y).Q}{x:=.(P{b:=(y).Q})}) (z).(N{b:=(y).Q}{x:=.(P{b:=(y).Q})})" by simp + have "(Cut .M (z).N){b:=(y).Q}{x:=.(P{b:=(y).Q})} = + (Cut .(M{b:=(y).Q}) (z).(N{b:=(y).Q})){x:=.(P{b:=(y).Q})}" using asm neq1 prems by simp + also have "\ = Cut .(M{b:=(y).Q}{x:=.(P{b:=(y).Q})}) (z).(N{b:=(y).Q}{x:=.(P{b:=(y).Q})})" + using asm neq2 prems by (simp add: fresh_prod fresh_atm subst_fresh) + finally have eq2: "(Cut .M (z).N){b:=(y).Q}{x:=.(P{b:=(y).Q})} = + Cut .(M{b:=(y).Q}{x:=.(P{b:=(y).Q})}) (z).(N{b:=(y).Q}{x:=.(P{b:=(y).Q})})" by simp + have "(Cut .M (z).N){x:=.P}{b:=(y).Q} = (Cut .M (z).N){b:=(y).Q}{x:=.(P{b:=(y).Q})}" + using eq1 eq2 by simp + } + ultimately show ?case by blast +next + case (NotR z M c) + then show ?case + apply(auto simp add: fresh_prod fresh_atm subst_fresh) + apply(subgoal_tac "\a'::coname. a'\(M{c:=(y).Q},M{c:=(y).Q}{x:=.P{c:=(y).Q}},Q,a,P,c,y)") + apply(erule exE) + apply(simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_NotR abs_fresh fresh_atm) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substn) + apply(simp add: fresh_prod fresh_atm subst_fresh abs_fresh) + apply(simp add: fresh_prod fresh_atm subst_fresh abs_fresh) + apply(simp add: forget) + apply(simp add: fresh_prod fresh_atm subst_fresh) + apply(rule exists_fresh'(2)[OF fs_coname1]) + done +next + case (NotL c M z) + then show ?case + apply(auto simp add: fresh_prod fresh_atm subst_fresh) + apply(subgoal_tac "\x'::name. x'\(P,M{x:=.P},P{b:=(y).Q},M{b:=(y).Q}{x:=.P{b:=(y).Q}},y,Q)") + apply(erule exE) + apply(simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_NotL abs_fresh fresh_atm) + apply(rule exists_fresh'(1)[OF fs_name1]) + done +next + case (AndR c1 M c2 N c3) + then show ?case + apply(auto simp add: fresh_prod fresh_atm subst_fresh) + apply(subgoal_tac "\a'::coname. a'\(Q,M{c3:=(y).Q},M{c3:=(y).Q}{x:=.P{c3:=(y).Q}},c2,c3,a, + P{c3:=(y).Q},N{c3:=(y).Q},N{c3:=(y).Q}{x:=.P{c3:=(y).Q}},c1)") + apply(erule exE) + apply(simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_AndR abs_fresh fresh_atm) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substn) + apply(simp_all add: fresh_atm abs_fresh subst_fresh) + apply(simp add: forget) + apply(rule exists_fresh'(2)[OF fs_coname1]) + done +next + case (AndL1 z1 M z2) + then show ?case + apply(auto simp add: fresh_prod fresh_atm subst_fresh) + apply(subgoal_tac "\x'::name. x'\(P,M{x:=.P},P{b:=(y).Q},z1,y,Q,M{b:=(y).Q}{x:=.P{b:=(y).Q}})") + apply(erule exE) + apply(simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_AndL1 abs_fresh fresh_atm) + apply(rule exists_fresh'(1)[OF fs_name1]) + done +next + case (AndL2 z1 M z2) + then show ?case + apply(auto simp add: fresh_prod fresh_atm subst_fresh) + apply(subgoal_tac "\x'::name. x'\(P,M{x:=.P},P{b:=(y).Q},z1,y,Q,M{b:=(y).Q}{x:=.P{b:=(y).Q}})") + apply(erule exE) + apply(simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_AndL2 abs_fresh fresh_atm) + apply(rule exists_fresh'(1)[OF fs_name1]) + done +next + case (OrL z1 M z2 N z3) + then show ?case + apply(auto simp add: fresh_prod fresh_atm subst_fresh) + apply(subgoal_tac "\x'::name. x'\(P,M{x:=.P},M{b:=(y).Q}{x:=.P{b:=(y).Q}},z2,z3,a,y,Q, + P{b:=(y).Q},N{x:=.P},N{b:=(y).Q}{x:=.P{b:=(y).Q}},z1)") + apply(erule exE) + apply(simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_OrL abs_fresh fresh_atm) + apply(simp add: trm.inject alpha) + apply(rule trans) + apply(rule substc.simps) + apply(simp_all add: fresh_atm subst_fresh) + apply(rule exists_fresh'(1)[OF fs_name1]) + done +next + case (OrR1 c1 M c2) + then show ?case + apply(auto simp add: fresh_prod fresh_atm subst_fresh) + apply(subgoal_tac "\a'::coname. a'\(Q,M{c2:=(y).Q},a,P{c2:=(y).Q},c1, + M{c2:=(y).Q}{x:=.P{c2:=(y).Q}})") + apply(erule exE) + apply(simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_OrR1 abs_fresh fresh_atm) + apply(simp_all add: fresh_atm subst_fresh abs_fresh) + apply(simp add: forget) + apply(rule exists_fresh'(2)[OF fs_coname1]) + done +next + case (OrR2 c1 M c2) + then show ?case + apply(auto simp add: fresh_prod fresh_atm subst_fresh) + apply(subgoal_tac "\a'::coname. a'\(Q,M{c2:=(y).Q},a,P{c2:=(y).Q},c1, + M{c2:=(y).Q}{x:=.P{c2:=(y).Q}})") + apply(erule exE) + apply(simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_OrR2 abs_fresh fresh_atm) + apply(simp_all add: fresh_atm subst_fresh abs_fresh) + apply(simp add: forget) + apply(rule exists_fresh'(2)[OF fs_coname1]) + done +next + case (ImpR z c M d) + then show ?case + apply(auto simp add: fresh_prod fresh_atm subst_fresh) + apply(subgoal_tac "\a'::coname. a'\(Q,M{d:=(y).Q},a,P{d:=(y).Q},c, + M{d:=(y).Q}{x:=.P{d:=(y).Q}})") + apply(erule exE) + apply(simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_ImpR abs_fresh fresh_atm) + apply(simp_all add: fresh_atm subst_fresh forget abs_fresh) + apply(rule exists_fresh'(2)[OF fs_coname1]) + done +next + case (ImpL c M z N u) + then show ?case + apply(auto simp add: fresh_prod fresh_atm subst_fresh) + apply(subgoal_tac "\z'::name. z'\(P,P{b:=(y).Q},M{u:=.P},N{u:=.P},y,Q, + M{b:=(y).Q}{u:=.P{b:=(y).Q}},N{b:=(y).Q}{u:=.P{b:=(y).Q}},z)") + apply(erule exE) + apply(simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_ImpL abs_fresh fresh_atm) + apply(simp add: trm.inject alpha) + apply(rule trans) + apply(rule substc.simps) + apply(simp_all add: fresh_atm subst_fresh forget) + apply(rule exists_fresh'(1)[OF fs_name1]) + done +qed + +lemma subst_subst2: + assumes a: "a\(b,P,N)" "x\(y,P,M)" "b\(M,N)" "y\P" + shows "M{a:=(x).N}{y:=.P} = M{y:=.P}{a:=(x).N{y:=.P}}" +using a +proof(nominal_induct M avoiding: a x N y b P rule: trm.induct) + case (Ax z c) + then show ?case + by (auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget trm.inject) +next + case (Cut d M' u M'') + then show ?case + apply(simp add: fresh_atm fresh_prod trm.inject abs_fresh) + apply(auto) + apply(simp add: fresh_atm) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substc) + apply(simp add: abs_fresh subst_fresh fresh_prod fresh_atm) + apply(simp add: fresh_prod subst_fresh fresh_atm abs_fresh) + apply(simp) + apply(simp add: forget) + apply(simp add: fresh_atm) + apply(case_tac "a\M'") + apply(simp add: forget) + apply(simp add: not_Ax1) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substc) + apply(simp add: abs_fresh subst_fresh fresh_prod fresh_atm) + apply(simp add: fresh_prod subst_fresh fresh_atm abs_fresh) + apply(auto)[1] + apply(case_tac "y\M''") + apply(simp add: forget) + apply(simp add: not_Ax2) + apply(simp add: forget) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substc) + apply(simp add: subst_fresh fresh_atm) + apply(simp add: abs_fresh subst_fresh) + apply(auto)[1] + apply(case_tac "y\M''") + apply(simp add: forget) + apply(simp add: not_Ax2) + apply(case_tac "a\M'") + apply(simp add: forget) + apply(simp add: not_Ax1) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substc) + apply(simp add: subst_fresh) + apply(simp add: subst_fresh abs_fresh) + apply(simp) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substc) + apply(simp add: subst_fresh fresh_atm) + apply(simp add: subst_fresh abs_fresh) + apply(auto)[1] + apply(case_tac "y\M''") + apply(simp add: forget) + apply(simp add: not_Ax2) + done +next + case (NotR z M' d) + then show ?case + apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget) + apply(subgoal_tac "\a'::coname. a'\(y,P,N,N{y:=.P},M'{d:=(x).N},M'{y:=.P}{d:=(x).N{y:=.P}})") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_NotR) + apply(rule trans) + apply(rule better_Cut_substn) + apply(simp add: abs_fresh subst_fresh) + apply(simp add: fresh_prod subst_fresh fresh_atm) + apply(simp) + apply(simp add: trm.inject alpha) + apply(rule trans) + apply(rule substn.simps) + apply(simp add: fresh_prod fresh_atm) + apply(simp add: fresh_atm) + apply(rule exists_fresh'(2)[OF fs_coname1]) + done +next + case (NotL d M' z) + then show ?case + apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget trm.inject) + apply(subgoal_tac "\x'::name. x'\(z,y,P,N,N{y:=.P},M'{y:=.P},M'{y:=.P}{a:=(x).N{y:=.P}})") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_NotL) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substc) + apply(simp add: abs_fresh subst_fresh fresh_atm) + apply(simp add: fresh_prod subst_fresh fresh_atm abs_fresh) + apply(simp) + apply(simp add: trm.inject alpha forget) + apply(rule trans) + apply(rule substc.simps) + apply(simp add: fresh_prod fresh_atm) + apply(simp add: fresh_atm subst_fresh) + apply(simp) + apply(rule exists_fresh'(1)[OF fs_name1]) + done +next + case (AndR d M' e M'' f) + then show ?case + apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget trm.inject) + apply(subgoal_tac "\a'::coname. a'\(P,b,d,e,N,N{y:=.P},M'{f:=(x).N},M''{f:=(x).N}, + M'{y:=.P}{f:=(x).N{y:=.P}},M''{y:=.P}{f:=(x).N{y:=.P}})") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_AndR) + apply(rule trans) + apply(rule better_Cut_substn) + apply(simp add: abs_fresh subst_fresh fresh_atm) + apply(simp add: fresh_prod subst_fresh fresh_atm abs_fresh) + apply(simp) + apply(simp add: trm.inject alpha forget) + apply(rule trans) + apply(rule substn.simps) + apply(simp add: fresh_prod fresh_atm subst_fresh) + apply(simp add: fresh_atm subst_fresh) + apply(simp add: fresh_atm) + apply(simp) + apply(rule exists_fresh'(2)[OF fs_coname1]) + done +next + case (AndL1 z M' u) + then show ?case + apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget trm.inject) + apply(subgoal_tac "\x'::name. x'\(P,b,z,u,x,N,M'{y:=.P},M'{y:=.P}{a:=(x).N{y:=.P}})") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_AndL1) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substc) + apply(simp add: abs_fresh subst_fresh fresh_atm) + apply(simp add: fresh_prod subst_fresh fresh_atm abs_fresh) + apply(simp) + apply(simp add: trm.inject alpha forget) + apply(rule trans) + apply(rule substc.simps) + apply(simp add: fresh_prod fresh_atm subst_fresh) + apply(simp) + apply(rule exists_fresh'(1)[OF fs_name1]) + done +next + case (AndL2 z M' u) + then show ?case + apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget trm.inject) + apply(subgoal_tac "\x'::name. x'\(P,b,z,u,x,N,M'{y:=.P},M'{y:=.P}{a:=(x).N{y:=.P}})") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_AndL2) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substc) + apply(simp add: abs_fresh subst_fresh fresh_atm) + apply(simp add: fresh_prod subst_fresh fresh_atm abs_fresh) + apply(simp) + apply(simp add: trm.inject alpha forget) + apply(rule trans) + apply(rule substc.simps) + apply(simp add: fresh_prod fresh_atm subst_fresh) + apply(simp) + apply(rule exists_fresh'(1)[OF fs_name1]) + done +next + case (OrL u M' v M'' w) + then show ?case + apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget trm.inject) + apply(subgoal_tac "\z'::name. z'\(P,b,u,w,v,N,N{y:=.P},M'{y:=.P},M''{y:=.P}, + M'{y:=.P}{a:=(x).N{y:=.P}},M''{y:=.P}{a:=(x).N{y:=.P}})") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_OrL) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substc) + apply(simp add: abs_fresh subst_fresh fresh_atm) + apply(simp add: fresh_prod subst_fresh fresh_atm abs_fresh) + apply(simp) + apply(simp add: trm.inject alpha forget) + apply(rule trans) + apply(rule substc.simps) + apply(simp add: fresh_prod fresh_atm subst_fresh) + apply(simp add: fresh_atm subst_fresh) + apply(simp add: fresh_atm) + apply(simp) + apply(rule exists_fresh'(1)[OF fs_name1]) + done +next + case (OrR1 e M' f) + then show ?case + apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget trm.inject) + apply(subgoal_tac "\c'::coname. c'\(P,b,e,f,x,N,N{y:=.P}, + M'{f:=(x).N},M'{y:=.P}{f:=(x).N{y:=.P}})") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_OrR1) + apply(rule trans) + apply(rule better_Cut_substn) + apply(simp add: abs_fresh subst_fresh fresh_atm) + apply(simp add: fresh_prod subst_fresh fresh_atm abs_fresh) + apply(simp) + apply(simp add: trm.inject alpha forget) + apply(rule trans) + apply(rule substn.simps) + apply(simp add: fresh_prod fresh_atm subst_fresh) + apply(simp) + apply(rule exists_fresh'(2)[OF fs_coname1]) + done +next + case (OrR2 e M' f) + then show ?case + apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget trm.inject) + apply(subgoal_tac "\c'::coname. c'\(P,b,e,f,x,N,N{y:=.P}, + M'{f:=(x).N},M'{y:=.P}{f:=(x).N{y:=.P}})") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_OrR2) + apply(rule trans) + apply(rule better_Cut_substn) + apply(simp add: abs_fresh subst_fresh fresh_atm) + apply(simp add: fresh_prod subst_fresh fresh_atm abs_fresh) + apply(simp) + apply(simp add: trm.inject alpha forget) + apply(rule trans) + apply(rule substn.simps) + apply(simp add: fresh_prod fresh_atm subst_fresh) + apply(simp) + apply(rule exists_fresh'(2)[OF fs_coname1]) + done +next + case (ImpR x e M' f) + then show ?case + apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget trm.inject) + apply(subgoal_tac "\c'::coname. c'\(P,b,e,f,x,N,N{y:=.P}, + M'{f:=(x).N},M'{y:=.P}{f:=(x).N{y:=.P}})") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_ImpR) + apply(rule trans) + apply(rule better_Cut_substn) + apply(simp add: abs_fresh subst_fresh fresh_atm) + apply(simp add: fresh_prod subst_fresh fresh_atm abs_fresh) + apply(simp) + apply(simp add: trm.inject alpha forget) + apply(rule trans) + apply(rule substn.simps) + apply(simp add: fresh_prod fresh_atm subst_fresh) + apply(simp add: fresh_atm) + apply(simp add: fresh_atm trm.inject alpha abs_fresh fin_supp abs_supp) + apply(rule exists_fresh'(2)[OF fs_coname1]) + apply(simp add: fresh_atm trm.inject alpha abs_fresh fin_supp abs_supp) + done +next + case (ImpL e M' v M'' w) + then show ?case + apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget trm.inject) + apply(subgoal_tac "\z'::name. z'\(P,b,e,w,v,N,N{y:=.P},M'{w:=.P},M''{w:=.P}, + M'{w:=.P}{a:=(x).N{w:=.P}},M''{w:=.P}{a:=(x).N{w:=.P}})") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_ImpL) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substc) + apply(simp add: abs_fresh subst_fresh fresh_atm) + apply(simp add: fresh_prod subst_fresh fresh_atm abs_fresh) + apply(simp) + apply(simp add: trm.inject alpha forget) + apply(rule trans) + apply(rule substc.simps) + apply(simp add: fresh_prod fresh_atm subst_fresh) + apply(simp add: fresh_atm subst_fresh) + apply(simp add: fresh_atm) + apply(rule exists_fresh'(1)[OF fs_name1]) + done +qed + +lemma subst_subst3: + assumes a: "a\(P,N,c)" "c\(M,N)" "x\(y,P,M)" "y\(P,x)" "M\Ax y a" + shows "N{x:=.M}{y:=.P} = N{y:=.P}{x:=.(M{y:=.P})}" +using a +proof(nominal_induct N avoiding: x y a c M P rule: trm.induct) + case (Ax z c) + then show ?case + by(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget) +next + case (Cut d M' u M'') + then show ?case + apply(simp add: fresh_atm fresh_prod trm.inject abs_fresh) + apply(auto) + apply(simp add: fresh_atm) + apply(simp add: trm.inject) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substn) + apply(simp add: abs_fresh subst_fresh) + apply(simp add: fresh_prod subst_fresh fresh_atm) + apply(subgoal_tac "P \ Ax x c") + apply(simp) + apply(simp add: forget) + apply(clarify) + apply(simp add: fresh_atm) + apply(case_tac "x\M'") + apply(simp add: forget) + apply(simp add: not_Ax2) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substn) + apply(simp add: abs_fresh subst_fresh) + apply(simp add: fresh_prod subst_fresh fresh_atm) + apply(simp) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substn) + apply(simp add: abs_fresh subst_fresh) + apply(simp add: fresh_prod subst_fresh fresh_atm) + apply(auto) + apply(case_tac "y\M'") + apply(simp add: forget) + apply(simp add: not_Ax2) + done +next + case NotR + then show ?case + by(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget) +next + case (NotL d M' u) + then show ?case + apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget) + apply(subgoal_tac "\x'::name. x'\(y,P,M,M{y:=.P},M'{x:=.M},M'{y:=.P}{x:=.M{y:=.P}})") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_NotL) + apply(rule trans) + apply(rule better_Cut_substn) + apply(simp add: abs_fresh subst_fresh) + apply(simp add: fresh_prod subst_fresh fresh_atm) + apply(simp) + apply(simp add: trm.inject alpha) + apply(rule trans) + apply(rule substn.simps) + apply(simp add: fresh_prod fresh_atm) + apply(simp add: fresh_atm) + apply(rule exists_fresh'(1)[OF fs_name1]) + apply(subgoal_tac "\x'::name. x'\(x,y,P,M,M'{y:=.P},M'{y:=.P}{x:=.M{y:=.P}})") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_NotL) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substn) + apply(simp add: abs_fresh subst_fresh) + apply(simp add: fresh_atm subst_fresh fresh_prod) + apply(subgoal_tac "P \ Ax x c") + apply(simp) + apply(simp add: forget trm.inject alpha) + apply(rule trans) + apply(rule substn.simps) + apply(simp add: fresh_atm subst_fresh) + apply(simp add: fresh_atm) + apply(clarify) + apply(simp add: fresh_atm) + apply(rule exists_fresh'(1)[OF fs_name1]) + done +next + case AndR + then show ?case + by(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget) +next + case (AndL1 u M' v) + then show ?case + apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget) + apply(subgoal_tac "\x'::name. x'\(u,y,v,P,M,M{y:=.P},M'{x:=.M},M'{y:=.P}{x:=.M{y:=.P}})") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_AndL1) + apply(rule trans) + apply(rule better_Cut_substn) + apply(simp add: abs_fresh subst_fresh) + apply(simp add: fresh_prod subst_fresh fresh_atm) + apply(simp) + apply(simp add: trm.inject alpha) + apply(rule trans) + apply(rule substn.simps) + apply(simp add: fresh_prod fresh_atm) + apply(simp add: fresh_atm) + apply(rule exists_fresh'(1)[OF fs_name1]) + apply(subgoal_tac "\x'::name. x'\(x,y,u,v,P,M,M'{y:=.P},M'{y:=.P}{x:=.M{y:=.P}})") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_AndL1) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substn) + apply(simp add: abs_fresh subst_fresh) + apply(simp add: fresh_atm subst_fresh fresh_prod) + apply(subgoal_tac "P \ Ax x c") + apply(simp) + apply(simp add: forget trm.inject alpha) + apply(rule trans) + apply(rule substn.simps) + apply(simp add: fresh_atm subst_fresh) + apply(simp add: fresh_atm) + apply(clarify) + apply(simp add: fresh_atm) + apply(rule exists_fresh'(1)[OF fs_name1]) + done +next + case (AndL2 u M' v) + then show ?case + apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget) + apply(subgoal_tac "\x'::name. x'\(u,y,v,P,M,M{y:=.P},M'{x:=.M},M'{y:=.P}{x:=.M{y:=.P}})") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_AndL2) + apply(rule trans) + apply(rule better_Cut_substn) + apply(simp add: abs_fresh subst_fresh) + apply(simp add: fresh_prod subst_fresh fresh_atm) + apply(simp) + apply(simp add: trm.inject alpha) + apply(rule trans) + apply(rule substn.simps) + apply(simp add: fresh_prod fresh_atm) + apply(simp add: fresh_atm) + apply(rule exists_fresh'(1)[OF fs_name1]) + apply(subgoal_tac "\x'::name. x'\(x,y,u,v,P,M,M'{y:=.P},M'{y:=.P}{x:=.M{y:=.P}})") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_AndL2) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substn) + apply(simp add: abs_fresh subst_fresh) + apply(simp add: fresh_atm subst_fresh fresh_prod) + apply(subgoal_tac "P \ Ax x c") + apply(simp) + apply(simp add: forget trm.inject alpha) + apply(rule trans) + apply(rule substn.simps) + apply(simp add: fresh_atm subst_fresh) + apply(simp add: fresh_atm) + apply(clarify) + apply(simp add: fresh_atm) + apply(rule exists_fresh'(1)[OF fs_name1]) + done +next + case OrR1 + then show ?case + by(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget) +next + case OrR2 + then show ?case + by(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget) +next + case (OrL x1 M' x2 M'' x3) + then show ?case + apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget) + apply(subgoal_tac "\x'::name. x'\(y,P,M,M{y:=.P},M'{x:=.M},M'{y:=.P}{x:=.M{y:=.P}}, + x1,x2,x3,M''{x:=.M},M''{y:=.P}{x:=.M{y:=.P}})") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_OrL) + apply(rule trans) + apply(rule better_Cut_substn) + apply(simp add: abs_fresh subst_fresh) + apply(simp add: fresh_prod subst_fresh fresh_atm) + apply(simp) + apply(simp add: trm.inject alpha) + apply(rule trans) + apply(rule substn.simps) + apply(simp add: fresh_prod fresh_atm subst_fresh) + apply(simp add: fresh_prod fresh_atm subst_fresh) + apply(simp add: fresh_atm) + apply(simp add: fresh_atm) + apply(rule exists_fresh'(1)[OF fs_name1]) + apply(subgoal_tac "\x'::name. x'\(x,y,P,M,M'{y:=.P},M'{y:=.P}{x:=.M{y:=.P}}, + x1,x2,x3,M''{y:=.P},M''{y:=.P}{x:=.M{y:=.P}})") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_OrL) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substn) + apply(simp add: abs_fresh subst_fresh) + apply(simp add: fresh_atm subst_fresh fresh_prod) + apply(simp add: fresh_prod fresh_atm) + apply(auto) + apply(simp add: fresh_atm) + apply(simp add: forget trm.inject alpha) + apply(rule trans) + apply(rule substn.simps) + apply(simp add: fresh_atm subst_fresh) + apply(simp add: fresh_atm subst_fresh) + apply(simp add: fresh_atm) + apply(simp add: fresh_atm) + apply(rule exists_fresh'(1)[OF fs_name1]) + done +next + case ImpR + then show ?case + by(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget) +next + case (ImpL d M' x1 M'' x2) + then show ?case + apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget) + apply(subgoal_tac "\x'::name. x'\(y,P,M,M{y:=.P},M'{x2:=.M},M'{y:=.P}{x2:=.M{y:=.P}}, + x1,x2,M''{x2:=.M},M''{y:=.P}{x2:=.M{y:=.P}})") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_ImpL) + apply(rule trans) + apply(rule better_Cut_substn) + apply(simp add: abs_fresh subst_fresh) + apply(simp add: fresh_prod subst_fresh fresh_atm) + apply(simp) + apply(simp add: trm.inject alpha) + apply(rule trans) + apply(rule substn.simps) + apply(simp add: fresh_prod fresh_atm subst_fresh) + apply(simp add: fresh_prod fresh_atm subst_fresh) + apply(simp add: fresh_atm) + apply(rule exists_fresh'(1)[OF fs_name1]) + apply(subgoal_tac "\x'::name. x'\(x,y,P,M,M'{x2:=.P},M'{x2:=.P}{x:=.M{x2:=.P}}, + x1,x2,M''{x2:=.P},M''{x2:=.P}{x:=.M{x2:=.P}})") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_ImpL) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substn) + apply(simp add: abs_fresh subst_fresh) + apply(simp add: fresh_atm subst_fresh fresh_prod) + apply(simp add: fresh_prod fresh_atm) + apply(auto) + apply(simp add: fresh_atm) + apply(simp add: forget trm.inject alpha) + apply(rule trans) + apply(rule substn.simps) + apply(simp add: fresh_atm subst_fresh) + apply(simp add: fresh_atm subst_fresh) + apply(simp add: fresh_atm) + apply(rule exists_fresh'(1)[OF fs_name1]) + done +qed + +lemma subst_subst4: + assumes a: "x\(P,N,y)" "y\(M,N)" "a\(c,P,M)" "c\(P,a)" "M\Ax x c" + shows "N{a:=(x).M}{c:=(y).P} = N{c:=(y).P}{a:=(x).(M{c:=(y).P})}" +using a +proof(nominal_induct N avoiding: x y a c M P rule: trm.induct) + case (Ax z c) + then show ?case + by (auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget) +next + case (Cut d M' u M'') + then show ?case + apply(simp add: fresh_atm fresh_prod trm.inject abs_fresh) + apply(auto) + apply(simp add: fresh_atm) + apply(simp add: trm.inject) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substc) + apply(simp add: abs_fresh subst_fresh fresh_atm) + apply(simp add: fresh_prod subst_fresh abs_fresh fresh_atm) + apply(subgoal_tac "P \ Ax y a") + apply(simp) + apply(simp add: forget) + apply(clarify) + apply(simp add: fresh_atm) + apply(case_tac "a\M''") + apply(simp add: forget) + apply(simp add: not_Ax1) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substc) + apply(simp add: fresh_prod subst_fresh fresh_atm) + apply(simp add: abs_fresh subst_fresh) + apply(simp) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substc) + apply(simp add: fresh_prod subst_fresh fresh_atm) + apply(simp add: abs_fresh subst_fresh) + apply(auto) + apply(case_tac "c\M''") + apply(simp add: forget) + apply(simp add: not_Ax1) + done +next + case NotL + then show ?case + by(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget) +next + case (NotR u M' d) + then show ?case + apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget) + apply(generate_fresh "coname") + apply(fresh_fun_simp) + apply(fresh_fun_simp) + apply(simp add: abs_fresh subst_fresh) + apply(rule trans) + apply(rule better_Cut_substc) + apply(simp) + apply(simp add: abs_fresh) + apply(simp) + apply(simp add: trm.inject alpha) + apply(rule trans) + apply(rule substc.simps) + apply(simp add: fresh_prod fresh_atm) + apply(auto simp add: fresh_atm fresh_prod)[1] + apply(generate_fresh "coname") + apply(fresh_fun_simp) + apply(fresh_fun_simp) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substc) + apply(simp add: fresh_prod fresh_atm subst_fresh) + apply(simp add: abs_fresh subst_fresh) + apply(auto simp add: fresh_atm) + apply(simp add: trm.inject alpha forget) + apply(rule trans) + apply(rule substc.simps) + apply(simp add: fresh_atm subst_fresh) + apply(auto simp add: fresh_prod fresh_atm) + done +next + case AndL1 + then show ?case + by(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget) +next + case AndL2 + then show ?case + by(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget) +next + case (AndR d M e M' f) + then show ?case + apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget) + apply(generate_fresh "coname") + apply(fresh_fun_simp) + apply(fresh_fun_simp) + apply(simp add: abs_fresh subst_fresh) + apply(rule trans) + apply(rule better_Cut_substc) + apply(simp) + apply(simp add: abs_fresh) + apply(simp) + apply(simp add: trm.inject alpha) + apply(rule trans) + apply(rule substc.simps) + apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1] + apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1] + apply(simp) + apply(auto simp add: fresh_atm fresh_prod)[1] + apply(generate_fresh "coname") + apply(fresh_fun_simp) + apply(fresh_fun_simp) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substc) + apply(simp add: subst_fresh fresh_atm fresh_prod) + apply(simp add: abs_fresh subst_fresh) + apply(auto simp add: fresh_atm)[1] + apply(simp add: trm.inject alpha forget) + apply(rule trans) + apply(rule substc.simps) + apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1] + apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1] + apply(simp) + apply(auto simp add: fresh_atm fresh_prod)[1] + done +next + case OrL + then show ?case + by(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget) +next + case (OrR1 d M' e) + then show ?case + apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget) + apply(generate_fresh "coname") + apply(fresh_fun_simp) + apply(fresh_fun_simp) + apply(simp add: abs_fresh subst_fresh) + apply(rule trans) + apply(rule better_Cut_substc) + apply(simp) + apply(simp add: abs_fresh) + apply(simp) + apply(simp add: trm.inject alpha) + apply(rule trans) + apply(rule substc.simps) + apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1] + apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1] + apply(generate_fresh "coname") + apply(fresh_fun_simp) + apply(fresh_fun_simp) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substc) + apply(simp add: subst_fresh fresh_atm fresh_prod) + apply(simp add: abs_fresh subst_fresh) + apply(auto simp add: fresh_atm)[1] + apply(simp add: trm.inject alpha forget) + apply(rule trans) + apply(rule substc.simps) + apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1] + apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1] + done +next + case (OrR2 d M' e) + then show ?case + apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget) + apply(generate_fresh "coname") + apply(fresh_fun_simp) + apply(fresh_fun_simp) + apply(simp add: abs_fresh subst_fresh) + apply(rule trans) + apply(rule better_Cut_substc) + apply(simp) + apply(simp add: abs_fresh) + apply(simp) + apply(simp add: trm.inject alpha) + apply(rule trans) + apply(rule substc.simps) + apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1] + apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1] + apply(generate_fresh "coname") + apply(fresh_fun_simp) + apply(fresh_fun_simp) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substc) + apply(simp add: subst_fresh fresh_atm fresh_prod) + apply(simp add: abs_fresh subst_fresh) + apply(auto simp add: fresh_atm)[1] + apply(simp add: trm.inject alpha forget) + apply(rule trans) + apply(rule substc.simps) + apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1] + apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1] + done +next + case ImpL + then show ?case + by(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget) +next + case (ImpR u d M' e) + then show ?case + apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget) + apply(generate_fresh "coname") + apply(fresh_fun_simp) + apply(fresh_fun_simp) + apply(simp add: abs_fresh subst_fresh) + apply(rule trans) + apply(rule better_Cut_substc) + apply(simp) + apply(simp add: abs_fresh) + apply(simp) + apply(simp add: trm.inject alpha) + apply(rule trans) + apply(rule substc.simps) + apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1] + apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1] + apply(auto simp add: fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp)[1] + apply(generate_fresh "coname") + apply(fresh_fun_simp) + apply(fresh_fun_simp) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substc) + apply(simp add: subst_fresh fresh_atm fresh_prod) + apply(simp add: abs_fresh subst_fresh) + apply(auto simp add: fresh_atm)[1] + apply(simp add: trm.inject alpha forget) + apply(rule trans) + apply(rule substc.simps) + apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1] + apply(auto simp add: fresh_prod fresh_atm subst_fresh)[1] + apply(auto simp add: fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp)[1] + apply(auto simp add: fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp)[1] + done +qed + +text {* Reduction *} + +lemma fin_not_Cut: + assumes a: "fin M x" + shows "\(\a M' x N'. M = Cut .M' (x).N')" +using a +by (induct) (auto) + +lemma fresh_not_fin: + assumes a: "x\M" + shows "\fin M x" +proof - + have "fin M x \ x\M \ False" by (induct rule: fin.induct) (auto simp add: abs_fresh fresh_atm) + with a show "\fin M x" by blast +qed + +lemma fresh_not_fic: + assumes a: "a\M" + shows "\fic M a" +proof - + have "fic M a \ a\M \ False" by (induct rule: fic.induct) (auto simp add: abs_fresh fresh_atm) + with a show "\fic M a" by blast +qed + +lemma c_redu_subst1: + assumes a: "M \\<^isub>c M'" "c\M" "y\P" + shows "M{y:=.P} \\<^isub>c M'{y:=.P}" +using a +proof(nominal_induct avoiding: y c P rule: c_redu.strong_induct) + case (left M a N x) + then show ?case + apply - + apply(simp) + apply(rule conjI) + apply(force) + apply(auto) + apply(subgoal_tac "M{a:=(x).N}{y:=.P} = M{y:=.P}{a:=(x).(N{y:=.P})}")(*A*) + apply(simp) + apply(rule c_redu.intros) + apply(rule not_fic_subst1) + apply(simp) + apply(simp add: subst_fresh) + apply(simp add: subst_fresh) + apply(simp add: abs_fresh fresh_atm) + apply(rule subst_subst2) + apply(simp add: fresh_prod fresh_atm) + apply(simp add: fresh_prod fresh_atm) + apply(simp add: fresh_prod fresh_atm) + apply(simp) + done +next + case (right N x a M) + then show ?case + apply - + apply(simp) + apply(rule conjI) + (* case M = Ax y a *) + apply(rule impI) + apply(subgoal_tac "N{x:=.Ax y a}{y:=.P} = N{y:=.P}{x:=.P}") + apply(simp) + apply(rule c_redu.right) + apply(rule not_fin_subst2) + apply(simp) + apply(rule subst_fresh) + apply(simp add: abs_fresh) + apply(simp add: abs_fresh) + apply(rule sym) + apply(rule interesting_subst1') + apply(simp add: fresh_atm) + apply(simp) + apply(simp) + (* case M \ Ax y a*) + apply(rule impI) + apply(subgoal_tac "N{x:=.M}{y:=.P} = N{y:=.P}{x:=.(M{y:=.P})}") + apply(simp) + apply(rule c_redu.right) + apply(rule not_fin_subst2) + apply(simp) + apply(simp add: subst_fresh) + apply(simp add: subst_fresh) + apply(simp add: abs_fresh fresh_atm) + apply(rule subst_subst3) + apply(simp_all add: fresh_atm fresh_prod) + done +qed + +lemma c_redu_subst2: + assumes a: "M \\<^isub>c M'" "c\P" "y\M" + shows "M{c:=(y).P} \\<^isub>c M'{c:=(y).P}" +using a +proof(nominal_induct avoiding: y c P rule: c_redu.strong_induct) + case (right N x a M) + then show ?case + apply - + apply(simp) + apply(rule conjI) + apply(force) + apply(auto) + apply(subgoal_tac "N{x:=.M}{c:=(y).P} = N{c:=(y).P}{x:=.(M{c:=(y).P})}")(*A*) + apply(simp) + apply(rule c_redu.intros) + apply(rule not_fin_subst1) + apply(simp) + apply(simp add: subst_fresh) + apply(simp add: subst_fresh) + apply(simp add: abs_fresh fresh_atm) + apply(rule subst_subst1) + apply(simp add: fresh_prod fresh_atm) + apply(simp add: fresh_prod fresh_atm) + apply(simp add: fresh_prod fresh_atm) + apply(simp) + done +next + case (left M a N x) + then show ?case + apply - + apply(simp) + apply(rule conjI) + (* case N = Ax x c *) + apply(rule impI) + apply(subgoal_tac "M{a:=(x).Ax x c}{c:=(y).P} = M{c:=(y).P}{a:=(y).P}") + apply(simp) + apply(rule c_redu.left) + apply(rule not_fic_subst2) + apply(simp) + apply(simp) + apply(rule subst_fresh) + apply(simp add: abs_fresh) + apply(rule sym) + apply(rule interesting_subst2') + apply(simp add: fresh_atm) + apply(simp) + apply(simp) + (* case M \ Ax y a*) + apply(rule impI) + apply(subgoal_tac "M{a:=(x).N}{c:=(y).P} = M{c:=(y).P}{a:=(x).(N{c:=(y).P})}") + apply(simp) + apply(rule c_redu.left) + apply(rule not_fic_subst2) + apply(simp) + apply(simp add: subst_fresh) + apply(simp add: subst_fresh) + apply(simp add: abs_fresh fresh_atm) + apply(rule subst_subst4) + apply(simp add: fresh_prod fresh_atm) + apply(simp add: fresh_prod fresh_atm) + apply(simp add: fresh_prod fresh_atm) + apply(simp add: fresh_prod fresh_atm) + apply(simp) + done +qed + +lemma c_redu_subst1': + assumes a: "M \\<^isub>c M'" + shows "M{y:=.P} \\<^isub>c M'{y:=.P}" +using a +proof - + obtain y'::"name" where fs1: "y'\(M,M',P,P,y)" by (rule exists_fresh(1), rule fin_supp, blast) + obtain c'::"coname" where fs2: "c'\(M,M',P,P,c)" by (rule exists_fresh(2), rule fin_supp, blast) + have "M{y:=.P} = ([(y',y)]\M){y':=.([(c',c)]\P)}" using fs1 fs2 + apply - + apply(rule trans) + apply(rule_tac y="y'" in subst_rename(3)) + apply(simp) + apply(rule subst_rename(4)) + apply(simp) + done + also have "\ \\<^isub>c ([(y',y)]\M'){y':=.([(c',c)]\P)}" using fs1 fs2 + apply - + apply(rule c_redu_subst1) + apply(simp add: c_redu.eqvt a) + apply(simp_all add: fresh_left calc_atm fresh_prod) + done + also have "\ = M'{y:=.P}" using fs1 fs2 + apply - + apply(rule sym) + apply(rule trans) + apply(rule_tac y="y'" in subst_rename(3)) + apply(simp) + apply(rule subst_rename(4)) + apply(simp) + done + finally show ?thesis by simp +qed + +lemma c_redu_subst2': + assumes a: "M \\<^isub>c M'" + shows "M{c:=(y).P} \\<^isub>c M'{c:=(y).P}" +using a +proof - + obtain y'::"name" where fs1: "y'\(M,M',P,P,y)" by (rule exists_fresh(1), rule fin_supp, blast) + obtain c'::"coname" where fs2: "c'\(M,M',P,P,c)" by (rule exists_fresh(2), rule fin_supp, blast) + have "M{c:=(y).P} = ([(c',c)]\M){c':=(y').([(y',y)]\P)}" using fs1 fs2 + apply - + apply(rule trans) + apply(rule_tac c="c'" in subst_rename(1)) + apply(simp) + apply(rule subst_rename(2)) + apply(simp) + done + also have "\ \\<^isub>c ([(c',c)]\M'){c':=(y').([(y',y)]\P)}" using fs1 fs2 + apply - + apply(rule c_redu_subst2) + apply(simp add: c_redu.eqvt a) + apply(simp_all add: fresh_left calc_atm fresh_prod) + done + also have "\ = M'{c:=(y).P}" using fs1 fs2 + apply - + apply(rule sym) + apply(rule trans) + apply(rule_tac c="c'" in subst_rename(1)) + apply(simp) + apply(rule subst_rename(2)) + apply(simp) + done + + finally show ?thesis by simp +qed + +lemma aux1: + assumes a: "M = M'" "M' \\<^isub>l M''" + shows "M \\<^isub>l M''" +using a by simp + +lemma aux2: + assumes a: "M \\<^isub>l M'" "M' = M''" + shows "M \\<^isub>l M''" +using a by simp + +lemma aux3: + assumes a: "M = M'" "M' \\<^isub>a* M''" + shows "M \\<^isub>a* M''" +using a by simp + +lemma aux4: + assumes a: "M = M'" + shows "M \\<^isub>a* M'" +using a by blast + +lemma l_redu_subst1: + assumes a: "M \\<^isub>l M'" + shows "M{y:=.P} \\<^isub>a* M'{y:=.P}" +using a +proof(nominal_induct M M' avoiding: y c P rule: l_redu.strong_induct) + case LAxR + then show ?case + apply - + apply(rule aux3) + apply(rule better_Cut_substn) + apply(simp add: abs_fresh) + apply(simp) + apply(simp add: fresh_atm) + apply(auto) + apply(rule aux4) + apply(simp add: trm.inject alpha calc_atm fresh_atm) + apply(rule a_star_trans) + apply(rule a_starI) + apply(rule al_redu) + apply(rule l_redu.intros) + apply(simp add: subst_fresh) + apply(simp add: fresh_atm) + apply(rule fic_subst2) + apply(simp_all) + apply(rule aux4) + apply(rule subst_comm') + apply(simp_all) + done +next + case LAxL + then show ?case + apply - + apply(rule aux3) + apply(rule better_Cut_substn) + apply(simp add: abs_fresh) + apply(simp) + apply(simp add: trm.inject fresh_atm) + apply(auto) + apply(rule aux4) + apply(rule sym) + apply(rule fin_substn_nrename) + apply(simp_all) + apply(rule a_starI) + apply(rule al_redu) + apply(rule aux2) + apply(rule l_redu.intros) + apply(simp add: subst_fresh) + apply(simp add: fresh_atm) + apply(rule fin_subst1) + apply(simp_all) + apply(rule subst_comm') + apply(simp_all) + done +next + case (LNot v M N u a b) + then show ?case + proof - + { assume asm: "N\Ax y b" + have "(Cut .NotR (u).M a (v).NotL .N v){y:=.P} = + (Cut .NotR (u).(M{y:=.P}) a (v).NotL .(N{y:=.P}) v)" using prems + by (simp add: subst_fresh abs_fresh fresh_atm) + also have "\ \\<^isub>l (Cut .(N{y:=.P}) (u).(M{y:=.P}))" using prems + by (auto intro: l_redu.intros simp add: subst_fresh) + also have "\ = (Cut .N (u).M){y:=.P}" using prems + by (simp add: subst_fresh abs_fresh fresh_atm) + finally have ?thesis by auto + } + moreover + { assume asm: "N=Ax y b" + have "(Cut .NotR (u).M a (v).NotL .N v){y:=.P} = + (Cut .NotR (u).(M{y:=.P}) a (v).NotL .(N{y:=.P}) v)" using prems + by (simp add: subst_fresh abs_fresh fresh_atm) + also have "\ \\<^isub>a* (Cut .(N{y:=.P}) (u).(M{y:=.P}))" using prems + apply - + apply(rule a_starI) + apply(rule al_redu) + apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh) + done + also have "\ = (Cut .(Cut .P (y).Ax y b) (u).(M{y:=.P}))" using prems + by simp + also have "\ \\<^isub>a* (Cut .(P[c\c>b]) (u).(M{y:=.P}))" + proof (cases "fic P c") + case True + assume "fic P c" + then show ?thesis using prems + apply - + apply(rule a_starI) + apply(rule better_CutL_intro) + apply(rule al_redu) + apply(rule better_LAxR_intro) + apply(simp) + done + next + case False + assume "\fic P c" + then show ?thesis + apply - + apply(rule a_star_CutL) + apply(rule a_star_trans) + apply(rule a_starI) + apply(rule ac_redu) + apply(rule better_left) + apply(simp) + apply(simp add: subst_with_ax2) + done + qed + also have "\ = (Cut .N (u).M){y:=.P}" using prems + apply - + apply(auto simp add: subst_fresh abs_fresh) + apply(simp add: trm.inject) + apply(simp add: alpha fresh_atm) + apply(rule sym) + apply(rule crename_swap) + apply(simp) + done + finally have "(Cut .NotR (u).M a (v).NotL .N v){y:=.P} \\<^isub>a* (Cut .N (u).M){y:=.P}" + by simp + } + ultimately show ?thesis by blast + qed +next + case (LAnd1 b a1 M1 a2 M2 N z u) + then show ?case + proof - + { assume asm: "M1\Ax y a1" + have "(Cut .AndR .M1 .M2 b (z).AndL1 (u).N z){y:=.P} = + Cut .AndR .(M1{y:=.P}) .(M2{y:=.P}) b (z).AndL1 (u).(N{y:=.P}) z" + using prems by (simp add: subst_fresh abs_fresh fresh_atm) + also have "\ \\<^isub>a* Cut .(M1{y:=.P}) (u).(N{y:=.P})" + using prems + apply - + apply(rule a_starI) + apply(rule al_redu) + apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh) + done + also have "\ = (Cut .M1 (u).N){y:=.P}" using prems + by (simp add: subst_fresh abs_fresh fresh_atm) + finally + have "(Cut .AndR .M1 .M2 b (z).AndL1 (u).N z){y:=.P} \\<^isub>a* (Cut .M1 (u).N){y:=.P}" + by simp + } + moreover + { assume asm: "M1=Ax y a1" + have "(Cut .AndR .M1 .M2 b (z).AndL1 (u).N z){y:=.P} = + Cut .AndR .(M1{y:=.P}) .(M2{y:=.P}) b (z).AndL1 (u).(N{y:=.P}) z" + using prems by (simp add: subst_fresh abs_fresh fresh_atm) + also have "\ \\<^isub>a* Cut .(M1{y:=.P}) (u).(N{y:=.P})" + using prems + apply - + apply(rule a_starI) + apply(rule al_redu) + apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh) + done + also have "\ = Cut .(Cut .P (y). Ax y a1) (u).(N{y:=.P})" + using prems by simp + also have "\ \\<^isub>a* Cut .P[c\c>a1] (u).(N{y:=.P})" + proof (cases "fic P c") + case True + assume "fic P c" + then show ?thesis using prems + apply - + apply(rule a_starI) + apply(rule better_CutL_intro) + apply(rule al_redu) + apply(rule better_LAxR_intro) + apply(simp) + done + next + case False + assume "\fic P c" + then show ?thesis + apply - + apply(rule a_star_CutL) + apply(rule a_star_trans) + apply(rule a_starI) + apply(rule ac_redu) + apply(rule better_left) + apply(simp) + apply(simp add: subst_with_ax2) + done + qed + also have "\ = (Cut .M1 (u).N){y:=.P}" using prems + apply - + apply(auto simp add: subst_fresh abs_fresh) + apply(simp add: trm.inject) + apply(simp add: alpha fresh_atm) + apply(rule sym) + apply(rule crename_swap) + apply(simp) + done + finally + have "(Cut .AndR .M1 .M2 b (z).AndL1 (u).N z){y:=.P} \\<^isub>a* (Cut .M1 (u).N){y:=.P}" + by simp + } + ultimately show ?thesis by blast + qed +next + case (LAnd2 b a1 M1 a2 M2 N z u) + then show ?case + proof - + { assume asm: "M2\Ax y a2" + have "(Cut .AndR .M1 .M2 b (z).AndL2 (u).N z){y:=.P} = + Cut .AndR .(M1{y:=.P}) .(M2{y:=.P}) b (z).AndL2 (u).(N{y:=.P}) z" + using prems by (simp add: subst_fresh abs_fresh fresh_atm) + also have "\ \\<^isub>a* Cut .(M2{y:=.P}) (u).(N{y:=.P})" + using prems + apply - + apply(rule a_starI) + apply(rule al_redu) + apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh) + done + also have "\ = (Cut .M2 (u).N){y:=.P}" using prems + by (simp add: subst_fresh abs_fresh fresh_atm) + finally + have "(Cut .AndR .M1 .M2 b (z).AndL2 (u).N z){y:=.P} \\<^isub>a* (Cut .M2 (u).N){y:=.P}" + by simp + } + moreover + { assume asm: "M2=Ax y a2" + have "(Cut .AndR .M1 .M2 b (z).AndL2 (u).N z){y:=.P} = + Cut .AndR .(M1{y:=.P}) .(M2{y:=.P}) b (z).AndL2 (u).(N{y:=.P}) z" + using prems by (simp add: subst_fresh abs_fresh fresh_atm) + also have "\ \\<^isub>a* Cut .(M2{y:=.P}) (u).(N{y:=.P})" + using prems + apply - + apply(rule a_starI) + apply(rule al_redu) + apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh) + done + also have "\ = Cut .(Cut .P (y). Ax y a2) (u).(N{y:=.P})" + using prems by simp + also have "\ \\<^isub>a* Cut .P[c\c>a2] (u).(N{y:=.P})" + proof (cases "fic P c") + case True + assume "fic P c" + then show ?thesis using prems + apply - + apply(rule a_starI) + apply(rule better_CutL_intro) + apply(rule al_redu) + apply(rule better_LAxR_intro) + apply(simp) + done + next + case False + assume "\fic P c" + then show ?thesis + apply - + apply(rule a_star_CutL) + apply(rule a_star_trans) + apply(rule a_starI) + apply(rule ac_redu) + apply(rule better_left) + apply(simp) + apply(simp add: subst_with_ax2) + done + qed + also have "\ = (Cut .M2 (u).N){y:=.P}" using prems + apply - + apply(auto simp add: subst_fresh abs_fresh) + apply(simp add: trm.inject) + apply(simp add: alpha fresh_atm) + apply(rule sym) + apply(rule crename_swap) + apply(simp) + done + finally + have "(Cut .AndR .M1 .M2 b (z).AndL2 (u).N z){y:=.P} \\<^isub>a* (Cut .M2 (u).N){y:=.P}" + by simp + } + ultimately show ?thesis by blast + qed +next + case (LOr1 b a M N1 N2 z x1 x2 y c P) + then show ?case + proof - + { assume asm: "M\Ax y a" + have "(Cut .OrR1 .M b (z).OrL (x1).N1 (x2).N2 z){y:=.P} = + Cut .OrR1 .(M{y:=.P}) b (z).OrL (x1).(N1{y:=.P}) (x2).(N2{y:=.P}) z" + using prems by (simp add: subst_fresh abs_fresh fresh_atm) + also have "\ \\<^isub>a* Cut .(M{y:=.P}) (x1).(N1{y:=.P})" + using prems + apply - + apply(rule a_starI) + apply(rule al_redu) + apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh) + done + also have "\ = (Cut .M (x1).N1){y:=.P}" using prems + by (simp add: subst_fresh abs_fresh fresh_atm) + finally + have "(Cut .OrR1 .M b (z).OrL (x1).N1 (x2).N2 z){y:=.P} \\<^isub>a* (Cut .M (x1).N1){y:=.P}" + by simp + } + moreover + { assume asm: "M=Ax y a" + have "(Cut .OrR1 .M b (z).OrL (x1).N1 (x2).N2 z){y:=.P} = + Cut .OrR1 .(M{y:=.P}) b (z).OrL (x1).(N1{y:=.P}) (x2).(N2{y:=.P}) z" + using prems by (simp add: subst_fresh abs_fresh fresh_atm) + also have "\ \\<^isub>a* Cut .(M{y:=.P}) (x1).(N1{y:=.P})" + using prems + apply - + apply(rule a_starI) + apply(rule al_redu) + apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh) + done + also have "\ = Cut .(Cut .P (y). Ax y a) (x1).(N1{y:=.P})" + using prems by simp + also have "\ \\<^isub>a* Cut .P[c\c>a] (x1).(N1{y:=.P})" + proof (cases "fic P c") + case True + assume "fic P c" + then show ?thesis using prems + apply - + apply(rule a_starI) + apply(rule better_CutL_intro) + apply(rule al_redu) + apply(rule better_LAxR_intro) + apply(simp) + done + next + case False + assume "\fic P c" + then show ?thesis + apply - + apply(rule a_star_CutL) + apply(rule a_star_trans) + apply(rule a_starI) + apply(rule ac_redu) + apply(rule better_left) + apply(simp) + apply(simp add: subst_with_ax2) + done + qed + also have "\ = (Cut .M (x1).N1){y:=.P}" using prems + apply - + apply(auto simp add: subst_fresh abs_fresh) + apply(simp add: trm.inject) + apply(simp add: alpha fresh_atm) + apply(rule sym) + apply(rule crename_swap) + apply(simp) + done + finally + have "(Cut .OrR1 .M b (z).OrL (x1).N1 (x2).N2 z){y:=.P} \\<^isub>a* (Cut .M (x1).N1){y:=.P}" + by simp + } + ultimately show ?thesis by blast + qed +next + case (LOr2 b a M N1 N2 z x1 x2 y c P) + then show ?case + proof - + { assume asm: "M\Ax y a" + have "(Cut .OrR2 .M b (z).OrL (x1).N1 (x2).N2 z){y:=.P} = + Cut .OrR2 .(M{y:=.P}) b (z).OrL (x1).(N1{y:=.P}) (x2).(N2{y:=.P}) z" + using prems by (simp add: subst_fresh abs_fresh fresh_atm) + also have "\ \\<^isub>a* Cut .(M{y:=.P}) (x2).(N2{y:=.P})" + using prems + apply - + apply(rule a_starI) + apply(rule al_redu) + apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh) + done + also have "\ = (Cut .M (x2).N2){y:=.P}" using prems + by (simp add: subst_fresh abs_fresh fresh_atm) + finally + have "(Cut .OrR2 .M b (z).OrL (x1).N1 (x2).N2 z){y:=.P} \\<^isub>a* (Cut .M (x2).N2){y:=.P}" + by simp + } + moreover + { assume asm: "M=Ax y a" + have "(Cut .OrR2 .M b (z).OrL (x1).N1 (x2).N2 z){y:=.P} = + Cut .OrR2 .(M{y:=.P}) b (z).OrL (x1).(N1{y:=.P}) (x2).(N2{y:=.P}) z" + using prems by (simp add: subst_fresh abs_fresh fresh_atm) + also have "\ \\<^isub>a* Cut .(M{y:=.P}) (x2).(N2{y:=.P})" + using prems + apply - + apply(rule a_starI) + apply(rule al_redu) + apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh) + done + also have "\ = Cut .(Cut .P (y). Ax y a) (x2).(N2{y:=.P})" + using prems by simp + also have "\ \\<^isub>a* Cut .P[c\c>a] (x2).(N2{y:=.P})" + proof (cases "fic P c") + case True + assume "fic P c" + then show ?thesis using prems + apply - + apply(rule a_starI) + apply(rule better_CutL_intro) + apply(rule al_redu) + apply(rule better_LAxR_intro) + apply(simp) + done + next + case False + assume "\fic P c" + then show ?thesis + apply - + apply(rule a_star_CutL) + apply(rule a_star_trans) + apply(rule a_starI) + apply(rule ac_redu) + apply(rule better_left) + apply(simp) + apply(simp add: subst_with_ax2) + done + qed + also have "\ = (Cut .M (x2).N2){y:=.P}" using prems + apply - + apply(auto simp add: subst_fresh abs_fresh) + apply(simp add: trm.inject) + apply(simp add: alpha fresh_atm) + apply(rule sym) + apply(rule crename_swap) + apply(simp) + done + finally + have "(Cut .OrR2 .M b (z).OrL (x1).N1 (x2).N2 z){y:=.P} \\<^isub>a* (Cut .M (x2).N2){y:=.P}" + by simp + } + ultimately show ?thesis by blast + qed +next + case (LImp z N u Q x M b a d y c P) + then show ?case + proof - + { assume asm: "N\Ax y d" + have "(Cut .ImpR (x)..M b (z).ImpL .N (u).Q z){y:=.P} = + Cut .ImpR (x)..(M{y:=.P}) b (z).ImpL .(N{y:=.P}) (u).(Q{y:=.P}) z" + using prems by (simp add: fresh_prod abs_fresh fresh_atm) + also have "\ \\<^isub>a* Cut .(Cut .(N{y:=.P}) (x).(M{y:=.P})) (u).(Q{y:=.P})" + using prems + apply - + apply(rule a_starI) + apply(rule al_redu) + apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh) + done + also have "\ = (Cut .(Cut .N (x).M) (u).Q){y:=.P}" using prems + by (simp add: subst_fresh abs_fresh fresh_atm) + finally + have "(Cut .ImpR (x)..M b (z).ImpL .N (u).Q z){y:=.P} \\<^isub>a* + (Cut .(Cut .N (x).M) (u).Q){y:=.P}" + by simp + } + moreover + { assume asm: "N=Ax y d" + have "(Cut .ImpR (x)..M b (z).ImpL .N (u).Q z){y:=.P} = + Cut .ImpR (x)..(M{y:=.P}) b (z).ImpL .(N{y:=.P}) (u).(Q{y:=.P}) z" + using prems by (simp add: subst_fresh abs_fresh fresh_atm fresh_prod) + also have "\ \\<^isub>a* Cut .(Cut .(N{y:=.P}) (x).(M{y:=.P})) (u).(Q{y:=.P})" + using prems + apply - + apply(rule a_starI) + apply(rule al_redu) + apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh) + done + also have "\ = Cut .(Cut .(Cut .P (y).Ax y d) (x).(M{y:=.P})) (u).(Q{y:=.P})" + using prems by simp + also have "\ \\<^isub>a* Cut .(Cut .(P[c\c>d]) (x).(M{y:=.P})) (u).(Q{y:=.P})" + proof (cases "fic P c") + case True + assume "fic P c" + then show ?thesis using prems + apply - + apply(rule a_starI) + apply(rule better_CutL_intro) + apply(rule a_Cut_l) + apply(simp add: subst_fresh abs_fresh) + apply(simp add: abs_fresh fresh_atm) + apply(rule al_redu) + apply(rule better_LAxR_intro) + apply(simp) + done + next + case False + assume "\fic P c" + then show ?thesis using prems + apply - + apply(rule a_star_CutL) + apply(rule a_star_CutL) + apply(rule a_star_trans) + apply(rule a_starI) + apply(rule ac_redu) + apply(rule better_left) + apply(simp) + apply(simp add: subst_with_ax2) + done + qed + also have "\ = (Cut .(Cut .N (x).M) (u).Q){y:=.P}" using prems + apply - + apply(auto simp add: subst_fresh abs_fresh) + apply(simp add: trm.inject) + apply(simp add: alpha fresh_atm) + apply(simp add: trm.inject) + apply(simp add: alpha) + apply(rule sym) + apply(rule crename_swap) + apply(simp) + done + finally + have "(Cut .ImpR (x)..M b (z).ImpL .N (u).Q z){y:=.P} \\<^isub>a* + (Cut .(Cut .N (x).M) (u).Q){y:=.P}" + by simp + } + ultimately show ?thesis by blast + qed +qed + +lemma l_redu_subst2: + assumes a: "M \\<^isub>l M'" + shows "M{c:=(y).P} \\<^isub>a* M'{c:=(y).P}" +using a +proof(nominal_induct M M' avoiding: y c P rule: l_redu.strong_induct) + case LAxR + then show ?case + apply - + apply(rule aux3) + apply(rule better_Cut_substc) + apply(simp add: abs_fresh) + apply(simp add: abs_fresh) + apply(simp add: trm.inject fresh_atm) + apply(auto) + apply(rule aux4) + apply(rule sym) + apply(rule fic_substc_crename) + apply(simp_all) + apply(rule a_starI) + apply(rule al_redu) + apply(rule aux2) + apply(rule l_redu.intros) + apply(simp add: subst_fresh) + apply(simp add: fresh_atm) + apply(rule fic_subst1) + apply(simp_all) + apply(rule subst_comm') + apply(simp_all) + done +next + case LAxL + then show ?case + apply - + apply(rule aux3) + apply(rule better_Cut_substc) + apply(simp) + apply(simp add: abs_fresh) + apply(simp add: fresh_atm) + apply(auto) + apply(rule aux4) + apply(simp add: trm.inject alpha calc_atm fresh_atm) + apply(rule a_star_trans) + apply(rule a_starI) + apply(rule al_redu) + apply(rule l_redu.intros) + apply(simp add: subst_fresh) + apply(simp add: fresh_atm) + apply(rule fin_subst2) + apply(simp_all) + apply(rule aux4) + apply(rule subst_comm') + apply(simp_all) + done +next + case (LNot v M N u a b) + then show ?case + proof - + { assume asm: "M\Ax u c" + have "(Cut .NotR (u).M a (v).NotL .N v){c:=(y).P} = + (Cut .NotR (u).(M{c:=(y).P}) a (v).NotL .(N{c:=(y).P}) v)" using prems + by (simp add: subst_fresh abs_fresh fresh_atm) + also have "\ \\<^isub>l (Cut .(N{c:=(y).P}) (u).(M{c:=(y).P}))" using prems + by (auto intro: l_redu.intros simp add: subst_fresh) + also have "\ = (Cut .N (u).M){c:=(y).P}" using prems + by (simp add: subst_fresh abs_fresh fresh_atm) + finally have ?thesis by auto + } + moreover + { assume asm: "M=Ax u c" + have "(Cut .NotR (u).M a (v).NotL .N v){c:=(y).P} = + (Cut .NotR (u).(M{c:=(y).P}) a (v).NotL .(N{c:=(y).P}) v)" using prems + by (simp add: subst_fresh abs_fresh fresh_atm) + also have "\ \\<^isub>a* (Cut .(N{c:=(y).P}) (u).(M{c:=(y).P}))" using prems + apply - + apply(rule a_starI) + apply(rule al_redu) + apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh) + done + also have "\ = (Cut .(N{c:=(y).P}) (u).(Cut .(Ax u c) (y).P))" using prems + by simp + also have "\ \\<^isub>a* (Cut .(N{c:=(y).P}) (u).(P[y\n>u]))" + proof (cases "fin P y") + case True + assume "fin P y" + then show ?thesis using prems + apply - + apply(rule a_starI) + apply(rule better_CutR_intro) + apply(rule al_redu) + apply(rule better_LAxL_intro) + apply(simp) + done + next + case False + assume "\fin P y" + then show ?thesis + apply - + apply(rule a_star_CutR) + apply(rule a_star_trans) + apply(rule a_starI) + apply(rule ac_redu) + apply(rule better_right) + apply(simp) + apply(simp add: subst_with_ax1) + done + qed + also have "\ = (Cut .N (u).M){c:=(y).P}" using prems + apply - + apply(auto simp add: subst_fresh abs_fresh) + apply(simp add: trm.inject) + apply(simp add: alpha fresh_atm) + apply(rule sym) + apply(rule nrename_swap) + apply(simp) + done + finally have "(Cut .NotR (u).M a (v).NotL .N v){c:=(y).P} \\<^isub>a* (Cut .N (u).M){c:=(y).P}" + by simp + } + ultimately show ?thesis by blast + qed +next + case (LAnd1 b a1 M1 a2 M2 N z u) + then show ?case + proof - + { assume asm: "N\Ax u c" + have "(Cut .AndR .M1 .M2 b (z).AndL1 (u).N z){c:=(y).P} = + Cut .AndR .(M1{c:=(y).P}) .(M2{c:=(y).P}) b (z).AndL1 (u).(N{c:=(y).P}) z" + using prems by (simp add: subst_fresh abs_fresh fresh_atm) + also have "\ \\<^isub>a* Cut .(M1{c:=(y).P}) (u).(N{c:=(y).P})" + using prems + apply - + apply(rule a_starI) + apply(rule al_redu) + apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh) + done + also have "\ = (Cut .M1 (u).N){c:=(y).P}" using prems + by (simp add: subst_fresh abs_fresh fresh_atm) + finally + have "(Cut .AndR .M1 .M2 b (z).AndL1 (u).N z){c:=(y).P} \\<^isub>a* (Cut .M1 (u).N){c:=(y).P}" + by simp + } + moreover + { assume asm: "N=Ax u c" + have "(Cut .AndR .M1 .M2 b (z).AndL1 (u).N z){c:=(y).P} = + Cut .AndR .(M1{c:=(y).P}) .(M2{c:=(y).P}) b (z).AndL1 (u).(N{c:=(y).P}) z" + using prems by (simp add: subst_fresh abs_fresh fresh_atm) + also have "\ \\<^isub>a* Cut .(M1{c:=(y).P}) (u).(N{c:=(y).P})" + using prems + apply - + apply(rule a_starI) + apply(rule al_redu) + apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh) + done + also have "\ = Cut .(M1{c:=(y).P}) (u).(Cut .(Ax u c) (y).P)" + using prems by simp + also have "\ \\<^isub>a* Cut .(M1{c:=(y).P}) (u).(P[y\n>u])" + proof (cases "fin P y") + case True + assume "fin P y" + then show ?thesis using prems + apply - + apply(rule a_starI) + apply(rule better_CutR_intro) + apply(rule al_redu) + apply(rule better_LAxL_intro) + apply(simp) + done + next + case False + assume "\fin P y" + then show ?thesis + apply - + apply(rule a_star_CutR) + apply(rule a_star_trans) + apply(rule a_starI) + apply(rule ac_redu) + apply(rule better_right) + apply(simp) + apply(simp add: subst_with_ax1) + done + qed + also have "\ = (Cut .M1 (u).N){c:=(y).P}" using prems + apply - + apply(auto simp add: subst_fresh abs_fresh) + apply(simp add: trm.inject) + apply(simp add: alpha fresh_atm) + apply(rule sym) + apply(rule nrename_swap) + apply(simp) + done + finally + have "(Cut .AndR .M1 .M2 b (z).AndL1 (u).N z){c:=(y).P} \\<^isub>a* (Cut .M1 (u).N){c:=(y).P}" + by simp + } + ultimately show ?thesis by blast + qed +next + case (LAnd2 b a1 M1 a2 M2 N z u) + then show ?case + proof - + { assume asm: "N\Ax u c" + have "(Cut .AndR .M1 .M2 b (z).AndL2 (u).N z){c:=(y).P} = + Cut .AndR .(M1{c:=(y).P}) .(M2{c:=(y).P}) b (z).AndL2 (u).(N{c:=(y).P}) z" + using prems by (simp add: subst_fresh abs_fresh fresh_atm) + also have "\ \\<^isub>a* Cut .(M2{c:=(y).P}) (u).(N{c:=(y).P})" + using prems + apply - + apply(rule a_starI) + apply(rule al_redu) + apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh) + done + also have "\ = (Cut .M2 (u).N){c:=(y).P}" using prems + by (simp add: subst_fresh abs_fresh fresh_atm) + finally + have "(Cut .AndR .M1 .M2 b (z).AndL2 (u).N z){c:=(y).P} \\<^isub>a* (Cut .M2 (u).N){c:=(y).P}" + by simp + } + moreover + { assume asm: "N=Ax u c" + have "(Cut .AndR .M1 .M2 b (z).AndL2 (u).N z){c:=(y).P} = + Cut .AndR .(M1{c:=(y).P}) .(M2{c:=(y).P}) b (z).AndL2 (u).(N{c:=(y).P}) z" + using prems by (simp add: subst_fresh abs_fresh fresh_atm) + also have "\ \\<^isub>a* Cut .(M2{c:=(y).P}) (u).(N{c:=(y).P})" + using prems + apply - + apply(rule a_starI) + apply(rule al_redu) + apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh) + done + also have "\ = Cut .(M2{c:=(y).P}) (u).(Cut .(Ax u c) (y).P)" + using prems by simp + also have "\ \\<^isub>a* Cut .(M2{c:=(y).P}) (u).(P[y\n>u])" + proof (cases "fin P y") + case True + assume "fin P y" + then show ?thesis using prems + apply - + apply(rule a_starI) + apply(rule better_CutR_intro) + apply(rule al_redu) + apply(rule better_LAxL_intro) + apply(simp) + done + next + case False + assume "\fin P y" + then show ?thesis + apply - + apply(rule a_star_CutR) + apply(rule a_star_trans) + apply(rule a_starI) + apply(rule ac_redu) + apply(rule better_right) + apply(simp) + apply(simp add: subst_with_ax1) + done + qed + also have "\ = (Cut .M2 (u).N){c:=(y).P}" using prems + apply - + apply(auto simp add: subst_fresh abs_fresh) + apply(simp add: trm.inject) + apply(simp add: alpha fresh_atm) + apply(rule sym) + apply(rule nrename_swap) + apply(simp) + done + finally + have "(Cut .AndR .M1 .M2 b (z).AndL2 (u).N z){c:=(y).P} \\<^isub>a* (Cut .M2 (u).N){c:=(y).P}" + by simp + } + ultimately show ?thesis by blast + qed +next + case (LOr1 b a M N1 N2 z x1 x2 y c P) + then show ?case + proof - + { assume asm: "N1\Ax x1 c" + have "(Cut .OrR1 .M b (z).OrL (x1).N1 (x2).N2 z){c:=(y).P} = + Cut .OrR1 .(M{c:=(y).P}) b (z).OrL (x1).(N1{c:=(y).P}) (x2).(N2{c:=(y).P}) z" + using prems by (simp add: subst_fresh abs_fresh fresh_atm) + also have "\ \\<^isub>a* Cut .(M{c:=(y).P}) (x1).(N1{c:=(y).P})" + using prems + apply - + apply(rule a_starI) + apply(rule al_redu) + apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh) + done + also have "\ = (Cut .M (x1).N1){c:=(y).P}" using prems + by (simp add: subst_fresh abs_fresh fresh_atm) + finally + have "(Cut .OrR1 .M b (z).OrL (x1).N1 (x2).N2 z){c:=(y).P} \\<^isub>a* (Cut .M (x1).N1){c:=(y).P}" + by simp + } + moreover + { assume asm: "N1=Ax x1 c" + have "(Cut .OrR1 .M b (z).OrL (x1).N1 (x2).N2 z){c:=(y).P} = + Cut .OrR1 .(M{c:=(y).P}) b (z).OrL (x1).(N1{c:=(y).P}) (x2).(N2{c:=(y).P}) z" + using prems by (simp add: subst_fresh abs_fresh fresh_atm) + also have "\ \\<^isub>a* Cut .(M{c:=(y).P}) (x1).(N1{c:=(y).P})" + using prems + apply - + apply(rule a_starI) + apply(rule al_redu) + apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh) + done + also have "\ = Cut .(M{c:=(y).P}) (x1).(Cut .(Ax x1 c) (y).P)" + using prems by simp + also have "\ \\<^isub>a* Cut .(M{c:=(y).P}) (x1).(P[y\n>x1])" + proof (cases "fin P y") + case True + assume "fin P y" + then show ?thesis using prems + apply - + apply(rule a_starI) + apply(rule better_CutR_intro) + apply(rule al_redu) + apply(rule better_LAxL_intro) + apply(simp) + done + next + case False + assume "\fin P y" + then show ?thesis + apply - + apply(rule a_star_CutR) + apply(rule a_star_trans) + apply(rule a_starI) + apply(rule ac_redu) + apply(rule better_right) + apply(simp) + apply(simp add: subst_with_ax1) + done + qed + also have "\ = (Cut .M (x1).N1){c:=(y).P}" using prems + apply - + apply(auto simp add: subst_fresh abs_fresh) + apply(simp add: trm.inject) + apply(simp add: alpha fresh_atm) + apply(rule sym) + apply(rule nrename_swap) + apply(simp) + done + finally + have "(Cut .OrR1 .M b (z).OrL (x1).N1 (x2).N2 z){c:=(y).P} \\<^isub>a* (Cut .M (x1).N1){c:=(y).P}" + by simp + } + ultimately show ?thesis by blast + qed +next + case (LOr2 b a M N1 N2 z x1 x2 y c P) + then show ?case + proof - + { assume asm: "N2\Ax x2 c" + have "(Cut .OrR2 .M b (z).OrL (x1).N1 (x2).N2 z){c:=(y).P} = + Cut .OrR2 .(M{c:=(y).P}) b (z).OrL (x1).(N1{c:=(y).P}) (x2).(N2{c:=(y).P}) z" + using prems by (simp add: subst_fresh abs_fresh fresh_atm) + also have "\ \\<^isub>a* Cut .(M{c:=(y).P}) (x2).(N2{c:=(y).P})" + using prems + apply - + apply(rule a_starI) + apply(rule al_redu) + apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh) + done + also have "\ = (Cut .M (x2).N2){c:=(y).P}" using prems + by (simp add: subst_fresh abs_fresh fresh_atm) + finally + have "(Cut .OrR2 .M b (z).OrL (x1).N1 (x2).N2 z){c:=(y).P} \\<^isub>a* (Cut .M (x2).N2){c:=(y).P}" + by simp + } + moreover + { assume asm: "N2=Ax x2 c" + have "(Cut .OrR2 .M b (z).OrL (x1).N1 (x2).N2 z){c:=(y).P} = + Cut .OrR2 .(M{c:=(y).P}) b (z).OrL (x1).(N1{c:=(y).P}) (x2).(N2{c:=(y).P}) z" + using prems by (simp add: subst_fresh abs_fresh fresh_atm) + also have "\ \\<^isub>a* Cut .(M{c:=(y).P}) (x2).(N2{c:=(y).P})" + using prems + apply - + apply(rule a_starI) + apply(rule al_redu) + apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh) + done + also have "\ = Cut .(M{c:=(y).P}) (x2).(Cut .(Ax x2 c) (y).P)" + using prems by simp + also have "\ \\<^isub>a* Cut .(M{c:=(y).P}) (x2).(P[y\n>x2])" + proof (cases "fin P y") + case True + assume "fin P y" + then show ?thesis using prems + apply - + apply(rule a_starI) + apply(rule better_CutR_intro) + apply(rule al_redu) + apply(rule better_LAxL_intro) + apply(simp) + done + next + case False + assume "\fin P y" + then show ?thesis + apply - + apply(rule a_star_CutR) + apply(rule a_star_trans) + apply(rule a_starI) + apply(rule ac_redu) + apply(rule better_right) + apply(simp) + apply(simp add: subst_with_ax1) + done + qed + also have "\ = (Cut .M (x2).N2){c:=(y).P}" using prems + apply - + apply(auto simp add: subst_fresh abs_fresh) + apply(simp add: trm.inject) + apply(simp add: alpha fresh_atm) + apply(rule sym) + apply(rule nrename_swap) + apply(simp) + done + finally + have "(Cut .OrR2 .M b (z).OrL (x1).N1 (x2).N2 z){c:=(y).P} \\<^isub>a* (Cut .M (x2).N2){c:=(y).P}" + by simp + } + ultimately show ?thesis by blast + qed +next + case (LImp z N u Q x M b a d y c P) + then show ?case + proof - + { assume asm: "M\Ax x c \ Q\Ax u c" + have "(Cut .ImpR (x)..M b (z).ImpL .N (u).Q z){c:=(y).P} = + Cut .ImpR (x)..(M{c:=(y).P}) b (z).ImpL .(N{c:=(y).P}) (u).(Q{c:=(y).P}) z" + using prems by (simp add: fresh_prod abs_fresh fresh_atm) + also have "\ \\<^isub>a* Cut .(Cut .(N{c:=(y).P}) (x).(M{c:=(y).P})) (u).(Q{c:=(y).P})" + using prems + apply - + apply(rule a_starI) + apply(rule al_redu) + apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh) + done + also have "\ = (Cut .(Cut .N (x).M) (u).Q){c:=(y).P}" using prems + by (simp add: subst_fresh abs_fresh fresh_atm) + finally + have "(Cut .ImpR (x)..M b (z).ImpL .N (u).Q z){c:=(y).P} \\<^isub>a* + (Cut .(Cut .N (x).M) (u).Q){c:=(y).P}" + by simp + } + moreover + { assume asm: "M=Ax x c \ Q\Ax u c" + have "(Cut .ImpR (x)..M b (z).ImpL .N (u).Q z){c:=(y).P} = + Cut .ImpR (x)..(M{c:=(y).P}) b (z).ImpL .(N{c:=(y).P}) (u).(Q{c:=(y).P}) z" + using prems by (simp add: subst_fresh abs_fresh fresh_atm fresh_prod) + also have "\ \\<^isub>a* Cut .(Cut .(N{c:=(y).P}) (x).(M{c:=(y).P})) (u).(Q{c:=(y).P})" + using prems + apply - + apply(rule a_starI) + apply(rule al_redu) + apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh) + done + also have "\ = Cut .(Cut .(N{c:=(y).P}) (x).(Cut .Ax x c (y).P)) (u).(Q{c:=(y).P})" + using prems by simp + also have "\ \\<^isub>a* Cut .(Cut .(N{c:=(y).P}) (x).(P[y\n>x])) (u).(Q{c:=(y).P})" + proof (cases "fin P y") + case True + assume "fin P y" + then show ?thesis using prems + apply - + apply(rule a_star_CutL) + apply(rule a_star_CutR) + apply(rule a_star_trans) + apply(rule a_starI) + apply(rule al_redu) + apply(rule better_LAxL_intro) + apply(simp) + apply(simp) + done + next + case False + assume "\fin P y" + then show ?thesis using prems + apply - + apply(rule a_star_CutL) + apply(rule a_star_CutR) + apply(rule a_star_trans) + apply(rule a_starI) + apply(rule ac_redu) + apply(rule better_right) + apply(simp) + apply(simp add: subst_with_ax1) + done + qed + also have "\ = (Cut .(Cut .N (x).M) (u).Q){c:=(y).P}" using prems + apply - + apply(auto simp add: subst_fresh abs_fresh) + apply(simp add: trm.inject) + apply(simp add: alpha fresh_atm) + apply(simp add: trm.inject) + apply(simp add: alpha) + apply(simp add: nrename_swap) + done + finally + have "(Cut .ImpR (x)..M b (z).ImpL .N (u).Q z){c:=(y).P} \\<^isub>a* + (Cut .(Cut .N (x).M) (u).Q){c:=(y).P}" + by simp + } + moreover + { assume asm: "M\Ax x c \ Q=Ax u c" + have "(Cut .ImpR (x)..M b (z).ImpL .N (u).Q z){c:=(y).P} = + Cut .ImpR (x)..(M{c:=(y).P}) b (z).ImpL .(N{c:=(y).P}) (u).(Q{c:=(y).P}) z" + using prems by (simp add: subst_fresh abs_fresh fresh_atm fresh_prod) + also have "\ \\<^isub>a* Cut .(Cut .(N{c:=(y).P}) (x).(M{c:=(y).P})) (u).(Q{c:=(y).P})" + using prems + apply - + apply(rule a_starI) + apply(rule al_redu) + apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh) + done + also have "\ = Cut .(Cut .(N{c:=(y).P}) (x).(M{c:=(y).P})) (u).(Cut .Ax u c (y).P)" + using prems by simp + also have "\ \\<^isub>a* Cut .(Cut .(N{c:=(y).P}) (x).(M{c:=(y).P})) (u).(P[y\n>u])" + proof (cases "fin P y") + case True + assume "fin P y" + then show ?thesis using prems + apply - + apply(rule a_star_CutR) + apply(rule a_starI) + apply(rule al_redu) + apply(rule better_LAxL_intro) + apply(simp) + done + next + case False + assume "\fin P y" + then show ?thesis using prems + apply - + apply(rule a_star_CutR) + apply(rule a_star_trans) + apply(rule a_starI) + apply(rule ac_redu) + apply(rule better_right) + apply(simp) + apply(simp add: subst_with_ax1) + done + qed + also have "\ = (Cut .(Cut .N (x).M) (u).Q){c:=(y).P}" using prems + apply - + apply(auto simp add: subst_fresh abs_fresh) + apply(simp add: trm.inject) + apply(simp add: alpha fresh_atm) + apply(simp add: nrename_swap) + done + finally + have "(Cut .ImpR (x)..M b (z).ImpL .N (u).Q z){c:=(y).P} \\<^isub>a* + (Cut .(Cut .N (x).M) (u).Q){c:=(y).P}" + by simp + } + moreover + { assume asm: "M=Ax x c \ Q=Ax u c" + have "(Cut .ImpR (x)..M b (z).ImpL .N (u).Q z){c:=(y).P} = + Cut .ImpR (x)..(M{c:=(y).P}) b (z).ImpL .(N{c:=(y).P}) (u).(Q{c:=(y).P}) z" + using prems by (simp add: subst_fresh abs_fresh fresh_atm fresh_prod) + also have "\ \\<^isub>a* Cut .(Cut .(N{c:=(y).P}) (x).(M{c:=(y).P})) (u).(Q{c:=(y).P})" + using prems + apply - + apply(rule a_starI) + apply(rule al_redu) + apply(auto intro: l_redu.intros simp add: subst_fresh abs_fresh) + done + also have "\ = Cut .(Cut .(N{c:=(y).P}) (x).(Cut .Ax x c (y).P)) (u).(Cut .Ax u c (y).P)" + using prems by simp + also have "\ \\<^isub>a* Cut .(Cut .(N{c:=(y).P}) (x).(Cut .Ax x c (y).P)) (u).(P[y\n>u])" + proof (cases "fin P y") + case True + assume "fin P y" + then show ?thesis using prems + apply - + apply(rule a_star_CutR) + apply(rule a_starI) + apply(rule al_redu) + apply(rule better_LAxL_intro) + apply(simp) + done + next + case False + assume "\fin P y" + then show ?thesis using prems + apply - + apply(rule a_star_CutR) + apply(rule a_star_trans) + apply(rule a_starI) + apply(rule ac_redu) + apply(rule better_right) + apply(simp) + apply(simp add: subst_with_ax1) + done + qed + also have "\ \\<^isub>a* Cut .(Cut .(N{c:=(y).P}) (x).(P[y\n>x])) (u).(P[y\n>u])" + proof (cases "fin P y") + case True + assume "fin P y" + then show ?thesis using prems + apply - + apply(rule a_star_CutL) + apply(rule a_star_CutR) + apply(rule a_starI) + apply(rule al_redu) + apply(rule better_LAxL_intro) + apply(simp) + done + next + case False + assume "\fin P y" + then show ?thesis using prems + apply - + apply(rule a_star_CutL) + apply(rule a_star_CutR) + apply(rule a_star_trans) + apply(rule a_starI) + apply(rule ac_redu) + apply(rule better_right) + apply(simp) + apply(simp add: subst_with_ax1) + done + qed + also have "\ = (Cut .(Cut .N (x).M) (u).Q){c:=(y).P}" using prems + apply - + apply(auto simp add: subst_fresh abs_fresh) + apply(simp add: trm.inject) + apply(rule conjI) + apply(simp add: alpha fresh_atm trm.inject) + apply(simp add: nrename_swap) + apply(simp add: alpha fresh_atm trm.inject) + apply(simp add: nrename_swap) + done + finally + have "(Cut .ImpR (x)..M b (z).ImpL .N (u).Q z){c:=(y).P} \\<^isub>a* + (Cut .(Cut .N (x).M) (u).Q){c:=(y).P}" + by simp + } + ultimately show ?thesis by blast + qed +qed + +lemma a_redu_subst1: + assumes a: "M \\<^isub>a M'" + shows "M{y:=.P} \\<^isub>a* M'{y:=.P}" +using a +proof(nominal_induct avoiding: y c P rule: a_redu.strong_induct) + case al_redu + then show ?case by (simp only: l_redu_subst1) +next + case ac_redu + then show ?case + apply - + apply(rule a_starI) + apply(rule a_redu.ac_redu) + apply(simp only: c_redu_subst1') + done +next + case (a_Cut_l a N x M M' y c P) + then show ?case + apply(simp add: subst_fresh fresh_a_redu) + apply(rule conjI) + apply(rule impI)+ + apply(simp) + apply(drule ax_do_not_a_reduce) + apply(simp) + apply(rule impI) + apply(rule conjI) + apply(rule impI) + apply(simp) + apply(drule_tac x="y" in meta_spec) + apply(drule_tac x="c" in meta_spec) + apply(drule_tac x="P" in meta_spec) + apply(simp) + apply(rule a_star_trans) + apply(rule a_star_CutL) + apply(assumption) + apply(rule a_star_trans) + apply(rule_tac M'="P[c\c>a]" in a_star_CutL) + apply(case_tac "fic P c") + apply(rule a_starI) + apply(rule al_redu) + apply(rule better_LAxR_intro) + apply(simp) + apply(rule a_star_trans) + apply(rule a_starI) + apply(rule ac_redu) + apply(rule better_left) + apply(simp) + apply(rule subst_with_ax2) + apply(rule aux4) + apply(simp add: trm.inject) + apply(simp add: alpha fresh_atm) + apply(simp add: crename_swap) + apply(rule impI) + apply(rule a_star_CutL) + apply(auto) + done +next + case (a_Cut_r a N x M M' y c P) + then show ?case + apply(auto simp add: subst_fresh fresh_a_redu) + apply(rule a_star_CutR) + apply(auto)[1] + apply(rule a_star_CutR) + apply(auto)[1] + done +next + case a_NotL + then show ?case + apply(auto) + apply(generate_fresh "name") + apply(fresh_fun_simp) + apply(fresh_fun_simp) + apply(simp add: subst_fresh) + apply(rule a_star_CutR) + apply(rule a_star_NotL) + apply(auto)[1] + apply(rule a_star_NotL) + apply(auto)[1] + done +next + case a_NotR + then show ?case + apply(auto) + apply(rule a_star_NotR) + apply(auto)[1] + done +next + case a_AndR_l + then show ?case + apply(auto simp add: subst_fresh fresh_a_redu) + apply(rule a_star_AndR) + apply(auto) + done +next + case a_AndR_r + then show ?case + apply(auto simp add: subst_fresh fresh_a_redu) + apply(rule a_star_AndR) + apply(auto) + done +next + case a_AndL1 + then show ?case + apply(auto) + apply(generate_fresh "name") + apply(fresh_fun_simp) + apply(fresh_fun_simp) + apply(simp add: subst_fresh) + apply(rule a_star_CutR) + apply(rule a_star_AndL1) + apply(auto)[1] + apply(rule a_star_AndL1) + apply(auto)[1] + done +next + case a_AndL2 + then show ?case + apply(auto) + apply(generate_fresh "name") + apply(fresh_fun_simp) + apply(fresh_fun_simp) + apply(simp add: subst_fresh) + apply(rule a_star_CutR) + apply(rule a_star_AndL2) + apply(auto)[1] + apply(rule a_star_AndL2) + apply(auto)[1] + done +next + case a_OrR1 + then show ?case + apply(auto simp add: subst_fresh fresh_a_redu) + apply(rule a_star_OrR1) + apply(auto) + done +next + case a_OrR2 + then show ?case + apply(auto simp add: subst_fresh fresh_a_redu) + apply(rule a_star_OrR2) + apply(auto) + done +next + case a_OrL_l + then show ?case + apply(auto simp add: subst_fresh fresh_a_redu) + apply(generate_fresh "name") + apply(fresh_fun_simp) + apply(fresh_fun_simp) + apply(simp add: subst_fresh) + apply(rule a_star_CutR) + apply(rule a_star_OrL) + apply(auto) + apply(rule a_star_OrL) + apply(auto) + done +next + case a_OrL_r + then show ?case + apply(auto simp add: subst_fresh fresh_a_redu) + apply(generate_fresh "name") + apply(fresh_fun_simp) + apply(fresh_fun_simp) + apply(simp add: subst_fresh) + apply(rule a_star_CutR) + apply(rule a_star_OrL) + apply(auto) + apply(rule a_star_OrL) + apply(auto) + done +next + case a_ImpR + then show ?case + apply(auto simp add: subst_fresh fresh_a_redu) + apply(rule a_star_ImpR) + apply(auto) + done +next + case a_ImpL_r + then show ?case + apply(auto simp add: subst_fresh fresh_a_redu) + apply(generate_fresh "name") + apply(fresh_fun_simp) + apply(fresh_fun_simp) + apply(simp add: subst_fresh) + apply(rule a_star_CutR) + apply(rule a_star_ImpL) + apply(auto) + apply(rule a_star_ImpL) + apply(auto) + done +next + case a_ImpL_l + then show ?case + apply(auto simp add: subst_fresh fresh_a_redu) + apply(generate_fresh "name") + apply(fresh_fun_simp) + apply(fresh_fun_simp) + apply(simp add: subst_fresh) + apply(rule a_star_CutR) + apply(rule a_star_ImpL) + apply(auto) + apply(rule a_star_ImpL) + apply(auto) + done +qed + +lemma a_redu_subst2: + assumes a: "M \\<^isub>a M'" + shows "M{c:=(y).P} \\<^isub>a* M'{c:=(y).P}" +using a +proof(nominal_induct avoiding: y c P rule: a_redu.strong_induct) + case al_redu + then show ?case by (simp only: l_redu_subst2) +next + case ac_redu + then show ?case + apply - + apply(rule a_starI) + apply(rule a_redu.ac_redu) + apply(simp only: c_redu_subst2') + done +next + case (a_Cut_r a N x M M' y c P) + then show ?case + apply(simp add: subst_fresh fresh_a_redu) + apply(rule conjI) + apply(rule impI)+ + apply(simp) + apply(drule ax_do_not_a_reduce) + apply(simp) + apply(rule impI) + apply(rule conjI) + apply(rule impI) + apply(simp) + apply(drule_tac x="c" in meta_spec) + apply(drule_tac x="y" in meta_spec) + apply(drule_tac x="P" in meta_spec) + apply(simp) + apply(rule a_star_trans) + apply(rule a_star_CutR) + apply(assumption) + apply(rule a_star_trans) + apply(rule_tac N'="P[y\n>x]" in a_star_CutR) + apply(case_tac "fin P y") + apply(rule a_starI) + apply(rule al_redu) + apply(rule better_LAxL_intro) + apply(simp) + apply(rule a_star_trans) + apply(rule a_starI) + apply(rule ac_redu) + apply(rule better_right) + apply(simp) + apply(rule subst_with_ax1) + apply(rule aux4) + apply(simp add: trm.inject) + apply(simp add: alpha fresh_atm) + apply(simp add: nrename_swap) + apply(rule impI) + apply(rule a_star_CutR) + apply(auto) + done +next + case (a_Cut_l a N x M M' y c P) + then show ?case + apply(auto simp add: subst_fresh fresh_a_redu) + apply(rule a_star_CutL) + apply(auto)[1] + apply(rule a_star_CutL) + apply(auto)[1] + done +next + case a_NotR + then show ?case + apply(auto) + apply(generate_fresh "coname") + apply(fresh_fun_simp) + apply(fresh_fun_simp) + apply(simp add: subst_fresh) + apply(rule a_star_CutL) + apply(rule a_star_NotR) + apply(auto)[1] + apply(rule a_star_NotR) + apply(auto)[1] + done +next + case a_NotL + then show ?case + apply(auto) + apply(rule a_star_NotL) + apply(auto)[1] + done +next + case a_AndR_l + then show ?case + apply(auto simp add: subst_fresh fresh_a_redu) + apply(generate_fresh "coname") + apply(fresh_fun_simp) + apply(fresh_fun_simp) + apply(simp add: subst_fresh) + apply(rule a_star_CutL) + apply(rule a_star_AndR) + apply(auto) + apply(rule a_star_AndR) + apply(auto) + done +next + case a_AndR_r + then show ?case + apply(auto simp add: subst_fresh fresh_a_redu) + apply(generate_fresh "coname") + apply(fresh_fun_simp) + apply(fresh_fun_simp) + apply(simp add: subst_fresh) + apply(rule a_star_CutL) + apply(rule a_star_AndR) + apply(auto) + apply(rule a_star_AndR) + apply(auto) + done +next + case a_AndL1 + then show ?case + apply(auto simp add: subst_fresh fresh_a_redu) + apply(rule a_star_AndL1) + apply(auto) + done +next + case a_AndL2 + then show ?case + apply(auto simp add: subst_fresh fresh_a_redu) + apply(rule a_star_AndL2) + apply(auto) + done +next + case a_OrR1 + then show ?case + apply(auto) + apply(generate_fresh "coname") + apply(fresh_fun_simp) + apply(fresh_fun_simp) + apply(simp add: subst_fresh) + apply(rule a_star_CutL) + apply(rule a_star_OrR1) + apply(auto)[1] + apply(rule a_star_OrR1) + apply(auto)[1] + done +next + case a_OrR2 + then show ?case + apply(auto) + apply(generate_fresh "coname") + apply(fresh_fun_simp) + apply(fresh_fun_simp) + apply(simp add: subst_fresh) + apply(rule a_star_CutL) + apply(rule a_star_OrR2) + apply(auto)[1] + apply(rule a_star_OrR2) + apply(auto)[1] + done +next + case a_OrL_l + then show ?case + apply(auto simp add: subst_fresh fresh_a_redu) + apply(rule a_star_OrL) + apply(auto) + done +next + case a_OrL_r + then show ?case + apply(auto simp add: subst_fresh fresh_a_redu) + apply(rule a_star_OrL) + apply(auto) + done +next + case a_ImpR + then show ?case + apply(auto simp add: subst_fresh fresh_a_redu) + apply(generate_fresh "coname") + apply(fresh_fun_simp) + apply(fresh_fun_simp) + apply(simp add: subst_fresh) + apply(rule a_star_CutL) + apply(rule a_star_ImpR) + apply(auto) + apply(rule a_star_ImpR) + apply(auto) + done +next + case a_ImpL_l + then show ?case + apply(auto simp add: subst_fresh fresh_a_redu) + apply(rule a_star_ImpL) + apply(auto) + done +next + case a_ImpL_r + then show ?case + apply(auto simp add: subst_fresh fresh_a_redu) + apply(rule a_star_ImpL) + apply(auto) + done +qed + +lemma a_star_subst1: + assumes a: "M \\<^isub>a* M'" + shows "M{y:=.P} \\<^isub>a* M'{y:=.P}" +using a +apply(induct) +apply(blast) +apply(drule_tac y="y" and c="c" and P="P" in a_redu_subst1) +apply(auto) +done + +lemma a_star_subst2: + assumes a: "M \\<^isub>a* M'" + shows "M{c:=(y).P} \\<^isub>a* M'{c:=(y).P}" +using a +apply(induct) +apply(blast) +apply(drule_tac y="y" and c="c" and P="P" in a_redu_subst2) +apply(auto) +done + +text {* Candidates and SN *} + +text {* SNa *} + +inductive2 + SNa :: "trm \ bool" +where + SNaI: "(\M'. M \\<^isub>a M' \ SNa M') \ SNa M" + +lemma SNa_induct[consumes 1]: + assumes major: "SNa M" + assumes hyp: "\M'. SNa M' \ (\M''. M'\\<^isub>a M'' \ P M'' \ P M')" + shows "P M" + apply (rule major[THEN SNa.induct]) + apply (rule hyp) + apply (rule SNaI) + apply (blast)+ + done + + +lemma double_SNa_aux: + assumes a_SNa: "SNa a" + and b_SNa: "SNa b" + and hyp: "\x z. + (\y. x\\<^isub>a y \ SNa y) \ + (\y. x\\<^isub>a y \ P y z) \ + (\u. z\\<^isub>a u \ SNa u) \ + (\u. z\\<^isub>a u \ P x u) \ P x z" + shows "P a b" +proof - + from a_SNa + have r: "\b. SNa b \ P a b" + proof (induct a rule: SNa.induct) + case (SNaI x) + note SNa' = this + have "SNa b" . + thus ?case + proof (induct b rule: SNa.induct) + case (SNaI y) + show ?case + apply (rule hyp) + apply (erule SNa') + apply (erule SNa') + apply (rule SNa.SNaI) + apply (erule SNaI)+ + done + qed + qed + from b_SNa show ?thesis by (rule r) +qed + +lemma double_SNa: + "\SNa a; SNa b; \x z. ((\y. x\\<^isub>ay \ P y z) \ (\u. z\\<^isub>a u \ P x u)) \ P x z\ \ P a b" +apply(rule_tac double_SNa_aux) +apply(assumption)+ +apply(blast) +done + +lemma a_preserves_SNa: + assumes a: "SNa M" "M\\<^isub>a M'" + shows "SNa M'" +using a +by (erule_tac SNa.cases) (simp) + +lemma a_star_preserves_SNa: + assumes a: "SNa M" and b: "M\\<^isub>a* M'" + shows "SNa M'" +using b a +by (induct) (auto simp add: a_preserves_SNa) + +lemma Ax_in_SNa: + shows "SNa (Ax x a)" +apply(rule SNaI) +apply(erule a_redu.cases, auto) +apply(erule l_redu.cases, auto) +apply(erule c_redu.cases, auto) +done + +lemma NotL_in_SNa: + assumes a: "SNa M" + shows "SNa (NotL .M x)" +using a +apply(induct) +apply(rule SNaI) +apply(erule a_redu.cases, auto) +apply(erule l_redu.cases, auto) +apply(erule c_redu.cases, auto) +apply(auto simp add: trm.inject alpha) +apply(rotate_tac 1) +apply(drule_tac x="[(a,aa)]\M'a" in meta_spec) +apply(simp add: a_redu.eqvt) +apply(subgoal_tac "NotL .([(a,aa)]\M'a) x = NotL .M'a x") +apply(simp) +apply(simp add: trm.inject alpha fresh_a_redu) +done + +lemma NotR_in_SNa: + assumes a: "SNa M" + shows "SNa (NotR (x).M a)" +using a +apply(induct) +apply(rule SNaI) +apply(erule a_redu.cases, auto) +apply(erule l_redu.cases, auto) +apply(erule c_redu.cases, auto) +apply(auto simp add: trm.inject alpha) +apply(rotate_tac 1) +apply(drule_tac x="[(x,xa)]\M'a" in meta_spec) +apply(simp add: a_redu.eqvt) +apply(rule_tac s="NotR (x).([(x,xa)]\M'a) a" in subst) +apply(simp add: trm.inject alpha fresh_a_redu) +apply(simp) +done + +lemma AndL1_in_SNa: + assumes a: "SNa M" + shows "SNa (AndL1 (x).M y)" +using a +apply(induct) +apply(rule SNaI) +apply(erule a_redu.cases, auto) +apply(erule l_redu.cases, auto) +apply(erule c_redu.cases, auto) +apply(auto simp add: trm.inject alpha) +apply(rotate_tac 1) +apply(drule_tac x="[(x,xa)]\M'a" in meta_spec) +apply(simp add: a_redu.eqvt) +apply(rule_tac s="AndL1 x.([(x,xa)]\M'a) y" in subst) +apply(simp add: trm.inject alpha fresh_a_redu) +apply(simp) +done + +lemma AndL2_in_SNa: + assumes a: "SNa M" + shows "SNa (AndL2 (x).M y)" +using a +apply(induct) +apply(rule SNaI) +apply(erule a_redu.cases, auto) +apply(erule l_redu.cases, auto) +apply(erule c_redu.cases, auto) +apply(auto simp add: trm.inject alpha) +apply(rotate_tac 1) +apply(drule_tac x="[(x,xa)]\M'a" in meta_spec) +apply(simp add: a_redu.eqvt) +apply(rule_tac s="AndL2 x.([(x,xa)]\M'a) y" in subst) +apply(simp add: trm.inject alpha fresh_a_redu) +apply(simp) +done + +lemma OrR1_in_SNa: + assumes a: "SNa M" + shows "SNa (OrR1 .M b)" +using a +apply(induct) +apply(rule SNaI) +apply(erule a_redu.cases, auto) +apply(erule l_redu.cases, auto) +apply(erule c_redu.cases, auto) +apply(auto simp add: trm.inject alpha) +apply(rotate_tac 1) +apply(drule_tac x="[(a,aa)]\M'a" in meta_spec) +apply(simp add: a_redu.eqvt) +apply(rule_tac s="OrR1 .([(a,aa)]\M'a) b" in subst) +apply(simp add: trm.inject alpha fresh_a_redu) +apply(simp) +done + +lemma OrR2_in_SNa: + assumes a: "SNa M" + shows "SNa (OrR2 .M b)" +using a +apply(induct) +apply(rule SNaI) +apply(erule a_redu.cases, auto) +apply(erule l_redu.cases, auto) +apply(erule c_redu.cases, auto) +apply(auto simp add: trm.inject alpha) +apply(rotate_tac 1) +apply(drule_tac x="[(a,aa)]\M'a" in meta_spec) +apply(simp add: a_redu.eqvt) +apply(rule_tac s="OrR2 .([(a,aa)]\M'a) b" in subst) +apply(simp add: trm.inject alpha fresh_a_redu) +apply(simp) +done + +lemma ImpR_in_SNa: + assumes a: "SNa M" + shows "SNa (ImpR (x)..M b)" +using a +apply(induct) +apply(rule SNaI) +apply(erule a_redu.cases, auto) +apply(erule l_redu.cases, auto) +apply(erule c_redu.cases, auto) +apply(auto simp add: trm.inject alpha abs_fresh abs_perm calc_atm) +apply(rotate_tac 1) +apply(drule_tac x="[(a,aa)]\M'a" in meta_spec) +apply(simp add: a_redu.eqvt) +apply(rule_tac s="ImpR (x)..([(a,aa)]\M'a) b" in subst) +apply(simp add: trm.inject alpha fresh_a_redu) +apply(simp) +apply(rotate_tac 1) +apply(drule_tac x="[(x,xa)]\M'a" in meta_spec) +apply(simp add: a_redu.eqvt) +apply(rule_tac s="ImpR (x)..([(x,xa)]\M'a) b" in subst) +apply(simp add: trm.inject alpha fresh_a_redu abs_fresh abs_perm calc_atm) +apply(simp) +apply(rotate_tac 1) +apply(drule_tac x="[(a,aa)]\[(x,xa)]\M'a" in meta_spec) +apply(simp add: a_redu.eqvt) +apply(rule_tac s="ImpR (x)..([(a,aa)]\[(x,xa)]\M'a) b" in subst) +apply(simp add: trm.inject alpha fresh_a_redu abs_fresh abs_perm calc_atm) +apply(simp add: fresh_left calc_atm fresh_a_redu) +apply(simp) +done + +lemma AndR_in_SNa: + assumes a: "SNa M" "SNa N" + shows "SNa (AndR .M .N c)" +apply(rule_tac a="M" and b="N" in double_SNa) +apply(rule a)+ +apply(auto) +apply(rule SNaI) +apply(drule a_redu_AndR_elim) +apply(auto) +done + +lemma OrL_in_SNa: + assumes a: "SNa M" "SNa N" + shows "SNa (OrL (x).M (y).N z)" +apply(rule_tac a="M" and b="N" in double_SNa) +apply(rule a)+ +apply(auto) +apply(rule SNaI) +apply(drule a_redu_OrL_elim) +apply(auto) +done + +lemma ImpL_in_SNa: + assumes a: "SNa M" "SNa N" + shows "SNa (ImpL .M (y).N z)" +apply(rule_tac a="M" and b="N" in double_SNa) +apply(rule a)+ +apply(auto) +apply(rule SNaI) +apply(drule a_redu_ImpL_elim) +apply(auto) +done + +lemma SNa_eqvt: + fixes pi1::"name prm" + and pi2::"coname prm" + shows "SNa M \ SNa (pi1\M)" + and "SNa M \ SNa (pi2\M)" +apply - +apply(induct rule: SNa.induct) +apply(rule SNaI) +apply(drule_tac pi="(rev pi1)" in a_redu.eqvt(1)) +apply(rotate_tac 1) +apply(drule_tac x="(rev pi1)\M'" in meta_spec) +apply(perm_simp) +apply(induct rule: SNa.induct) +apply(rule SNaI) +apply(drule_tac pi="(rev pi2)" in a_redu.eqvt(2)) +apply(rotate_tac 1) +apply(drule_tac x="(rev pi2)\M'" in meta_spec) +apply(perm_simp) +done + +text {* set operators *} + +constdefs + AXIOMSn::"ty \ ntrm set" + "AXIOMSn B \ { (x):(Ax y b) | x y b. True }" + + AXIOMSc::"ty \ ctrm set" + "AXIOMSc B \ { :(Ax y b) | a y b. True }" + + BINDINGn::"ty \ ctrm set \ ntrm set" + "BINDINGn B X \ { (x):M | x M. \a P. :P\X \ SNa (M{x:=.P})}" + + BINDINGc::"ty \ ntrm set \ ctrm set" + "BINDINGc B X \ { :M | a M. \x P. (x):P\X \ SNa (M{a:=(x).P})}" + +lemma BINDINGn_decreasing: + shows "X\Y \ BINDINGn B Y \ BINDINGn B X" +by (simp add: BINDINGn_def) (blast) + +lemma BINDINGc_decreasing: + shows "X\Y \ BINDINGc B Y \ BINDINGc B X" +by (simp add: BINDINGc_def) (blast) + +consts + NOTRIGHT::"ty \ ntrm set \ ctrm set" + +nominal_primrec + "NOTRIGHT (NOT B) X = { :NotR (x).M a | a x M. fic (NotR (x).M a) a \ (x):M \ X }" +apply(rule TrueI)+ +done + +lemma NOTRIGHT_eqvt_name: + fixes pi::"name prm" + shows "(pi\(NOTRIGHT (NOT B) X)) = NOTRIGHT (NOT B) (pi\X)" +apply(auto simp add: perm_set_def) +apply(rule_tac x="pi\a" in exI) +apply(rule_tac x="pi\xb" in exI) +apply(rule_tac x="pi\M" in exI) +apply(simp) +apply(rule conjI) +apply(drule_tac pi="pi" in fic.eqvt(1)) +apply(simp) +apply(rule_tac x="(xb):M" in exI) +apply(simp) +apply(rule_tac x="(rev pi)\(:NotR (xa).M a)" in exI) +apply(perm_simp) +apply(rule_tac x="(rev pi)\a" in exI) +apply(rule_tac x="(rev pi)\xa" in exI) +apply(rule_tac x="(rev pi)\M" in exI) +apply(simp) +apply(drule_tac pi="rev pi" in fic.eqvt(1)) +apply(simp) +apply(drule sym) +apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst]) +apply(simp) +done + +lemma NOTRIGHT_eqvt_coname: + fixes pi::"coname prm" + shows "(pi\(NOTRIGHT (NOT B) X)) = NOTRIGHT (NOT B) (pi\X)" +apply(auto simp add: perm_set_def) +apply(rule_tac x="pi\a" in exI) +apply(rule_tac x="pi\xb" in exI) +apply(rule_tac x="pi\M" in exI) +apply(simp) +apply(rule conjI) +apply(drule_tac pi="pi" in fic.eqvt(2)) +apply(simp) +apply(rule_tac x="(xb):M" in exI) +apply(simp) +apply(rule_tac x="<((rev pi)\a)>:NotR ((rev pi)\xa).((rev pi)\M) ((rev pi)\a)" in exI) +apply(perm_simp) +apply(rule_tac x="(rev pi)\a" in exI) +apply(rule_tac x="(rev pi)\xa" in exI) +apply(rule_tac x="(rev pi)\M" in exI) +apply(simp) +apply(drule_tac pi="rev pi" in fic.eqvt(2)) +apply(simp) +apply(drule sym) +apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst]) +apply(simp) +done + +consts + NOTLEFT::"ty \ ctrm set \ ntrm set" + +nominal_primrec + "NOTLEFT (NOT B) X = { (x):NotL .M x | a x M. fin (NotL .M x) x \ :M \ X }" +apply(rule TrueI)+ +done + +lemma NOTLEFT_eqvt_name: + fixes pi::"name prm" + shows "(pi\(NOTLEFT (NOT B) X)) = NOTLEFT (NOT B) (pi\X)" +apply(auto simp add: perm_set_def) +apply(rule_tac x="pi\a" in exI) +apply(rule_tac x="pi\xb" in exI) +apply(rule_tac x="pi\M" in exI) +apply(simp) +apply(rule conjI) +apply(drule_tac pi="pi" in fin.eqvt(1)) +apply(simp) +apply(rule_tac x=":M" in exI) +apply(simp) +apply(rule_tac x="(((rev pi)\xa)):NotL <((rev pi)\a)>.((rev pi)\M) ((rev pi)\xa)" in exI) +apply(perm_simp) +apply(rule_tac x="(rev pi)\a" in exI) +apply(rule_tac x="(rev pi)\xa" in exI) +apply(rule_tac x="(rev pi)\M" in exI) +apply(simp) +apply(drule_tac pi="rev pi" in fin.eqvt(1)) +apply(simp) +apply(drule sym) +apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst]) +apply(simp) +done + +lemma NOTLEFT_eqvt_coname: + fixes pi::"coname prm" + shows "(pi\(NOTLEFT (NOT B) X)) = NOTLEFT (NOT B) (pi\X)" +apply(auto simp add: perm_set_def) +apply(rule_tac x="pi\a" in exI) +apply(rule_tac x="pi\xb" in exI) +apply(rule_tac x="pi\M" in exI) +apply(simp) +apply(rule conjI) +apply(drule_tac pi="pi" in fin.eqvt(2)) +apply(simp) +apply(rule_tac x=":M" in exI) +apply(simp) +apply(rule_tac x="(((rev pi)\xa)):NotL <((rev pi)\a)>.((rev pi)\M) ((rev pi)\xa)" in exI) +apply(perm_simp) +apply(rule_tac x="(rev pi)\a" in exI) +apply(rule_tac x="(rev pi)\xa" in exI) +apply(rule_tac x="(rev pi)\M" in exI) +apply(simp) +apply(drule_tac pi="rev pi" in fin.eqvt(2)) +apply(simp) +apply(drule sym) +apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst]) +apply(simp) +done + +consts + ANDRIGHT::"ty \ ctrm set \ ctrm set \ ctrm set" + +nominal_primrec + "ANDRIGHT (B AND C) X Y = + { :AndR .M .N c | c a b M N. fic (AndR .M .N c) c \ :M \ X \ :N \ Y }" +apply(rule TrueI)+ +done + +lemma ANDRIGHT_eqvt_name: + fixes pi::"name prm" + shows "(pi\(ANDRIGHT (A AND B) X Y)) = ANDRIGHT (A AND B) (pi\X) (pi\Y)" +apply(auto simp add: perm_set_def) +apply(rule_tac x="pi\c" in exI) +apply(rule_tac x="pi\a" in exI) +apply(rule_tac x="pi\b" in exI) +apply(rule_tac x="pi\M" in exI) +apply(rule_tac x="pi\N" in exI) +apply(simp) +apply(rule conjI) +apply(drule_tac pi="pi" in fic.eqvt(1)) +apply(simp) +apply(rule conjI) +apply(rule_tac x=":M" in exI) +apply(simp) +apply(rule_tac x=":N" in exI) +apply(simp) +apply(rule_tac x="(rev pi)\(:AndR .M .N c)" in exI) +apply(perm_simp) +apply(rule_tac x="(rev pi)\c" in exI) +apply(rule_tac x="(rev pi)\a" in exI) +apply(rule_tac x="(rev pi)\b" in exI) +apply(rule_tac x="(rev pi)\M" in exI) +apply(rule_tac x="(rev pi)\N" in exI) +apply(simp) +apply(drule_tac pi="rev pi" in fic.eqvt(1)) +apply(simp) +apply(drule sym) +apply(drule sym) +apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst]) +apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst]) +apply(simp) +done + +lemma ANDRIGHT_eqvt_coname: + fixes pi::"coname prm" + shows "(pi\(ANDRIGHT (A AND B) X Y)) = ANDRIGHT (A AND B) (pi\X) (pi\Y)" +apply(auto simp add: perm_set_def) +apply(rule_tac x="pi\c" in exI) +apply(rule_tac x="pi\a" in exI) +apply(rule_tac x="pi\b" in exI) +apply(rule_tac x="pi\M" in exI) +apply(rule_tac x="pi\N" in exI) +apply(simp) +apply(rule conjI) +apply(drule_tac pi="pi" in fic.eqvt(2)) +apply(simp) +apply(rule conjI) +apply(rule_tac x=":M" in exI) +apply(simp) +apply(rule_tac x=":N" in exI) +apply(simp) +apply(rule_tac x="(rev pi)\(:AndR .M .N c)" in exI) +apply(perm_simp) +apply(rule_tac x="(rev pi)\c" in exI) +apply(rule_tac x="(rev pi)\a" in exI) +apply(rule_tac x="(rev pi)\b" in exI) +apply(rule_tac x="(rev pi)\M" in exI) +apply(rule_tac x="(rev pi)\N" in exI) +apply(simp) +apply(drule_tac pi="rev pi" in fic.eqvt(2)) +apply(simp) +apply(drule sym) +apply(drule sym) +apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst]) +apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst]) +apply(simp) +done + +consts + ANDLEFT1::"ty \ ntrm set \ ntrm set" + +nominal_primrec + "ANDLEFT1 (B AND C) X = { (y):AndL1 (x).M y | x y M. fin (AndL1 (x).M y) y \ (x):M \ X }" +apply(rule TrueI)+ +done + +lemma ANDLEFT1_eqvt_name: + fixes pi::"name prm" + shows "(pi\(ANDLEFT1 (A AND B) X)) = ANDLEFT1 (A AND B) (pi\X)" +apply(auto simp add: perm_set_def) +apply(rule_tac x="pi\xb" in exI) +apply(rule_tac x="pi\y" in exI) +apply(rule_tac x="pi\M" in exI) +apply(simp) +apply(rule conjI) +apply(drule_tac pi="pi" in fin.eqvt(1)) +apply(simp) +apply(rule_tac x="(xb):M" in exI) +apply(simp) +apply(rule_tac x="(rev pi)\((y):AndL1 (xa).M y)" in exI) +apply(perm_simp) +apply(rule_tac x="(rev pi)\xa" in exI) +apply(rule_tac x="(rev pi)\y" in exI) +apply(rule_tac x="(rev pi)\M" in exI) +apply(simp) +apply(drule_tac pi="rev pi" in fin.eqvt(1)) +apply(simp) +apply(drule sym) +apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst]) +apply(simp) +done + +lemma ANDLEFT1_eqvt_coname: + fixes pi::"coname prm" + shows "(pi\(ANDLEFT1 (A AND B) X)) = ANDLEFT1 (A AND B) (pi\X)" +apply(auto simp add: perm_set_def) +apply(rule_tac x="pi\xb" in exI) +apply(rule_tac x="pi\y" in exI) +apply(rule_tac x="pi\M" in exI) +apply(simp) +apply(rule conjI) +apply(drule_tac pi="pi" in fin.eqvt(2)) +apply(simp) +apply(rule_tac x="(xb):M" in exI) +apply(simp) +apply(rule_tac x="(rev pi)\((y):AndL1 (xa).M y)" in exI) +apply(perm_simp) +apply(rule_tac x="(rev pi)\xa" in exI) +apply(rule_tac x="(rev pi)\y" in exI) +apply(rule_tac x="(rev pi)\M" in exI) +apply(simp) +apply(drule_tac pi="rev pi" in fin.eqvt(2)) +apply(simp) +apply(drule sym) +apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst]) +apply(simp) +done + +consts + ANDLEFT2::"ty \ ntrm set \ ntrm set" + +nominal_primrec + "ANDLEFT2 (B AND C) X = { (y):AndL2 (x).M y | x y M. fin (AndL2 (x).M y) y \ (x):M \ X }" +apply(rule TrueI)+ +done + +lemma ANDLEFT2_eqvt_name: + fixes pi::"name prm" + shows "(pi\(ANDLEFT2 (A AND B) X)) = ANDLEFT2 (A AND B) (pi\X)" +apply(auto simp add: perm_set_def) +apply(rule_tac x="pi\xb" in exI) +apply(rule_tac x="pi\y" in exI) +apply(rule_tac x="pi\M" in exI) +apply(simp) +apply(rule conjI) +apply(drule_tac pi="pi" in fin.eqvt(1)) +apply(simp) +apply(rule_tac x="(xb):M" in exI) +apply(simp) +apply(rule_tac x="(rev pi)\((y):AndL2 (xa).M y)" in exI) +apply(perm_simp) +apply(rule_tac x="(rev pi)\xa" in exI) +apply(rule_tac x="(rev pi)\y" in exI) +apply(rule_tac x="(rev pi)\M" in exI) +apply(simp) +apply(drule_tac pi="rev pi" in fin.eqvt(1)) +apply(simp) +apply(drule sym) +apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst]) +apply(simp) +done + +lemma ANDLEFT2_eqvt_coname: + fixes pi::"coname prm" + shows "(pi\(ANDLEFT2 (A AND B) X)) = ANDLEFT2 (A AND B) (pi\X)" +apply(auto simp add: perm_set_def) +apply(rule_tac x="pi\xb" in exI) +apply(rule_tac x="pi\y" in exI) +apply(rule_tac x="pi\M" in exI) +apply(simp) +apply(rule conjI) +apply(drule_tac pi="pi" in fin.eqvt(2)) +apply(simp) +apply(rule_tac x="(xb):M" in exI) +apply(simp) +apply(rule_tac x="(rev pi)\((y):AndL2 (xa).M y)" in exI) +apply(perm_simp) +apply(rule_tac x="(rev pi)\xa" in exI) +apply(rule_tac x="(rev pi)\y" in exI) +apply(rule_tac x="(rev pi)\M" in exI) +apply(simp) +apply(drule_tac pi="rev pi" in fin.eqvt(2)) +apply(simp) +apply(drule sym) +apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst]) +apply(simp) +done + +consts + ORLEFT::"ty \ ntrm set \ ntrm set \ ntrm set" + +nominal_primrec + "ORLEFT (B OR C) X Y = + { (z):OrL (x).M (y).N z | x y z M N. fin (OrL (x).M (y).N z) z \ (x):M \ X \ (y):N \ Y }" +apply(rule TrueI)+ +done + +lemma ORLEFT_eqvt_name: + fixes pi::"name prm" + shows "(pi\(ORLEFT (A OR B) X Y)) = ORLEFT (A OR B) (pi\X) (pi\Y)" +apply(auto simp add: perm_set_def) +apply(rule_tac x="pi\xb" in exI) +apply(rule_tac x="pi\y" in exI) +apply(rule_tac x="pi\z" in exI) +apply(rule_tac x="pi\M" in exI) +apply(rule_tac x="pi\N" in exI) +apply(simp) +apply(rule conjI) +apply(drule_tac pi="pi" in fin.eqvt(1)) +apply(simp) +apply(rule conjI) +apply(rule_tac x="(xb):M" in exI) +apply(simp) +apply(rule_tac x="(y):N" in exI) +apply(simp) +apply(rule_tac x="(rev pi)\((z):OrL (xa).M (y).N z)" in exI) +apply(perm_simp) +apply(rule_tac x="(rev pi)\xa" in exI) +apply(rule_tac x="(rev pi)\y" in exI) +apply(rule_tac x="(rev pi)\z" in exI) +apply(rule_tac x="(rev pi)\M" in exI) +apply(rule_tac x="(rev pi)\N" in exI) +apply(simp) +apply(drule_tac pi="rev pi" in fin.eqvt(1)) +apply(simp) +apply(drule sym) +apply(drule sym) +apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst]) +apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst]) +apply(simp) +done + +lemma ORLEFT_eqvt_coname: + fixes pi::"coname prm" + shows "(pi\(ORLEFT (A OR B) X Y)) = ORLEFT (A OR B) (pi\X) (pi\Y)" +apply(auto simp add: perm_set_def) +apply(rule_tac x="pi\xb" in exI) +apply(rule_tac x="pi\y" in exI) +apply(rule_tac x="pi\z" in exI) +apply(rule_tac x="pi\M" in exI) +apply(rule_tac x="pi\N" in exI) +apply(simp) +apply(rule conjI) +apply(drule_tac pi="pi" in fin.eqvt(2)) +apply(simp) +apply(rule conjI) +apply(rule_tac x="(xb):M" in exI) +apply(simp) +apply(rule_tac x="(y):N" in exI) +apply(simp) +apply(rule_tac x="(rev pi)\((z):OrL (xa).M (y).N z)" in exI) +apply(perm_simp) +apply(rule_tac x="(rev pi)\xa" in exI) +apply(rule_tac x="(rev pi)\y" in exI) +apply(rule_tac x="(rev pi)\z" in exI) +apply(rule_tac x="(rev pi)\M" in exI) +apply(rule_tac x="(rev pi)\N" in exI) +apply(simp) +apply(drule_tac pi="rev pi" in fin.eqvt(2)) +apply(simp) +apply(drule sym) +apply(drule sym) +apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst]) +apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst]) +apply(simp) +done + +consts + ORRIGHT1::"ty \ ctrm set \ ctrm set" + +nominal_primrec + "ORRIGHT1 (B OR C) X = { :OrR1 .M b | a b M. fic (OrR1 .M b) b \ :M \ X }" +apply(rule TrueI)+ +done + +lemma ORRIGHT1_eqvt_name: + fixes pi::"name prm" + shows "(pi\(ORRIGHT1 (A OR B) X)) = ORRIGHT1 (A OR B) (pi\X)" +apply(auto simp add: perm_set_def) +apply(rule_tac x="pi\a" in exI) +apply(rule_tac x="pi\b" in exI) +apply(rule_tac x="pi\M" in exI) +apply(simp) +apply(rule conjI) +apply(drule_tac pi="pi" in fic.eqvt(1)) +apply(simp) +apply(rule_tac x=":M" in exI) +apply(perm_simp) +apply(rule_tac x="(rev pi)\(:OrR1 .M b)" in exI) +apply(perm_simp) +apply(rule_tac x="(rev pi)\a" in exI) +apply(rule_tac x="(rev pi)\b" in exI) +apply(rule_tac x="(rev pi)\M" in exI) +apply(simp) +apply(drule_tac pi="rev pi" in fic.eqvt(1)) +apply(simp) +apply(drule sym) +apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst]) +apply(simp) +done + +lemma ORRIGHT1_eqvt_coname: + fixes pi::"coname prm" + shows "(pi\(ORRIGHT1 (A OR B) X)) = ORRIGHT1 (A OR B) (pi\X)" +apply(auto simp add: perm_set_def) +apply(rule_tac x="pi\a" in exI) +apply(rule_tac x="pi\b" in exI) +apply(rule_tac x="pi\M" in exI) +apply(simp) +apply(rule conjI) +apply(drule_tac pi="pi" in fic.eqvt(2)) +apply(simp) +apply(rule_tac x=":M" in exI) +apply(perm_simp) +apply(rule_tac x="(rev pi)\(:OrR1 .M b)" in exI) +apply(perm_simp) +apply(rule_tac x="(rev pi)\a" in exI) +apply(rule_tac x="(rev pi)\b" in exI) +apply(rule_tac x="(rev pi)\M" in exI) +apply(simp) +apply(drule_tac pi="rev pi" in fic.eqvt(2)) +apply(simp) +apply(drule sym) +apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst]) +apply(simp) +done + +consts + ORRIGHT2::"ty \ ctrm set \ ctrm set" + +nominal_primrec + "ORRIGHT2 (B OR C) X = { :OrR2 .M b | a b M. fic (OrR2 .M b) b \ :M \ X }" +apply(rule TrueI)+ +done + +lemma ORRIGHT2_eqvt_name: + fixes pi::"name prm" + shows "(pi\(ORRIGHT2 (A OR B) X)) = ORRIGHT2 (A OR B) (pi\X)" +apply(auto simp add: perm_set_def) +apply(rule_tac x="pi\a" in exI) +apply(rule_tac x="pi\b" in exI) +apply(rule_tac x="pi\M" in exI) +apply(simp) +apply(rule conjI) +apply(drule_tac pi="pi" in fic.eqvt(1)) +apply(simp) +apply(rule_tac x=":M" in exI) +apply(perm_simp) +apply(rule_tac x="(rev pi)\(:OrR2 .M b)" in exI) +apply(perm_simp) +apply(rule_tac x="(rev pi)\a" in exI) +apply(rule_tac x="(rev pi)\b" in exI) +apply(rule_tac x="(rev pi)\M" in exI) +apply(simp) +apply(drule_tac pi="rev pi" in fic.eqvt(1)) +apply(simp) +apply(drule sym) +apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst]) +apply(simp) +done + +lemma ORRIGHT2_eqvt_coname: + fixes pi::"coname prm" + shows "(pi\(ORRIGHT2 (A OR B) X)) = ORRIGHT2 (A OR B) (pi\X)" +apply(auto simp add: perm_set_def) +apply(rule_tac x="pi\a" in exI) +apply(rule_tac x="pi\b" in exI) +apply(rule_tac x="pi\M" in exI) +apply(simp) +apply(rule conjI) +apply(drule_tac pi="pi" in fic.eqvt(2)) +apply(simp) +apply(rule_tac x=":M" in exI) +apply(perm_simp) +apply(rule_tac x="(rev pi)\(:OrR2 .M b)" in exI) +apply(perm_simp) +apply(rule_tac x="(rev pi)\a" in exI) +apply(rule_tac x="(rev pi)\b" in exI) +apply(rule_tac x="(rev pi)\M" in exI) +apply(simp) +apply(drule_tac pi="rev pi" in fic.eqvt(2)) +apply(simp) +apply(drule sym) +apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst]) +apply(simp) +done + +consts + IMPRIGHT::"ty \ ntrm set \ ctrm set \ ntrm set \ ctrm set \ ctrm set" + +nominal_primrec + "IMPRIGHT (B IMP C) X Y Z U= + { :ImpR (x)..M b | x a b M. fic (ImpR (x)..M b) b + \ (\z P. x\(z,P) \ (z):P \ Z \ (x):(M{a:=(z).P}) \ X) + \ (\c Q. a\(c,Q) \ :Q \ U \ :(M{x:=.Q}) \ Y)}" +apply(rule TrueI)+ +done + +lemma IMPRIGHT_eqvt_name: + fixes pi::"name prm" + shows "(pi\(IMPRIGHT (A IMP B) X Y Z U)) = IMPRIGHT (A IMP B) (pi\X) (pi\Y) (pi\Z) (pi\U)" +apply(auto simp add: perm_set_def) +apply(rule_tac x="pi\xb" in exI) +apply(rule_tac x="pi\a" in exI) +apply(rule_tac x="pi\b" in exI) +apply(rule_tac x="pi\M" in exI) +apply(simp) +apply(rule conjI) +apply(drule_tac pi="pi" in fic.eqvt(1)) +apply(simp) +apply(rule conjI) +apply(auto)[1] +apply(rule_tac x="(xb):(M{a:=((rev pi)\z).((rev pi)\P)})" in exI) +apply(perm_simp add: csubst_eqvt) +apply(drule sym) +apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst]) +apply(simp) +apply(simp add: fresh_right) +apply(auto)[1] +apply(rule_tac x=":(M{xb:=<((rev pi)\c)>.((rev pi)\Q)})" in exI) +apply(perm_simp add: nsubst_eqvt) +apply(drule sym) +apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst]) +apply(simp add: fresh_right) +apply(rule_tac x="(rev pi)\(:ImpR xa..M b)" in exI) +apply(perm_simp) +apply(rule_tac x="(rev pi)\xa" in exI) +apply(rule_tac x="(rev pi)\a" in exI) +apply(rule_tac x="(rev pi)\b" in exI) +apply(rule_tac x="(rev pi)\M" in exI) +apply(simp) +apply(drule_tac pi="rev pi" in fic.eqvt(1)) +apply(simp) +apply(rule conjI) +apply(auto)[1] +apply(drule_tac x="pi\z" in spec) +apply(drule_tac x="pi\P" in spec) +apply(drule mp) +apply(simp add: fresh_right) +apply(rule_tac x="(z):P" in exI) +apply(simp) +apply(auto)[1] +apply(drule sym) +apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst]) +apply(perm_simp add: csubst_eqvt fresh_right) +apply(auto)[1] +apply(drule_tac x="pi\c" in spec) +apply(drule_tac x="pi\Q" in spec) +apply(drule mp) +apply(simp add: fresh_right) +apply(rule_tac x=":Q" in exI) +apply(simp) +apply(auto)[1] +apply(drule sym) +apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst]) +apply(perm_simp add: nsubst_eqvt) +done + +lemma IMPRIGHT_eqvt_coname: + fixes pi::"coname prm" + shows "(pi\(IMPRIGHT (A IMP B) X Y Z U)) = IMPRIGHT (A IMP B) (pi\X) (pi\Y) (pi\Z) (pi\U)" +apply(auto simp add: perm_set_def) +apply(rule_tac x="pi\xb" in exI) +apply(rule_tac x="pi\a" in exI) +apply(rule_tac x="pi\b" in exI) +apply(rule_tac x="pi\M" in exI) +apply(simp) +apply(rule conjI) +apply(drule_tac pi="pi" in fic.eqvt(2)) +apply(simp) +apply(rule conjI) +apply(auto)[1] +apply(rule_tac x="(xb):(M{a:=((rev pi)\z).((rev pi)\P)})" in exI) +apply(perm_simp add: csubst_eqvt) +apply(drule sym) +apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: fresh_right) +apply(auto)[1] +apply(rule_tac x=":(M{xb:=<((rev pi)\c)>.((rev pi)\Q)})" in exI) +apply(perm_simp add: nsubst_eqvt) +apply(drule sym) +apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: fresh_right) +apply(rule_tac x="(rev pi)\(:ImpR xa..M b)" in exI) +apply(perm_simp) +apply(rule_tac x="(rev pi)\xa" in exI) +apply(rule_tac x="(rev pi)\a" in exI) +apply(rule_tac x="(rev pi)\b" in exI) +apply(rule_tac x="(rev pi)\M" in exI) +apply(simp) +apply(drule_tac pi="rev pi" in fic.eqvt(2)) +apply(simp) +apply(rule conjI) +apply(auto)[1] +apply(drule_tac x="pi\z" in spec) +apply(drule_tac x="pi\P" in spec) +apply(simp add: fresh_right) +apply(drule mp) +apply(rule_tac x="(z):P" in exI) +apply(simp) +apply(auto)[1] +apply(drule sym) +apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst]) +apply(perm_simp add: csubst_eqvt fresh_right) +apply(auto)[1] +apply(drule_tac x="pi\c" in spec) +apply(drule_tac x="pi\Q" in spec) +apply(simp add: fresh_right) +apply(drule mp) +apply(rule_tac x=":Q" in exI) +apply(simp) +apply(auto)[1] +apply(drule sym) +apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst]) +apply(perm_simp add: nsubst_eqvt fresh_right) +done + +consts + IMPLEFT::"ty \ ctrm set \ ntrm set \ ntrm set" + +nominal_primrec + "IMPLEFT (B IMP C) X Y = + { (y):ImpL .M (x).N y | x a y M N. fin (ImpL .M (x).N y) y \ :M \ X \ (x):N \ Y }" +apply(rule TrueI)+ +done + +lemma IMPLEFT_eqvt_name: + fixes pi::"name prm" + shows "(pi\(IMPLEFT (A IMP B) X Y)) = IMPLEFT (A IMP B) (pi\X) (pi\Y)" +apply(auto simp add: perm_set_def) +apply(rule_tac x="pi\xb" in exI) +apply(rule_tac x="pi\a" in exI) +apply(rule_tac x="pi\y" in exI) +apply(rule_tac x="pi\M" in exI) +apply(rule_tac x="pi\N" in exI) +apply(simp) +apply(rule conjI) +apply(drule_tac pi="pi" in fin.eqvt(1)) +apply(simp) +apply(rule conjI) +apply(rule_tac x=":M" in exI) +apply(simp) +apply(rule_tac x="(xb):N" in exI) +apply(simp) +apply(rule_tac x="(rev pi)\((y):ImpL .M (xa).N y)" in exI) +apply(perm_simp) +apply(rule_tac x="(rev pi)\xa" in exI) +apply(rule_tac x="(rev pi)\a" in exI) +apply(rule_tac x="(rev pi)\y" in exI) +apply(rule_tac x="(rev pi)\M" in exI) +apply(rule_tac x="(rev pi)\N" in exI) +apply(simp) +apply(drule_tac pi="rev pi" in fin.eqvt(1)) +apply(simp) +apply(drule sym) +apply(drule sym) +apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst]) +apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst]) +apply(simp) +done + +lemma IMPLEFT_eqvt_coname: + fixes pi::"coname prm" + shows "(pi\(IMPLEFT (A IMP B) X Y)) = IMPLEFT (A IMP B) (pi\X) (pi\Y)" +apply(auto simp add: perm_set_def) +apply(rule_tac x="pi\xb" in exI) +apply(rule_tac x="pi\a" in exI) +apply(rule_tac x="pi\y" in exI) +apply(rule_tac x="pi\M" in exI) +apply(rule_tac x="pi\N" in exI) +apply(simp) +apply(rule conjI) +apply(drule_tac pi="pi" in fin.eqvt(2)) +apply(simp) +apply(rule conjI) +apply(rule_tac x=":M" in exI) +apply(simp) +apply(rule_tac x="(xb):N" in exI) +apply(simp) +apply(rule_tac x="(rev pi)\((y):ImpL .M (xa).N y)" in exI) +apply(perm_simp) +apply(rule_tac x="(rev pi)\xa" in exI) +apply(rule_tac x="(rev pi)\a" in exI) +apply(rule_tac x="(rev pi)\y" in exI) +apply(rule_tac x="(rev pi)\M" in exI) +apply(rule_tac x="(rev pi)\N" in exI) +apply(simp) +apply(drule_tac pi="rev pi" in fin.eqvt(2)) +apply(simp) +apply(drule sym) +apply(drule sym) +apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst]) +apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst]) +apply(simp) +done + +lemma sum_cases: + shows "(\y. x=Inl y) \ (\y. x=Inr y)" +apply(rule_tac s="x" in sumE) +apply(auto) +done + +function + NEGc::"ty \ ntrm set \ ctrm set" +and + NEGn::"ty \ ctrm set \ ntrm set" +where + "NEGc (PR A) X = AXIOMSc (PR A) \ BINDINGc (PR A) X" +| "NEGc (NOT C) X = AXIOMSc (NOT C) \ BINDINGc (NOT C) X + \ NOTRIGHT (NOT C) (lfp (NEGn C \ NEGc C))" +| "NEGc (C AND D) X = AXIOMSc (C AND D) \ BINDINGc (C AND D) X + \ ANDRIGHT (C AND D) (NEGc C (lfp (NEGn C \ NEGc C))) (NEGc D (lfp (NEGn D \ NEGc D)))" +| "NEGc (C OR D) X = AXIOMSc (C OR D) \ BINDINGc (C OR D) X + \ ORRIGHT1 (C OR D) (NEGc C (lfp (NEGn C \ NEGc C))) + \ ORRIGHT2 (C OR D) (NEGc D (lfp (NEGn D \ NEGc D)))" +| "NEGc (C IMP D) X = AXIOMSc (C IMP D) \ BINDINGc (C IMP D) X + \ IMPRIGHT (C IMP D) (lfp (NEGn C \ NEGc C)) (NEGc D (lfp (NEGn D \ NEGc D))) + (lfp (NEGn D \ NEGc D)) (NEGc C (lfp (NEGn C \ NEGc C)))" +| "NEGn (PR A) X = AXIOMSn (PR A) \ BINDINGn (PR A) X" +| "NEGn (NOT C) X = AXIOMSn (NOT C) \ BINDINGn (NOT C) X + \ NOTLEFT (NOT C) (NEGc C (lfp (NEGn C \ NEGc C)))" +| "NEGn (C AND D) X = AXIOMSn (C AND D) \ BINDINGn (C AND D) X + \ ANDLEFT1 (C AND D) (lfp (NEGn C \ NEGc C)) + \ ANDLEFT2 (C AND D) (lfp (NEGn D \ NEGc D))" +| "NEGn (C OR D) X = AXIOMSn (C OR D) \ BINDINGn (C OR D) X + \ ORLEFT (C OR D) (lfp (NEGn C \ NEGc C)) (lfp (NEGn D \ NEGc D))" +| "NEGn (C IMP D) X = AXIOMSn (C IMP D) \ BINDINGn (C IMP D) X + \ IMPLEFT (C IMP D) (NEGc C (lfp (NEGn C \ NEGc C))) (lfp (NEGn D \ NEGc D))" +using ty_cases sum_cases +apply(auto simp add: ty.inject) +apply(drule_tac x="x" in meta_spec) +apply(auto simp add: ty.inject) +apply(rotate_tac 10) +apply(drule_tac x="a" in meta_spec) +apply(auto simp add: ty.inject) +apply(blast) +apply(blast) +apply(blast) +apply(rotate_tac 10) +apply(drule_tac x="a" in meta_spec) +apply(auto simp add: ty.inject) +apply(blast) +apply(blast) +apply(blast) +done + +termination +apply(relation "measure (sum_case (size\fst) (size\fst))") +apply(simp_all) +done + +text {* Candidates *} + +lemma test1: + shows "x\(X\Y) = (x\X \ x\Y)" +by blast + +lemma test2: + shows "x\(X\Y) = (x\X \ x\Y)" +by blast + +lemma pt_Collect_eqvt: + fixes pi::"'x prm" + assumes pt: "pt TYPE('a) TYPE('x)" + and at: "at TYPE('x)" + shows "pi\{x::'a. P x} = {x. P ((rev pi)\x)}" +apply(auto simp add: perm_set_def pt_rev_pi[OF pt, OF at]) +apply(rule_tac x="(rev pi)\x" in exI) +apply(simp add: pt_pi_rev[OF pt, OF at]) +done + +lemma big_inter_eqvt: + fixes pi1::"name prm" + and X::"('a::pt_name) set set" + and pi2::"coname prm" + and Y::"('b::pt_coname) set set" + shows "(pi1\(\ X)) = \ (pi1\X)" + and "(pi2\(\ Y)) = \ (pi2\Y)" +apply(auto simp add: perm_set_def) +apply(rule_tac x="(rev pi1)\x" in exI) +apply(perm_simp) +apply(rule ballI) +apply(drule_tac x="pi1\xa" in spec) +apply(auto) +apply(drule_tac x="xa" in spec) +apply(auto)[1] +apply(rule_tac x="(rev pi1)\xb" in exI) +apply(perm_simp) +apply(simp add: pt_set_bij1[OF pt_name_inst, OF at_name_inst]) +apply(simp add: pt_set_bij[OF pt_name_inst, OF at_name_inst]) +apply(simp add: pt_set_bij1[OF pt_name_inst, OF at_name_inst]) +apply(rule_tac x="(rev pi2)\x" in exI) +apply(perm_simp) +apply(rule ballI) +apply(drule_tac x="pi2\xa" in spec) +apply(auto) +apply(drule_tac x="xa" in spec) +apply(auto)[1] +apply(rule_tac x="(rev pi2)\xb" in exI) +apply(perm_simp) +apply(simp add: pt_set_bij1[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: pt_set_bij[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: pt_set_bij1[OF pt_coname_inst, OF at_coname_inst]) +done + +lemma lfp_eqvt: + fixes pi1::"name prm" + and f::"'a set \ ('a::pt_name) set" + and pi2::"coname prm" + and g::"'b set \ ('b::pt_coname) set" + shows "pi1\(lfp f) = lfp (pi1\f)" + and "pi2\(lfp g) = lfp (pi2\g)" +apply(simp add: lfp_def) +apply(simp add: Inf_set_def) +apply(simp add: big_inter_eqvt) +apply(simp add: pt_Collect_eqvt[OF pt_name_inst, OF at_name_inst]) +apply(subgoal_tac "{u. (pi1\f) u \ u} = {u. ((rev pi1)\((pi1\f) u)) \ ((rev pi1)\u)}") +apply(perm_simp) +apply(simp add: pt_subseteq_eqvt[OF pt_name_inst,OF at_name_inst]) +apply(simp add: lfp_def) +apply(simp add: Inf_set_def) +apply(simp add: big_inter_eqvt) +apply(simp add: pt_Collect_eqvt[OF pt_coname_inst, OF at_coname_inst]) +apply(subgoal_tac "{u. (pi2\g) u \ u} = {u. ((rev pi2)\((pi2\g) u)) \ ((rev pi2)\u)}") +apply(perm_simp) +apply(simp add: pt_subseteq_eqvt[OF pt_coname_inst,OF at_coname_inst]) +done + +abbreviation + CANDn::"ty \ ntrm set" ("\'(_')\" [60] 60) +where + "\(B)\ \ lfp (NEGn B \ NEGc B)" + +abbreviation + CANDc::"ty \ ctrm set" ("\<_>\" [60] 60) +where + "\\ \ NEGc B (\(B)\)" + +lemma NEGn_decreasing: + shows "X\Y \ NEGn B Y \ NEGn B X" +by (nominal_induct B rule: ty.induct) + (auto dest: BINDINGn_decreasing) + +lemma NEGc_decreasing: + shows "X\Y \ NEGc B Y \ NEGc B X" +by (nominal_induct B rule: ty.induct) + (auto dest: BINDINGc_decreasing) + +lemma mono_NEGn_NEGc: + shows "mono (NEGn B \ NEGc B)" + and "mono (NEGc B \ NEGn B)" +proof - + have "\X Y. X\Y \ NEGn B (NEGc B X) \ NEGn B (NEGc B Y)" + proof (intro strip) + fix X::"ntrm set" and Y::"ntrm set" + assume "X\Y" + then have "NEGc B Y \ NEGc B X" by (simp add: NEGc_decreasing) + then show "NEGn B (NEGc B X) \ NEGn B (NEGc B Y)" by (simp add: NEGn_decreasing) + qed + then show "mono (NEGn B \ NEGc B)" by (simp add: mono_def) +next + have "\X Y. X\Y \ NEGc B (NEGn B X) \ NEGc B (NEGn B Y)" + proof (intro strip) + fix X::"ctrm set" and Y::"ctrm set" + assume "X\Y" + then have "NEGn B Y \ NEGn B X" by (simp add: NEGn_decreasing) + then show "NEGc B (NEGn B X) \ NEGc B (NEGn B Y)" by (simp add: NEGc_decreasing) + qed + then show "mono (NEGc B \ NEGn B)" by (simp add: mono_def) +qed + +lemma NEG_simp: + shows "\\ = NEGc B (\(B)\)" + and "\(B)\ = NEGn B (\\)" +proof - + show "\\ = NEGc B (\(B)\)" by simp +next + have "\(B)\ \ lfp (NEGn B \ NEGc B)" by simp + then have "\(B)\ = (NEGn B \ NEGc B) (\(B)\)" using mono_NEGn_NEGc def_lfp_unfold by blast + then show "\(B)\ = NEGn B (\\)" by simp +qed + +lemma NEG_elim: + shows "M \ \\ \ M \ NEGc B (\(B)\)" + and "N \ \(B)\ \ N \ NEGn B (\\)" +using NEG_simp by (blast)+ + +lemma NEG_intro: + shows "M \ NEGc B (\(B)\) \ M \ \\" + and "N \ NEGn B (\\) \ N \ \(B)\" +using NEG_simp by (blast)+ + +lemma NEGc_simps: + shows "NEGc (PR A) (\(PR A)\) = AXIOMSc (PR A) \ BINDINGc (PR A) (\(PR A)\)" + and "NEGc (NOT C) (\(NOT C)\) = AXIOMSc (NOT C) \ BINDINGc (NOT C) (\(NOT C)\) + \ (NOTRIGHT (NOT C) (\(C)\))" + and "NEGc (C AND D) (\(C AND D)\) = AXIOMSc (C AND D) \ BINDINGc (C AND D) (\(C AND D)\) + \ (ANDRIGHT (C AND D) (\\) (\\))" + and "NEGc (C OR D) (\(C OR D)\) = AXIOMSc (C OR D) \ BINDINGc (C OR D) (\(C OR D)\) + \ (ORRIGHT1 (C OR D) (\\)) \ (ORRIGHT2 (C OR D) (\\))" + and "NEGc (C IMP D) (\(C IMP D)\) = AXIOMSc (C IMP D) \ BINDINGc (C IMP D) (\(C IMP D)\) + \ (IMPRIGHT (C IMP D) (\(C)\) (\\) (\(D)\) (\\))" +by (simp_all only: NEGc.simps) + +lemma AXIOMS_in_CANDs: + shows "AXIOMSn B \ (\(B)\)" + and "AXIOMSc B \ (\\)" +proof - + have "AXIOMSn B \ NEGn B (\\)" + by (nominal_induct B rule: ty.induct) (auto) + then show "AXIOMSn B \ \(B)\" using NEG_simp by blast +next + have "AXIOMSc B \ NEGc B (\(B)\)" + by (nominal_induct B rule: ty.induct) (auto) + then show "AXIOMSc B \ \\" using NEG_simp by blast +qed + +lemma Ax_in_CANDs: + shows "(y):Ax x a \ \(B)\" + and ":Ax x a \ \\" +proof - + have "(y):Ax x a \ AXIOMSn B" by (auto simp add: AXIOMSn_def) + also have "AXIOMSn B \ \(B)\" by (rule AXIOMS_in_CANDs) + finally show "(y):Ax x a \ \(B)\" by simp +next + have ":Ax x a \ AXIOMSc B" by (auto simp add: AXIOMSc_def) + also have "AXIOMSc B \ \\" by (rule AXIOMS_in_CANDs) + finally show ":Ax x a \ \\" by simp +qed + +lemma AXIOMS_eqvt_aux_name: + fixes pi::"name prm" + shows "M \ AXIOMSn B \ (pi\M) \ AXIOMSn B" + and "N \ AXIOMSc B \ (pi\N) \ AXIOMSc B" +apply(auto simp add: AXIOMSn_def AXIOMSc_def) +apply(rule_tac x="pi\x" in exI) +apply(rule_tac x="pi\y" in exI) +apply(rule_tac x="pi\b" in exI) +apply(simp) +apply(rule_tac x="pi\a" in exI) +apply(rule_tac x="pi\y" in exI) +apply(rule_tac x="pi\b" in exI) +apply(simp) +done + +lemma AXIOMS_eqvt_aux_coname: + fixes pi::"coname prm" + shows "M \ AXIOMSn B \ (pi\M) \ AXIOMSn B" + and "N \ AXIOMSc B \ (pi\N) \ AXIOMSc B" +apply(auto simp add: AXIOMSn_def AXIOMSc_def) +apply(rule_tac x="pi\x" in exI) +apply(rule_tac x="pi\y" in exI) +apply(rule_tac x="pi\b" in exI) +apply(simp) +apply(rule_tac x="pi\a" in exI) +apply(rule_tac x="pi\y" in exI) +apply(rule_tac x="pi\b" in exI) +apply(simp) +done + +lemma AXIOMS_eqvt_name: + fixes pi::"name prm" + shows "(pi\AXIOMSn B) = AXIOMSn B" + and "(pi\AXIOMSc B) = AXIOMSc B" +apply(auto) +apply(simp add: pt_set_bij1a[OF pt_name_inst, OF at_name_inst]) +apply(drule_tac pi="pi" in AXIOMS_eqvt_aux_name(1)) +apply(perm_simp) +apply(drule_tac pi="rev pi" in AXIOMS_eqvt_aux_name(1)) +apply(simp add: pt_set_bij1[OF pt_name_inst, OF at_name_inst]) +apply(simp add: pt_set_bij1a[OF pt_name_inst, OF at_name_inst]) +apply(drule_tac pi="pi" in AXIOMS_eqvt_aux_name(2)) +apply(perm_simp) +apply(drule_tac pi="rev pi" in AXIOMS_eqvt_aux_name(2)) +apply(simp add: pt_set_bij1[OF pt_name_inst, OF at_name_inst]) +done + +lemma AXIOMS_eqvt_coname: + fixes pi::"coname prm" + shows "(pi\AXIOMSn B) = AXIOMSn B" + and "(pi\AXIOMSc B) = AXIOMSc B" +apply(auto) +apply(simp add: pt_set_bij1a[OF pt_coname_inst, OF at_coname_inst]) +apply(drule_tac pi="pi" in AXIOMS_eqvt_aux_coname(1)) +apply(perm_simp) +apply(drule_tac pi="rev pi" in AXIOMS_eqvt_aux_coname(1)) +apply(simp add: pt_set_bij1[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: pt_set_bij1a[OF pt_coname_inst, OF at_coname_inst]) +apply(drule_tac pi="pi" in AXIOMS_eqvt_aux_coname(2)) +apply(perm_simp) +apply(drule_tac pi="rev pi" in AXIOMS_eqvt_aux_coname(2)) +apply(simp add: pt_set_bij1[OF pt_coname_inst, OF at_coname_inst]) +done + +lemma BINDING_eqvt_name: + fixes pi::"name prm" + shows "(pi\(BINDINGn B X)) = BINDINGn B (pi\X)" + and "(pi\(BINDINGc B Y)) = BINDINGc B (pi\Y)" +apply(auto simp add: BINDINGn_def BINDINGc_def perm_set_def) +apply(rule_tac x="pi\xb" in exI) +apply(rule_tac x="pi\M" in exI) +apply(simp) +apply(auto)[1] +apply(drule_tac x="(rev pi)\a" in spec) +apply(drule_tac x="(rev pi)\P" in spec) +apply(drule mp) +apply(drule sym) +apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst]) +apply(simp) +apply(drule_tac ?pi1.0="pi" in SNa_eqvt(1)) +apply(perm_simp add: nsubst_eqvt) +apply(rule_tac x="(rev pi\xa):(rev pi\M)" in exI) +apply(perm_simp) +apply(rule_tac x="rev pi\xa" in exI) +apply(rule_tac x="rev pi\M" in exI) +apply(simp) +apply(auto)[1] +apply(drule_tac x="pi\a" in spec) +apply(drule_tac x="pi\P" in spec) +apply(drule mp) +apply(force) +apply(drule_tac ?pi1.0="rev pi" in SNa_eqvt(1)) +apply(perm_simp add: nsubst_eqvt) +apply(rule_tac x="pi\a" in exI) +apply(rule_tac x="pi\M" in exI) +apply(simp) +apply(auto)[1] +apply(drule_tac x="(rev pi)\x" in spec) +apply(drule_tac x="(rev pi)\P" in spec) +apply(drule mp) +apply(drule sym) +apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst]) +apply(simp) +apply(drule_tac ?pi1.0="pi" in SNa_eqvt(1)) +apply(perm_simp add: csubst_eqvt) +apply(rule_tac x="<(rev pi\a)>:(rev pi\M)" in exI) +apply(perm_simp) +apply(rule_tac x="rev pi\a" in exI) +apply(rule_tac x="rev pi\M" in exI) +apply(simp) +apply(auto)[1] +apply(drule_tac x="pi\x" in spec) +apply(drule_tac x="pi\P" in spec) +apply(drule mp) +apply(force) +apply(drule_tac ?pi1.0="rev pi" in SNa_eqvt(1)) +apply(perm_simp add: csubst_eqvt) +done + +lemma BINDING_eqvt_coname: + fixes pi::"coname prm" + shows "(pi\(BINDINGn B X)) = BINDINGn B (pi\X)" + and "(pi\(BINDINGc B Y)) = BINDINGc B (pi\Y)" +apply(auto simp add: BINDINGn_def BINDINGc_def perm_set_def) +apply(rule_tac x="pi\xb" in exI) +apply(rule_tac x="pi\M" in exI) +apply(simp) +apply(auto)[1] +apply(drule_tac x="(rev pi)\a" in spec) +apply(drule_tac x="(rev pi)\P" in spec) +apply(drule mp) +apply(drule sym) +apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst]) +apply(simp) +apply(drule_tac ?pi2.0="pi" in SNa_eqvt(2)) +apply(perm_simp add: nsubst_eqvt) +apply(rule_tac x="(rev pi\xa):(rev pi\M)" in exI) +apply(perm_simp) +apply(rule_tac x="rev pi\xa" in exI) +apply(rule_tac x="rev pi\M" in exI) +apply(simp) +apply(auto)[1] +apply(drule_tac x="pi\a" in spec) +apply(drule_tac x="pi\P" in spec) +apply(drule mp) +apply(force) +apply(drule_tac ?pi2.0="rev pi" in SNa_eqvt(2)) +apply(perm_simp add: nsubst_eqvt) +apply(rule_tac x="pi\a" in exI) +apply(rule_tac x="pi\M" in exI) +apply(simp) +apply(auto)[1] +apply(drule_tac x="(rev pi)\x" in spec) +apply(drule_tac x="(rev pi)\P" in spec) +apply(drule mp) +apply(drule sym) +apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst]) +apply(simp) +apply(drule_tac ?pi2.0="pi" in SNa_eqvt(2)) +apply(perm_simp add: csubst_eqvt) +apply(rule_tac x="<(rev pi\a)>:(rev pi\M)" in exI) +apply(perm_simp) +apply(rule_tac x="rev pi\a" in exI) +apply(rule_tac x="rev pi\M" in exI) +apply(simp) +apply(auto)[1] +apply(drule_tac x="pi\x" in spec) +apply(drule_tac x="pi\P" in spec) +apply(drule mp) +apply(force) +apply(drule_tac ?pi2.0="rev pi" in SNa_eqvt(2)) +apply(perm_simp add: csubst_eqvt) +done + +lemma CAND_eqvt_name: + fixes pi::"name prm" + shows "(pi\(\(B)\)) = (\(B)\)" + and "(pi\(\\)) = (\\)" +proof (nominal_induct B rule: ty.induct) + case (PR X) + { case 1 show ?case + apply - + apply(simp add: lfp_eqvt) + apply(simp add: perm_fun_def) + apply(simp add: union_eqvt AXIOMS_eqvt_name BINDING_eqvt_name) + apply(perm_simp) + done + next + case 2 show ?case + apply - + apply(simp only: NEGc_simps) + apply(simp add: union_eqvt AXIOMS_eqvt_name BINDING_eqvt_name) + apply(simp add: lfp_eqvt) + apply(simp add: comp_def) + apply(simp add: perm_fun_def) + apply(simp add: union_eqvt AXIOMS_eqvt_name BINDING_eqvt_name) + apply(perm_simp) + done + } +next + case (NOT B) + have ih1: "pi\(\(B)\) = (\(B)\)" by fact + have ih2: "pi\(\\) = (\\)" by fact + have g: "pi\(\(NOT B)\) = (\(NOT B)\)" + apply - + apply(simp only: lfp_eqvt) + apply(simp only: comp_def) + apply(simp only: perm_fun_def) + apply(simp only: NEGc.simps NEGn.simps) + apply(simp only: union_eqvt AXIOMS_eqvt_name BINDING_eqvt_name NOTRIGHT_eqvt_name NOTLEFT_eqvt_name) + apply(perm_simp add: ih1 ih2) + done + { case 1 show ?case by (rule g) + next + case 2 show ?case + by (simp only: NEGc_simps union_eqvt AXIOMS_eqvt_name BINDING_eqvt_name NOTRIGHT_eqvt_name ih1 ih2 g) + } +next + case (AND A B) + have ih1: "pi\(\(A)\) = (\(A)\)" by fact + have ih2: "pi\(\\) = (\\)" by fact + have ih3: "pi\(\(B)\) = (\(B)\)" by fact + have ih4: "pi\(\\) = (\\)" by fact + have g: "pi\(\(A AND B)\) = (\(A AND B)\)" + apply - + apply(simp only: lfp_eqvt) + apply(simp only: comp_def) + apply(simp only: perm_fun_def) + apply(simp only: NEGc.simps NEGn.simps) + apply(simp only: union_eqvt AXIOMS_eqvt_name BINDING_eqvt_name ANDRIGHT_eqvt_name + ANDLEFT2_eqvt_name ANDLEFT1_eqvt_name) + apply(perm_simp add: ih1 ih2 ih3 ih4) + done + { case 1 show ?case by (rule g) + next + case 2 show ?case + by (simp only: NEGc_simps union_eqvt AXIOMS_eqvt_name BINDING_eqvt_name + ANDRIGHT_eqvt_name ANDLEFT1_eqvt_name ANDLEFT2_eqvt_name ih1 ih2 ih3 ih4 g) + } +next + case (OR A B) + have ih1: "pi\(\(A)\) = (\(A)\)" by fact + have ih2: "pi\(\\) = (\\)" by fact + have ih3: "pi\(\(B)\) = (\(B)\)" by fact + have ih4: "pi\(\\) = (\\)" by fact + have g: "pi\(\(A OR B)\) = (\(A OR B)\)" + apply - + apply(simp only: lfp_eqvt) + apply(simp only: comp_def) + apply(simp only: perm_fun_def) + apply(simp only: NEGc.simps NEGn.simps) + apply(simp only: union_eqvt AXIOMS_eqvt_name BINDING_eqvt_name ORRIGHT1_eqvt_name + ORRIGHT2_eqvt_name ORLEFT_eqvt_name) + apply(perm_simp add: ih1 ih2 ih3 ih4) + done + { case 1 show ?case by (rule g) + next + case 2 show ?case + by (simp only: NEGc_simps union_eqvt AXIOMS_eqvt_name BINDING_eqvt_name + ORRIGHT1_eqvt_name ORRIGHT2_eqvt_name ORLEFT_eqvt_name ih1 ih2 ih3 ih4 g) + } +next + case (IMP A B) + have ih1: "pi\(\(A)\) = (\(A)\)" by fact + have ih2: "pi\(\\) = (\\)" by fact + have ih3: "pi\(\(B)\) = (\(B)\)" by fact + have ih4: "pi\(\\) = (\\)" by fact + have g: "pi\(\(A IMP B)\) = (\(A IMP B)\)" + apply - + apply(simp only: lfp_eqvt) + apply(simp only: comp_def) + apply(simp only: perm_fun_def) + apply(simp only: NEGc.simps NEGn.simps) + apply(simp only: union_eqvt AXIOMS_eqvt_name BINDING_eqvt_name IMPRIGHT_eqvt_name IMPLEFT_eqvt_name) + apply(perm_simp add: ih1 ih2 ih3 ih4) + done + { case 1 show ?case by (rule g) + next + case 2 show ?case + by (simp only: NEGc_simps union_eqvt AXIOMS_eqvt_name BINDING_eqvt_name + IMPRIGHT_eqvt_name IMPLEFT_eqvt_name ih1 ih2 ih3 ih4 g) + } +qed + +lemma CAND_eqvt_coname: + fixes pi::"coname prm" + shows "(pi\(\(B)\)) = (\(B)\)" + and "(pi\(\\)) = (\\)" +proof (nominal_induct B rule: ty.induct) + case (PR X) + { case 1 show ?case + apply - + apply(simp add: lfp_eqvt) + apply(simp add: perm_fun_def) + apply(simp add: union_eqvt AXIOMS_eqvt_coname BINDING_eqvt_coname) + apply(perm_simp) + done + next + case 2 show ?case + apply - + apply(simp only: NEGc_simps) + apply(simp add: union_eqvt AXIOMS_eqvt_coname BINDING_eqvt_coname) + apply(simp add: lfp_eqvt) + apply(simp add: comp_def) + apply(simp add: perm_fun_def) + apply(simp add: union_eqvt AXIOMS_eqvt_coname BINDING_eqvt_coname) + apply(perm_simp) + done + } +next + case (NOT B) + have ih1: "pi\(\(B)\) = (\(B)\)" by fact + have ih2: "pi\(\\) = (\\)" by fact + have g: "pi\(\(NOT B)\) = (\(NOT B)\)" + apply - + apply(simp only: lfp_eqvt) + apply(simp only: comp_def) + apply(simp only: perm_fun_def) + apply(simp only: NEGc.simps NEGn.simps) + apply(simp only: union_eqvt AXIOMS_eqvt_coname BINDING_eqvt_coname + NOTRIGHT_eqvt_coname NOTLEFT_eqvt_coname) + apply(perm_simp add: ih1 ih2) + done + { case 1 show ?case by (rule g) + next + case 2 show ?case + by (simp only: NEGc_simps union_eqvt AXIOMS_eqvt_coname BINDING_eqvt_coname + NOTRIGHT_eqvt_coname ih1 ih2 g) + } +next + case (AND A B) + have ih1: "pi\(\(A)\) = (\(A)\)" by fact + have ih2: "pi\(\\) = (\\)" by fact + have ih3: "pi\(\(B)\) = (\(B)\)" by fact + have ih4: "pi\(\\) = (\\)" by fact + have g: "pi\(\(A AND B)\) = (\(A AND B)\)" + apply - + apply(simp only: lfp_eqvt) + apply(simp only: comp_def) + apply(simp only: perm_fun_def) + apply(simp only: NEGc.simps NEGn.simps) + apply(simp only: union_eqvt AXIOMS_eqvt_coname BINDING_eqvt_coname ANDRIGHT_eqvt_coname + ANDLEFT2_eqvt_coname ANDLEFT1_eqvt_coname) + apply(perm_simp add: ih1 ih2 ih3 ih4) + done + { case 1 show ?case by (rule g) + next + case 2 show ?case + by (simp only: NEGc_simps union_eqvt AXIOMS_eqvt_coname BINDING_eqvt_coname + ANDRIGHT_eqvt_coname ANDLEFT1_eqvt_coname ANDLEFT2_eqvt_coname ih1 ih2 ih3 ih4 g) + } +next + case (OR A B) + have ih1: "pi\(\(A)\) = (\(A)\)" by fact + have ih2: "pi\(\\) = (\\)" by fact + have ih3: "pi\(\(B)\) = (\(B)\)" by fact + have ih4: "pi\(\\) = (\\)" by fact + have g: "pi\(\(A OR B)\) = (\(A OR B)\)" + apply - + apply(simp only: lfp_eqvt) + apply(simp only: comp_def) + apply(simp only: perm_fun_def) + apply(simp only: NEGc.simps NEGn.simps) + apply(simp only: union_eqvt AXIOMS_eqvt_coname BINDING_eqvt_coname ORRIGHT1_eqvt_coname + ORRIGHT2_eqvt_coname ORLEFT_eqvt_coname) + apply(perm_simp add: ih1 ih2 ih3 ih4) + done + { case 1 show ?case by (rule g) + next + case 2 show ?case + by (simp only: NEGc_simps union_eqvt AXIOMS_eqvt_coname BINDING_eqvt_coname + ORRIGHT1_eqvt_coname ORRIGHT2_eqvt_coname ORLEFT_eqvt_coname ih1 ih2 ih3 ih4 g) + } +next + case (IMP A B) + have ih1: "pi\(\(A)\) = (\(A)\)" by fact + have ih2: "pi\(\\) = (\\)" by fact + have ih3: "pi\(\(B)\) = (\(B)\)" by fact + have ih4: "pi\(\\) = (\\)" by fact + have g: "pi\(\(A IMP B)\) = (\(A IMP B)\)" + apply - + apply(simp only: lfp_eqvt) + apply(simp only: comp_def) + apply(simp only: perm_fun_def) + apply(simp only: NEGc.simps NEGn.simps) + apply(simp only: union_eqvt AXIOMS_eqvt_coname BINDING_eqvt_coname IMPRIGHT_eqvt_coname + IMPLEFT_eqvt_coname) + apply(perm_simp add: ih1 ih2 ih3 ih4) + done + { case 1 show ?case by (rule g) + next + case 2 show ?case + by (simp only: NEGc_simps union_eqvt AXIOMS_eqvt_coname BINDING_eqvt_coname + IMPRIGHT_eqvt_coname IMPLEFT_eqvt_coname ih1 ih2 ih3 ih4 g) + } +qed + +text {* Elimination rules for the set-operators *} + +lemma BINDINGc_elim: + assumes a: ":M \ BINDINGc B (\(B)\)" + shows "\x P. ((x):P)\(\(B)\) \ SNa (M{a:=(x).P})" +using a +apply(auto simp add: BINDINGc_def) +apply(auto simp add: ctrm.inject alpha) +apply(drule_tac x="[(a,aa)]\x" in spec) +apply(drule_tac x="[(a,aa)]\P" in spec) +apply(drule mp) +apply(drule_tac pi="[(a,aa)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_coname) +apply(drule_tac ?pi2.0="[(a,aa)]" in SNa_eqvt(2)) +apply(perm_simp add: csubst_eqvt) +done + +lemma BINDINGn_elim: + assumes a: "(x):M \ BINDINGn B (\\)" + shows "\c P. (:P)\(\\) \ SNa (M{x:=.P})" +using a +apply(auto simp add: BINDINGn_def) +apply(auto simp add: ntrm.inject alpha) +apply(drule_tac x="[(x,xa)]\c" in spec) +apply(drule_tac x="[(x,xa)]\P" in spec) +apply(drule mp) +apply(drule_tac pi="[(x,xa)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name) +apply(drule_tac ?pi1.0="[(x,xa)]" in SNa_eqvt(1)) +apply(perm_simp add: nsubst_eqvt) +done + +lemma NOTRIGHT_elim: + assumes a: ":M \ NOTRIGHT (NOT B) (\(B)\)" + obtains x' M' where "M = NotR (x').M' a" and "fic (NotR (x').M' a) a" and "(x'):M' \ (\(B)\)" +using a +apply(auto simp add: ctrm.inject alpha abs_fresh calc_atm) +apply(drule_tac x="x" in meta_spec) +apply(drule_tac x="[(a,aa)]\M" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(a,aa)]" in fic.eqvt(2)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(a,aa)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(simp) +done + +lemma NOTLEFT_elim: + assumes a: "(x):M \ NOTLEFT (NOT B) (\\)" + obtains a' M' where "M = NotL .M' x" and "fin (NotL .M' x) x" and ":M' \ (\\)" +using a +apply(auto simp add: ntrm.inject alpha abs_fresh calc_atm) +apply(drule_tac x="a" in meta_spec) +apply(drule_tac x="[(x,xa)]\M" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(x,xa)]" in fin.eqvt(1)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(x,xa)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(simp) +done + +lemma ANDRIGHT_elim: + assumes a: ":M \ ANDRIGHT (B AND C) (\\) (\\)" + obtains d' M' e' N' where "M = AndR .M' .N' a" and "fic (AndR .M' .N' a) a" + and ":M' \ (\\)" and ":N' \ (\\)" +using a +apply(auto simp add: ctrm.inject alpha abs_fresh calc_atm fresh_atm) +apply(drule_tac x="c" in meta_spec) +apply(drule_tac x="[(a,c)]\M" in meta_spec) +apply(drule_tac x="c" in meta_spec) +apply(drule_tac x="[(a,c)]\N" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(a,c)]" in fic.eqvt(2)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(a,c)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(drule meta_mp) +apply(drule_tac pi="[(a,c)]" and x=":N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(simp) +apply(case_tac "a=b") +apply(simp) +apply(drule_tac x="c" in meta_spec) +apply(drule_tac x="[(b,c)]\M" in meta_spec) +apply(drule_tac x="c" in meta_spec) +apply(drule_tac x="[(b,c)]\N" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(b,c)]" in fic.eqvt(2)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(b,c)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(drule meta_mp) +apply(drule_tac pi="[(b,c)]" and x=":N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(simp) +apply(simp) +apply(case_tac "c=b") +apply(simp) +apply(drule_tac x="b" in meta_spec) +apply(drule_tac x="[(a,b)]\M" in meta_spec) +apply(drule_tac x="a" in meta_spec) +apply(drule_tac x="[(a,b)]\N" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(a,b)]" in fic.eqvt(2)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(a,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(drule meta_mp) +apply(drule_tac pi="[(a,b)]" and x=":N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(simp) +apply(simp) +apply(drule_tac x="c" in meta_spec) +apply(drule_tac x="[(a,c)]\M" in meta_spec) +apply(drule_tac x="b" in meta_spec) +apply(drule_tac x="[(a,c)]\N" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(a,c)]" in fic.eqvt(2)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(a,c)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(drule meta_mp) +apply(drule_tac pi="[(a,c)]" and x=":N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(simp) +apply(case_tac "a=aa") +apply(simp) +apply(drule_tac x="c" in meta_spec) +apply(drule_tac x="[(aa,c)]\M" in meta_spec) +apply(drule_tac x="c" in meta_spec) +apply(drule_tac x="[(aa,c)]\N" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(aa,c)]" in fic.eqvt(2)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(aa,c)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(drule meta_mp) +apply(drule_tac pi="[(aa,c)]" and x=":N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(simp) +apply(simp) +apply(case_tac "c=aa") +apply(simp) +apply(drule_tac x="a" in meta_spec) +apply(drule_tac x="[(a,aa)]\M" in meta_spec) +apply(drule_tac x="aa" in meta_spec) +apply(drule_tac x="[(a,aa)]\N" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(a,aa)]" in fic.eqvt(2)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(a,aa)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(drule meta_mp) +apply(drule_tac pi="[(a,aa)]" and x=":N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(simp) +apply(simp) +apply(drule_tac x="aa" in meta_spec) +apply(drule_tac x="[(a,c)]\M" in meta_spec) +apply(drule_tac x="c" in meta_spec) +apply(drule_tac x="[(a,c)]\N" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(a,c)]" in fic.eqvt(2)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(a,c)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(drule meta_mp) +apply(drule_tac pi="[(a,c)]" and x=":N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(simp) +apply(case_tac "a=aa") +apply(simp) +apply(case_tac "aa=b") +apply(simp) +apply(drule_tac x="c" in meta_spec) +apply(drule_tac x="[(b,c)]\M" in meta_spec) +apply(drule_tac x="c" in meta_spec) +apply(drule_tac x="[(b,c)]\N" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(b,c)]" in fic.eqvt(2)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(b,c)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(drule meta_mp) +apply(drule_tac pi="[(b,c)]" and x=":N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(simp) +apply(simp) +apply(case_tac "c=b") +apply(simp) +apply(drule_tac x="b" in meta_spec) +apply(drule_tac x="[(aa,b)]\M" in meta_spec) +apply(drule_tac x="aa" in meta_spec) +apply(drule_tac x="[(aa,b)]\N" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(aa,b)]" in fic.eqvt(2)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(aa,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(drule meta_mp) +apply(drule_tac pi="[(aa,b)]" and x=":N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(simp) +apply(simp) +apply(drule_tac x="c" in meta_spec) +apply(drule_tac x="[(aa,c)]\M" in meta_spec) +apply(drule_tac x="b" in meta_spec) +apply(drule_tac x="[(aa,c)]\N" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(aa,c)]" in fic.eqvt(2)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(aa,c)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(drule meta_mp) +apply(drule_tac pi="[(aa,c)]" and x=":N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(simp) +apply(simp) +apply(case_tac "c=aa") +apply(simp) +apply(case_tac "a=b") +apply(simp) +apply(drule_tac x="b" in meta_spec) +apply(drule_tac x="[(b,aa)]\M" in meta_spec) +apply(drule_tac x="aa" in meta_spec) +apply(drule_tac x="[(b,aa)]\N" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(b,aa)]" in fic.eqvt(2)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(b,aa)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(drule meta_mp) +apply(drule_tac pi="[(b,aa)]" and x=":N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(simp) +apply(simp) +apply(case_tac "aa=b") +apply(simp) +apply(drule_tac x="a" in meta_spec) +apply(drule_tac x="[(a,b)]\M" in meta_spec) +apply(drule_tac x="a" in meta_spec) +apply(drule_tac x="[(a,b)]\N" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(a,b)]" in fic.eqvt(2)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(a,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(drule meta_mp) +apply(drule_tac pi="[(a,b)]" and x=":N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(simp) +apply(simp) +apply(drule_tac x="a" in meta_spec) +apply(drule_tac x="[(a,aa)]\M" in meta_spec) +apply(drule_tac x="b" in meta_spec) +apply(drule_tac x="[(a,aa)]\N" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(a,aa)]" in fic.eqvt(2)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(a,aa)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(drule meta_mp) +apply(drule_tac pi="[(a,aa)]" and x=":N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(simp) +apply(simp) +apply(case_tac "a=b") +apply(simp) +apply(drule_tac x="aa" in meta_spec) +apply(drule_tac x="[(b,c)]\M" in meta_spec) +apply(drule_tac x="c" in meta_spec) +apply(drule_tac x="[(b,c)]\N" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(b,c)]" in fic.eqvt(2)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(b,c)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(drule meta_mp) +apply(drule_tac pi="[(b,c)]" and x=":N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(simp) +apply(simp) +apply(case_tac "c=b") +apply(simp) +apply(drule_tac x="aa" in meta_spec) +apply(drule_tac x="[(a,b)]\M" in meta_spec) +apply(drule_tac x="a" in meta_spec) +apply(drule_tac x="[(a,b)]\N" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(a,b)]" in fic.eqvt(2)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(a,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(drule meta_mp) +apply(drule_tac pi="[(a,b)]" and x=":N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(simp) +apply(simp) +apply(drule_tac x="aa" in meta_spec) +apply(drule_tac x="[(a,c)]\M" in meta_spec) +apply(drule_tac x="b" in meta_spec) +apply(drule_tac x="[(a,c)]\N" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(a,c)]" in fic.eqvt(2)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(a,c)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(drule meta_mp) +apply(drule_tac pi="[(a,c)]" and x=":N" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(simp) +done + +lemma ANDLEFT1_elim: + assumes a: "(x):M \ ANDLEFT1 (B AND C) (\(B)\)" + obtains x' M' where "M = AndL1 (x').M' x" and "fin (AndL1 (x').M' x) x" and "(x'):M' \ (\(B)\)" +using a +apply(auto simp add: ntrm.inject alpha abs_fresh calc_atm) +apply(drule_tac x="y" in meta_spec) +apply(drule_tac x="[(x,y)]\M" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(x,y)]" in fin.eqvt(1)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(x,y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(simp) +apply(case_tac "x=xa") +apply(simp) +apply(drule_tac x="y" in meta_spec) +apply(drule_tac x="[(xa,y)]\M" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(xa,y)]" in fin.eqvt(1)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(xa,y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(simp) +apply(simp) +apply(case_tac "y=xa") +apply(simp) +apply(drule_tac x="x" in meta_spec) +apply(drule_tac x="[(x,xa)]\M" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(x,xa)]" in fin.eqvt(1)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(x,xa)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(simp) +apply(simp) +apply(drule_tac x="xa" in meta_spec) +apply(drule_tac x="[(x,y)]\M" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(x,y)]" in fin.eqvt(1)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(x,y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(simp) +done + +lemma ANDLEFT2_elim: + assumes a: "(x):M \ ANDLEFT2 (B AND C) (\(C)\)" + obtains x' M' where "M = AndL2 (x').M' x" and "fin (AndL2 (x').M' x) x" and "(x'):M' \ (\(C)\)" +using a +apply(auto simp add: ntrm.inject alpha abs_fresh calc_atm) +apply(drule_tac x="y" in meta_spec) +apply(drule_tac x="[(x,y)]\M" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(x,y)]" in fin.eqvt(1)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(x,y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(simp) +apply(case_tac "x=xa") +apply(simp) +apply(drule_tac x="y" in meta_spec) +apply(drule_tac x="[(xa,y)]\M" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(xa,y)]" in fin.eqvt(1)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(xa,y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(simp) +apply(simp) +apply(case_tac "y=xa") +apply(simp) +apply(drule_tac x="x" in meta_spec) +apply(drule_tac x="[(x,xa)]\M" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(x,xa)]" in fin.eqvt(1)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(x,xa)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(simp) +apply(simp) +apply(drule_tac x="xa" in meta_spec) +apply(drule_tac x="[(x,y)]\M" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(x,y)]" in fin.eqvt(1)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(x,y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(simp) +done + +lemma ORRIGHT1_elim: + assumes a: ":M \ ORRIGHT1 (B OR C) (\\)" + obtains a' M' where "M = OrR1 .M' a" and "fic (OrR1 .M' a) a" and ":M' \ (\\)" +using a +apply(auto simp add: ctrm.inject alpha abs_fresh calc_atm) +apply(drule_tac x="b" in meta_spec) +apply(drule_tac x="[(a,b)]\M" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(a,b)]" in fic.eqvt(2)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(a,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(simp) +apply(case_tac "a=aa") +apply(simp) +apply(drule_tac x="b" in meta_spec) +apply(drule_tac x="[(aa,b)]\M" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(aa,b)]" in fic.eqvt(2)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(aa,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(simp) +apply(simp) +apply(case_tac "b=aa") +apply(simp) +apply(drule_tac x="a" in meta_spec) +apply(drule_tac x="[(a,aa)]\M" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(a,aa)]" in fic.eqvt(2)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(a,aa)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(simp) +apply(simp) +apply(drule_tac x="aa" in meta_spec) +apply(drule_tac x="[(a,b)]\M" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(a,b)]" in fic.eqvt(2)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(a,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(simp) +done + +lemma ORRIGHT2_elim: + assumes a: ":M \ ORRIGHT2 (B OR C) (\\)" + obtains a' M' where "M = OrR2 .M' a" and "fic (OrR2 .M' a) a" and ":M' \ (\\)" +using a +apply(auto simp add: ctrm.inject alpha abs_fresh calc_atm) +apply(drule_tac x="b" in meta_spec) +apply(drule_tac x="[(a,b)]\M" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(a,b)]" in fic.eqvt(2)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(a,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(simp) +apply(case_tac "a=aa") +apply(simp) +apply(drule_tac x="b" in meta_spec) +apply(drule_tac x="[(aa,b)]\M" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(aa,b)]" in fic.eqvt(2)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(aa,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(simp) +apply(simp) +apply(case_tac "b=aa") +apply(simp) +apply(drule_tac x="a" in meta_spec) +apply(drule_tac x="[(a,aa)]\M" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(a,aa)]" in fic.eqvt(2)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(a,aa)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(simp) +apply(simp) +apply(drule_tac x="aa" in meta_spec) +apply(drule_tac x="[(a,b)]\M" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(a,b)]" in fic.eqvt(2)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(a,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(simp) +done + +lemma ORLEFT_elim: + assumes a: "(x):M \ ORLEFT (B OR C) (\(B)\) (\(C)\)" + obtains y' M' z' N' where "M = OrL (y').M' (z').N' x" and "fin (OrL (y').M' (z').N' x) x" + and "(y'):M' \ (\(B)\)" and "(z'):N' \ (\(C)\)" +using a +apply(auto simp add: ntrm.inject alpha abs_fresh calc_atm fresh_atm) +apply(drule_tac x="z" in meta_spec) +apply(drule_tac x="[(x,z)]\M" in meta_spec) +apply(drule_tac x="z" in meta_spec) +apply(drule_tac x="[(x,z)]\N" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(x,z)]" in fin.eqvt(1)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(x,z)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(drule meta_mp) +apply(drule_tac pi="[(x,z)]" and x="(x):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(simp) +apply(case_tac "x=y") +apply(simp) +apply(drule_tac x="z" in meta_spec) +apply(drule_tac x="[(y,z)]\M" in meta_spec) +apply(drule_tac x="z" in meta_spec) +apply(drule_tac x="[(y,z)]\N" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(y,z)]" in fin.eqvt(1)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(y,z)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(drule meta_mp) +apply(drule_tac pi="[(y,z)]" and x="(y):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(simp) +apply(simp) +apply(case_tac "z=y") +apply(simp) +apply(drule_tac x="y" in meta_spec) +apply(drule_tac x="[(x,y)]\M" in meta_spec) +apply(drule_tac x="x" in meta_spec) +apply(drule_tac x="[(x,y)]\N" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(x,y)]" in fin.eqvt(1)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(x,y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(drule meta_mp) +apply(drule_tac pi="[(x,y)]" and x="(y):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(simp) +apply(simp) +apply(drule_tac x="z" in meta_spec) +apply(drule_tac x="[(x,z)]\M" in meta_spec) +apply(drule_tac x="y" in meta_spec) +apply(drule_tac x="[(x,z)]\N" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(x,z)]" in fin.eqvt(1)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(x,z)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(drule meta_mp) +apply(drule_tac pi="[(x,z)]" and x="(y):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(simp) +apply(case_tac "x=xa") +apply(simp) +apply(drule_tac x="z" in meta_spec) +apply(drule_tac x="[(xa,z)]\M" in meta_spec) +apply(drule_tac x="z" in meta_spec) +apply(drule_tac x="[(xa,z)]\N" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(xa,z)]" in fin.eqvt(1)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(xa,z)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(drule meta_mp) +apply(drule_tac pi="[(xa,z)]" and x="(xa):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(simp) +apply(simp) +apply(case_tac "z=xa") +apply(simp) +apply(drule_tac x="x" in meta_spec) +apply(drule_tac x="[(x,xa)]\M" in meta_spec) +apply(drule_tac x="xa" in meta_spec) +apply(drule_tac x="[(x,xa)]\N" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(x,xa)]" in fin.eqvt(1)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(x,xa)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(drule meta_mp) +apply(drule_tac pi="[(x,xa)]" and x="(x):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(simp) +apply(simp) +apply(drule_tac x="xa" in meta_spec) +apply(drule_tac x="[(x,z)]\M" in meta_spec) +apply(drule_tac x="z" in meta_spec) +apply(drule_tac x="[(x,z)]\N" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(x,z)]" in fin.eqvt(1)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(x,z)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(drule meta_mp) +apply(drule_tac pi="[(x,z)]" and x="(x):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(simp) +apply(case_tac "x=xa") +apply(simp) +apply(case_tac "xa=y") +apply(simp) +apply(drule_tac x="z" in meta_spec) +apply(drule_tac x="[(y,z)]\M" in meta_spec) +apply(drule_tac x="z" in meta_spec) +apply(drule_tac x="[(y,z)]\N" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(y,z)]" in fin.eqvt(1)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(y,z)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(drule meta_mp) +apply(drule_tac pi="[(y,z)]" and x="(y):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(simp) +apply(simp) +apply(case_tac "z=y") +apply(simp) +apply(drule_tac x="y" in meta_spec) +apply(drule_tac x="[(xa,y)]\M" in meta_spec) +apply(drule_tac x="xa" in meta_spec) +apply(drule_tac x="[(xa,y)]\N" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(xa,y)]" in fin.eqvt(1)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(xa,y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(drule meta_mp) +apply(drule_tac pi="[(xa,y)]" and x="(y):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(simp) +apply(simp) +apply(drule_tac x="z" in meta_spec) +apply(drule_tac x="[(xa,z)]\M" in meta_spec) +apply(drule_tac x="y" in meta_spec) +apply(drule_tac x="[(xa,z)]\N" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(xa,z)]" in fin.eqvt(1)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(xa,z)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(drule meta_mp) +apply(drule_tac pi="[(xa,z)]" and x="(y):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(simp) +apply(simp) +apply(case_tac "z=xa") +apply(simp) +apply(case_tac "x=y") +apply(simp) +apply(drule_tac x="y" in meta_spec) +apply(drule_tac x="[(y,xa)]\M" in meta_spec) +apply(drule_tac x="xa" in meta_spec) +apply(drule_tac x="[(y,xa)]\N" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(y,xa)]" in fin.eqvt(1)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(y,xa)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(drule meta_mp) +apply(drule_tac pi="[(y,xa)]" and x="(y):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(simp) +apply(simp) +apply(case_tac "xa=y") +apply(simp) +apply(drule_tac x="x" in meta_spec) +apply(drule_tac x="[(x,y)]\M" in meta_spec) +apply(drule_tac x="x" in meta_spec) +apply(drule_tac x="[(x,y)]\N" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(x,y)]" in fin.eqvt(1)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(x,y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(drule meta_mp) +apply(drule_tac pi="[(x,y)]" and x="(y):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(simp) +apply(simp) +apply(drule_tac x="x" in meta_spec) +apply(drule_tac x="[(x,xa)]\M" in meta_spec) +apply(drule_tac x="y" in meta_spec) +apply(drule_tac x="[(x,xa)]\N" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(x,xa)]" in fin.eqvt(1)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(x,xa)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(drule meta_mp) +apply(drule_tac pi="[(x,xa)]" and x="(y):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(simp) +apply(simp) +apply(case_tac "x=y") +apply(simp) +apply(drule_tac x="xa" in meta_spec) +apply(drule_tac x="[(y,z)]\M" in meta_spec) +apply(drule_tac x="z" in meta_spec) +apply(drule_tac x="[(y,z)]\N" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(y,z)]" in fin.eqvt(1)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(y,z)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(drule meta_mp) +apply(drule_tac pi="[(y,z)]" and x="(y):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(simp) +apply(simp) +apply(case_tac "z=y") +apply(simp) +apply(drule_tac x="xa" in meta_spec) +apply(drule_tac x="[(x,y)]\M" in meta_spec) +apply(drule_tac x="x" in meta_spec) +apply(drule_tac x="[(x,y)]\N" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(x,y)]" in fin.eqvt(1)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(x,y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(drule meta_mp) +apply(drule_tac pi="[(x,y)]" and x="(y):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(simp) +apply(simp) +apply(drule_tac x="xa" in meta_spec) +apply(drule_tac x="[(x,z)]\M" in meta_spec) +apply(drule_tac x="y" in meta_spec) +apply(drule_tac x="[(x,z)]\N" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(x,z)]" in fin.eqvt(1)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(x,z)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(drule meta_mp) +apply(drule_tac pi="[(x,z)]" and x="(y):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(simp) +done + +lemma IMPRIGHT_elim: + assumes a: ":M \ IMPRIGHT (B IMP C) (\(B)\) (\\) (\(C)\) (\\)" + obtains x' a' M' where "M = ImpR (x')..M' a" and "fic (ImpR (x')..M' a) a" + and "\z P. x'\(z,P) \ (z):P \ \(C)\ \ (x'):(M'{a':=(z).P}) \ \(B)\" + and "\c Q. a'\(c,Q) \ :Q \ \\ \ :(M'{x':=.Q}) \ \\" +using a +apply(auto simp add: ctrm.inject alpha abs_fresh calc_atm) +apply(drule_tac x="x" in meta_spec) +apply(drule_tac x="b" in meta_spec) +apply(drule_tac x="[(a,b)]\M" in meta_spec) +apply(simp) +apply(drule_tac pi="[(a,b)]" in fic.eqvt(2)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(auto)[1] +apply(drule_tac pi="[(a,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(drule_tac x="z" in spec) +apply(drule_tac x="[(a,b)]\P" in spec) +apply(simp add: fresh_prod fresh_left calc_atm) +apply(drule_tac pi="[(a,b)]" and x="(x):M{a:=(z).([(a,b)]\P)}" + in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(perm_simp add: calc_atm csubst_eqvt CAND_eqvt_coname) +apply(drule meta_mp) +apply(auto)[1] +apply(drule_tac pi="[(a,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_coname) +apply(rotate_tac 2) +apply(drule_tac x="[(a,b)]\c" in spec) +apply(drule_tac x="[(a,b)]\Q" in spec) +apply(simp add: fresh_prod fresh_left) +apply(drule mp) +apply(simp add: calc_atm) +apply(drule_tac pi="[(a,b)]" and x=":M{x:=<([(a,b)]\c)>.([(a,b)]\Q)}" + in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(perm_simp add: nsubst_eqvt CAND_eqvt_coname) +apply(simp add: calc_atm) +apply(simp) +apply(case_tac "a=aa") +apply(simp) +apply(drule_tac x="x" in meta_spec) +apply(drule_tac x="b" in meta_spec) +apply(drule_tac x="[(aa,b)]\M" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(aa,b)]" in fic.eqvt(2)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(auto)[1] +apply(drule_tac pi="[(a,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(drule_tac x="z" in spec) +apply(drule_tac x="[(a,b)]\P" in spec) +apply(simp add: fresh_prod fresh_left calc_atm) +apply(drule_tac pi="[(a,b)]" and x="(x):M{a:=(z).([(a,b)]\P)}" + in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(perm_simp add: calc_atm csubst_eqvt CAND_eqvt_coname) +apply(drule meta_mp) +apply(auto)[1] +apply(drule_tac pi="[(a,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_coname) +apply(drule_tac x="[(a,b)]\c" in spec) +apply(drule_tac x="[(a,b)]\Q" in spec) +apply(simp) +apply(simp add: fresh_prod fresh_left) +apply(drule mp) +apply(simp add: calc_atm) +apply(drule_tac pi="[(a,b)]" and x=":M{x:=<([(a,b)]\c)>.([(a,b)]\Q)}" + in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(perm_simp add: nsubst_eqvt CAND_eqvt_coname) +apply(simp add: calc_atm) +apply(simp) +apply(simp) +apply(case_tac "b=aa") +apply(simp) +apply(drule_tac x="x" in meta_spec) +apply(drule_tac x="a" in meta_spec) +apply(drule_tac x="[(a,aa)]\M" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(a,aa)]" in fic.eqvt(2)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(auto)[1] +apply(drule_tac pi="[(a,aa)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(drule_tac x="z" in spec) +apply(drule_tac x="[(a,aa)]\P" in spec) +apply(simp add: fresh_prod fresh_left calc_atm) +apply(drule_tac pi="[(a,aa)]" and x="(x):M{aa:=(z).([(a,aa)]\P)}" + in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(perm_simp add: calc_atm csubst_eqvt CAND_eqvt_coname) +apply(drule meta_mp) +apply(auto)[1] +apply(drule_tac pi="[(a,aa)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_coname) +apply(drule_tac x="[(a,aa)]\c" in spec) +apply(drule_tac x="[(a,aa)]\Q" in spec) +apply(simp) +apply(simp add: fresh_prod fresh_left) +apply(drule mp) +apply(simp add: calc_atm) +apply(drule_tac pi="[(a,aa)]" and x=":M{x:=<([(a,aa)]\c)>.([(a,aa)]\Q)}" + in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(perm_simp add: nsubst_eqvt CAND_eqvt_coname) +apply(simp add: calc_atm) +apply(simp) +apply(simp) +apply(drule_tac x="x" in meta_spec) +apply(drule_tac x="aa" in meta_spec) +apply(drule_tac x="[(a,b)]\M" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(a,b)]" in fic.eqvt(2)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(auto)[1] +apply(drule_tac pi="[(a,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm CAND_eqvt_coname) +apply(drule_tac x="z" in spec) +apply(drule_tac x="[(a,b)]\P" in spec) +apply(simp add: fresh_prod fresh_left calc_atm) +apply(drule_tac pi="[(a,b)]" and x="(x):M{aa:=(z).([(a,b)]\P)}" + in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(perm_simp add: calc_atm csubst_eqvt CAND_eqvt_coname) +apply(drule meta_mp) +apply(auto)[1] +apply(drule_tac pi="[(a,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_coname) +apply(drule_tac x="[(a,b)]\c" in spec) +apply(drule_tac x="[(a,b)]\Q" in spec) +apply(simp add: fresh_prod fresh_left) +apply(drule mp) +apply(simp add: calc_atm) +apply(drule_tac pi="[(a,b)]" and x=":M{x:=<([(a,b)]\c)>.([(a,b)]\Q)}" + in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(perm_simp add: nsubst_eqvt CAND_eqvt_coname) +apply(simp add: calc_atm) +apply(simp) +done + +lemma IMPLEFT_elim: + assumes a: "(x):M \ IMPLEFT (B IMP C) (\\) (\(C)\)" + obtains x' a' M' N' where "M = ImpL .M' (x').N' x" and "fin (ImpL .M' (x').N' x) x" + and ":M' \ \\" and "(x'):N' \ \(C)\" +using a +apply(auto simp add: ntrm.inject alpha abs_fresh calc_atm) +apply(drule_tac x="a" in meta_spec) +apply(drule_tac x="[(x,y)]\M" in meta_spec) +apply(drule_tac x="y" in meta_spec) +apply(drule_tac x="[(x,y)]\N" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(x,y)]" in fin.eqvt(1)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(x,y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(drule meta_mp) +apply(drule_tac pi="[(x,y)]" and x="(x):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(perm_simp add: calc_atm CAND_eqvt_name) +apply(simp) +apply(case_tac "x=xa") +apply(simp) +apply(drule_tac x="a" in meta_spec) +apply(drule_tac x="[(xa,y)]\M" in meta_spec) +apply(drule_tac x="y" in meta_spec) +apply(drule_tac x="[(xa,y)]\N" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(xa,y)]" in fin.eqvt(1)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(xa,y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(drule meta_mp) +apply(drule_tac pi="[(xa,y)]" and x="(xa):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(simp) +apply(simp) +apply(case_tac "y=xa") +apply(simp) +apply(drule_tac x="a" in meta_spec) +apply(drule_tac x="[(x,xa)]\M" in meta_spec) +apply(drule_tac x="x" in meta_spec) +apply(drule_tac x="[(x,xa)]\N" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(x,xa)]" in fin.eqvt(1)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(x,xa)]" and x=":M" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(drule meta_mp) +apply(drule_tac pi="[(x,xa)]" and x="(xa):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(simp) +apply(simp) +apply(drule_tac x="a" in meta_spec) +apply(drule_tac x="[(x,y)]\M" in meta_spec) +apply(drule_tac x="xa" in meta_spec) +apply(drule_tac x="[(x,y)]\N" in meta_spec) +apply(simp) +apply(drule meta_mp) +apply(drule_tac pi="[(x,y)]" in fin.eqvt(1)) +apply(simp add: calc_atm) +apply(drule meta_mp) +apply(drule_tac pi="[(x,y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(drule meta_mp) +apply(drule_tac pi="[(x,y)]" and x="(xa):N" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm CAND_eqvt_name) +apply(simp) +done + +lemma CANDs_alpha: + shows ":M \ (\\) \ [a].M = [b].N \ :N \ (\\)" + and "(x):M \ (\(B)\) \ [x].M = [y].N \ (y):N \ (\(B)\)" +apply(auto simp add: alpha) +apply(drule_tac pi="[(a,b)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(perm_simp add: CAND_eqvt_coname calc_atm) +apply(drule_tac pi="[(x,y)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(perm_simp add: CAND_eqvt_name calc_atm) +done + +lemma CAND_NotR_elim: + assumes a: ":NotR (x).M a \ (\\)" ":NotR (x).M a \ BINDINGc B (\(B)\)" + shows "\B'. B = NOT B' \ (x):M \ (\(B')\)" +using a +apply(nominal_induct B rule: ty.induct) +apply(simp_all add: ty.inject AXIOMSc_def ctrm.inject alpha) +apply(auto intro: CANDs_alpha simp add: trm.inject calc_atm abs_fresh fresh_atm) +apply(drule_tac pi="[(a,aa)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(auto simp add: CAND_eqvt_coname calc_atm intro: CANDs_alpha) +done + +lemma CAND_NotL_elim_aux: + assumes a: "(x):NotL .M x \ NEGn B (\\)" "(x):NotL .M x \ BINDINGn B (\\)" + shows "\B'. B = NOT B' \ :M \ (\\)" +using a +apply(nominal_induct B rule: ty.induct) +apply(simp_all add: ty.inject AXIOMSn_def ntrm.inject alpha) +apply(auto intro: CANDs_alpha simp add: trm.inject calc_atm abs_fresh fresh_atm) +apply(drule_tac pi="[(x,xa)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(auto simp add: CAND_eqvt_name calc_atm intro: CANDs_alpha) +done + +lemmas CAND_NotL_elim = CAND_NotL_elim_aux[OF NEG_elim(2)] + +lemma CAND_AndR_elim: + assumes a: ":AndR .M .N a \ (\\)" ":AndR .M .N a \ BINDINGc B (\(B)\)" + shows "\B1 B2. B = B1 AND B2 \ :M \ (\\) \ :N \ (\\)" +using a +apply(nominal_induct B rule: ty.induct) +apply(simp_all add: ty.inject AXIOMSc_def ctrm.inject alpha) +apply(auto intro: CANDs_alpha simp add: trm.inject calc_atm abs_fresh fresh_atm) +apply(drule_tac pi="[(a,ca)]" and x=":Ma" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_coname calc_atm) +apply(auto intro: CANDs_alpha)[1] +apply(drule_tac pi="[(a,ca)]" and x=":Na" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_coname calc_atm) +apply(auto intro: CANDs_alpha)[1] +apply(drule_tac pi="[(a,ca)]" and x=":Ma" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_coname calc_atm) +apply(auto intro: CANDs_alpha)[1] +apply(case_tac "a=ba") +apply(simp) +apply(drule_tac pi="[(ba,ca)]" and x=":Na" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_coname calc_atm) +apply(auto intro: CANDs_alpha)[1] +apply(simp) +apply(case_tac "ca=ba") +apply(simp) +apply(drule_tac pi="[(a,ba)]" and x=":Na" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_coname calc_atm) +apply(auto intro: CANDs_alpha)[1] +apply(simp) +apply(drule_tac pi="[(a,ca)]" and x=":Na" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_coname calc_atm) +apply(auto intro: CANDs_alpha)[1] +apply(case_tac "a=aa") +apply(simp) +apply(drule_tac pi="[(aa,ca)]" and x=":Ma" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_coname calc_atm) +apply(auto intro: CANDs_alpha)[1] +apply(simp) +apply(case_tac "ca=aa") +apply(simp) +apply(drule_tac pi="[(a,aa)]" and x=":Ma" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_coname calc_atm) +apply(auto intro: CANDs_alpha)[1] +apply(simp) +apply(drule_tac pi="[(a,ca)]" and x=":Ma" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_coname calc_atm) +apply(auto intro: CANDs_alpha)[1] +apply(drule_tac pi="[(a,ca)]" and x=":Na" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_coname calc_atm) +apply(auto intro: CANDs_alpha)[1] +apply(case_tac "a=aa") +apply(simp) +apply(drule_tac pi="[(aa,ca)]" and x=":Ma" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_coname calc_atm) +apply(auto intro: CANDs_alpha)[1] +apply(simp) +apply(case_tac "ca=aa") +apply(simp) +apply(drule_tac pi="[(a,aa)]" and x=":Ma" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_coname calc_atm) +apply(auto intro: CANDs_alpha)[1] +apply(simp) +apply(drule_tac pi="[(a,ca)]" and x=":Ma" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_coname calc_atm) +apply(auto intro: CANDs_alpha)[1] +apply(case_tac "a=ba") +apply(simp) +apply(drule_tac pi="[(ba,ca)]" and x=":Na" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_coname calc_atm) +apply(auto intro: CANDs_alpha)[1] +apply(simp) +apply(case_tac "ca=ba") +apply(simp) +apply(drule_tac pi="[(a,ba)]" and x=":Na" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_coname calc_atm) +apply(auto intro: CANDs_alpha)[1] +apply(simp) +apply(drule_tac pi="[(a,ca)]" and x=":Na" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_coname calc_atm) +apply(auto intro: CANDs_alpha)[1] +done + +lemma CAND_OrR1_elim: + assumes a: ":OrR1 .M a \ (\\)" ":OrR1 .M a \ BINDINGc B (\(B)\)" + shows "\B1 B2. B = B1 OR B2 \ :M \ (\\)" +using a +apply(nominal_induct B rule: ty.induct) +apply(simp_all add: ty.inject AXIOMSc_def ctrm.inject alpha) +apply(auto intro: CANDs_alpha simp add: trm.inject calc_atm abs_fresh fresh_atm) +apply(drule_tac pi="[(a,ba)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(auto simp add: CAND_eqvt_coname calc_atm intro: CANDs_alpha) +apply(case_tac "a=aa") +apply(simp) +apply(drule_tac pi="[(aa,ba)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(auto simp add: CAND_eqvt_coname calc_atm intro: CANDs_alpha) +apply(case_tac "ba=aa") +apply(simp) +apply(drule_tac pi="[(a,aa)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(auto simp add: CAND_eqvt_coname calc_atm intro: CANDs_alpha) +apply(drule_tac pi="[(a,ba)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(auto simp add: CAND_eqvt_coname calc_atm intro: CANDs_alpha) +done + +lemma CAND_OrR2_elim: + assumes a: ":OrR2 .M a \ (\\)" ":OrR2 .M a \ BINDINGc B (\(B)\)" + shows "\B1 B2. B = B1 OR B2 \ :M \ (\\)" +using a +apply(nominal_induct B rule: ty.induct) +apply(simp_all add: ty.inject AXIOMSc_def ctrm.inject alpha) +apply(auto intro: CANDs_alpha simp add: trm.inject calc_atm abs_fresh fresh_atm) +apply(drule_tac pi="[(a,ba)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(auto simp add: CAND_eqvt_coname calc_atm intro: CANDs_alpha) +apply(case_tac "a=aa") +apply(simp) +apply(drule_tac pi="[(aa,ba)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(auto simp add: CAND_eqvt_coname calc_atm intro: CANDs_alpha) +apply(case_tac "ba=aa") +apply(simp) +apply(drule_tac pi="[(a,aa)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(auto simp add: CAND_eqvt_coname calc_atm intro: CANDs_alpha) +apply(drule_tac pi="[(a,ba)]" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(auto simp add: CAND_eqvt_coname calc_atm intro: CANDs_alpha) +done + +lemma CAND_OrL_elim_aux: + assumes a: "(x):(OrL (y).M (z).N x) \ NEGn B (\\)" "(x):(OrL (y).M (z).N x) \ BINDINGn B (\\)" + shows "\B1 B2. B = B1 OR B2 \ (y):M \ (\(B1)\) \ (z):N \ (\(B2)\)" +using a +apply(nominal_induct B rule: ty.induct) +apply(simp_all add: ty.inject AXIOMSn_def ntrm.inject alpha) +apply(auto intro: CANDs_alpha simp add: trm.inject calc_atm abs_fresh fresh_atm) +apply(drule_tac pi="[(x,za)]" and x="(x):Ma" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name calc_atm) +apply(auto intro: CANDs_alpha)[1] +apply(drule_tac pi="[(x,za)]" and x="(x):Nb" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name calc_atm) +apply(auto intro: CANDs_alpha)[1] +apply(drule_tac pi="[(x,za)]" and x="(x):Ma" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name calc_atm) +apply(auto intro: CANDs_alpha)[1] +apply(case_tac "x=ya") +apply(simp) +apply(drule_tac pi="[(ya,za)]" and x="(ya):Nb" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name calc_atm) +apply(auto intro: CANDs_alpha)[1] +apply(simp) +apply(case_tac "za=ya") +apply(simp) +apply(drule_tac pi="[(x,ya)]" and x="(ya):Nb" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name calc_atm) +apply(auto intro: CANDs_alpha)[1] +apply(simp) +apply(drule_tac pi="[(x,za)]" and x="(ya):Nb" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name calc_atm) +apply(auto intro: CANDs_alpha)[1] +apply(case_tac "x=xa") +apply(simp) +apply(drule_tac pi="[(xa,za)]" and x="(xa):Ma" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name calc_atm) +apply(auto intro: CANDs_alpha)[1] +apply(simp) +apply(case_tac "za=xa") +apply(simp) +apply(drule_tac pi="[(x,xa)]" and x="(xa):Ma" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name calc_atm) +apply(auto intro: CANDs_alpha)[1] +apply(simp) +apply(drule_tac pi="[(x,za)]" and x="(xa):Ma" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name calc_atm) +apply(auto intro: CANDs_alpha)[1] +apply(drule_tac pi="[(x,za)]" and x="(x):Nb" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name calc_atm) +apply(auto intro: CANDs_alpha)[1] +apply(case_tac "x=xa") +apply(simp) +apply(drule_tac pi="[(xa,za)]" and x="(xa):Ma" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name calc_atm) +apply(auto intro: CANDs_alpha)[1] +apply(simp) +apply(case_tac "za=xa") +apply(simp) +apply(drule_tac pi="[(x,xa)]" and x="(xa):Ma" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name calc_atm) +apply(auto intro: CANDs_alpha)[1] +apply(simp) +apply(drule_tac pi="[(x,za)]" and x="(xa):Ma" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name calc_atm) +apply(auto intro: CANDs_alpha)[1] +apply(case_tac "x=ya") +apply(simp) +apply(drule_tac pi="[(ya,za)]" and x="(ya):Nb" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name calc_atm) +apply(auto intro: CANDs_alpha)[1] +apply(simp) +apply(case_tac "za=ya") +apply(simp) +apply(drule_tac pi="[(x,ya)]" and x="(ya):Nb" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name calc_atm) +apply(auto intro: CANDs_alpha)[1] +apply(simp) +apply(drule_tac pi="[(x,za)]" and x="(ya):Nb" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name calc_atm) +apply(auto intro: CANDs_alpha)[1] +done + +lemmas CAND_OrL_elim = CAND_OrL_elim_aux[OF NEG_elim(2)] + +lemma CAND_AndL1_elim_aux: + assumes a: "(x):(AndL1 (y).M x) \ NEGn B (\\)" "(x):(AndL1 (y).M x) \ BINDINGn B (\\)" + shows "\B1 B2. B = B1 AND B2 \ (y):M \ (\(B1)\)" +using a +apply(nominal_induct B rule: ty.induct) +apply(simp_all add: ty.inject AXIOMSn_def ntrm.inject alpha) +apply(auto intro: CANDs_alpha simp add: trm.inject calc_atm abs_fresh fresh_atm) +apply(drule_tac pi="[(x,ya)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(auto simp add: CAND_eqvt_name calc_atm intro: CANDs_alpha) +apply(case_tac "x=xa") +apply(simp) +apply(drule_tac pi="[(xa,ya)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(auto simp add: CAND_eqvt_name calc_atm intro: CANDs_alpha) +apply(case_tac "ya=xa") +apply(simp) +apply(drule_tac pi="[(x,xa)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(auto simp add: CAND_eqvt_name calc_atm intro: CANDs_alpha) +apply(drule_tac pi="[(x,ya)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(auto simp add: CAND_eqvt_name calc_atm intro: CANDs_alpha) +done + +lemmas CAND_AndL1_elim = CAND_AndL1_elim_aux[OF NEG_elim(2)] + +lemma CAND_AndL2_elim_aux: + assumes a: "(x):(AndL2 (y).M x) \ NEGn B (\\)" "(x):(AndL2 (y).M x) \ BINDINGn B (\\)" + shows "\B1 B2. B = B1 AND B2 \ (y):M \ (\(B2)\)" +using a +apply(nominal_induct B rule: ty.induct) +apply(simp_all add: ty.inject AXIOMSn_def ntrm.inject alpha) +apply(auto intro: CANDs_alpha simp add: trm.inject calc_atm abs_fresh fresh_atm) +apply(drule_tac pi="[(x,ya)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(auto simp add: CAND_eqvt_name calc_atm intro: CANDs_alpha) +apply(case_tac "x=xa") +apply(simp) +apply(drule_tac pi="[(xa,ya)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(auto simp add: CAND_eqvt_name calc_atm intro: CANDs_alpha) +apply(case_tac "ya=xa") +apply(simp) +apply(drule_tac pi="[(x,xa)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(auto simp add: CAND_eqvt_name calc_atm intro: CANDs_alpha) +apply(drule_tac pi="[(x,ya)]" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(auto simp add: CAND_eqvt_name calc_atm intro: CANDs_alpha) +done + +lemmas CAND_AndL2_elim = CAND_AndL2_elim_aux[OF NEG_elim(2)] + +lemma CAND_ImpL_elim_aux: + assumes a: "(x):(ImpL .M (z).N x) \ NEGn B (\\)" "(x):(ImpL .M (z).N x) \ BINDINGn B (\\)" + shows "\B1 B2. B = B1 IMP B2 \ :M \ (\\) \ (z):N \ (\(B2)\)" +using a +apply(nominal_induct B rule: ty.induct) +apply(simp_all add: ty.inject AXIOMSn_def ntrm.inject alpha) +apply(auto intro: CANDs_alpha simp add: trm.inject calc_atm abs_fresh fresh_atm) +apply(drule_tac pi="[(x,y)]" and x=":Ma" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name calc_atm) +apply(auto intro: CANDs_alpha)[1] +apply(drule_tac pi="[(x,y)]" and x="(x):Nb" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name calc_atm) +apply(auto intro: CANDs_alpha)[1] +apply(drule_tac pi="[(x,y)]" and x=":Ma" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name calc_atm) +apply(auto intro: CANDs_alpha)[1] +apply(case_tac "x=xa") +apply(simp) +apply(drule_tac pi="[(xa,y)]" and x="(xa):Nb" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name calc_atm) +apply(auto intro: CANDs_alpha)[1] +apply(simp) +apply(case_tac "y=xa") +apply(simp) +apply(drule_tac pi="[(x,xa)]" and x="(xa):Nb" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name calc_atm) +apply(auto intro: CANDs_alpha)[1] +apply(simp) +apply(drule_tac pi="[(x,y)]" and x="(xa):Nb" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name calc_atm) +apply(auto intro: CANDs_alpha)[1] +done + +lemmas CAND_ImpL_elim = CAND_ImpL_elim_aux[OF NEG_elim(2)] + +lemma CAND_ImpR_elim: + assumes a: ":ImpR (x)..M a \ (\\)" ":ImpR (x)..M a \ BINDINGc B (\(B)\)" + shows "\B1 B2. B = B1 IMP B2 \ + (\z P. x\(z,P) \ (z):P \ \(B2)\ \ (x):(M{b:=(z).P}) \ \(B1)\) \ + (\c Q. b\(c,Q) \ :Q \ \\ \ :(M{x:=.Q}) \ \\)" +using a +apply(nominal_induct B rule: ty.induct) +apply(simp_all add: ty.inject AXIOMSc_def ctrm.inject alpha) +apply(auto intro: CANDs_alpha simp add: trm.inject calc_atm abs_fresh fresh_atm fresh_prod fresh_bij) +apply(generate_fresh "name") +apply(generate_fresh "coname") +apply(drule_tac a="ca" and z="c" in alpha_name_coname) +apply(simp) +apply(simp) +apply(simp) +apply(drule_tac x="[(xa,c)]\[(aa,ca)]\[(b,ca)]\[(x,c)]\z" in spec) +apply(drule_tac x="[(xa,c)]\[(aa,ca)]\[(b,ca)]\[(x,c)]\P" in spec) +apply(drule mp) +apply(rule conjI) +apply(auto simp add: calc_atm fresh_prod fresh_atm)[1] +apply(rule conjI) +apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(drule_tac pi="[(x,c)]" and X="\(ty2)\" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(b,ca)]" and X="\(ty2)\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(aa,ca)]" and X="\(ty2)\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(xa,c)]" and X="\(ty2)\" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(xa,c)]" and X="\(ty1)\" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt) +apply(drule_tac pi="[(aa,ca)]" and X="\(ty1)\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt) +apply(drule_tac pi="[(b,ca)]" and X="\(ty1)\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(x,c)]" and X="\(ty1)\" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(perm_simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt) +apply(auto simp add: calc_atm fresh_prod fresh_atm)[1] +apply(generate_fresh "name") +apply(generate_fresh "coname") +apply(drule_tac a="cb" and z="ca" in alpha_name_coname) +apply(simp) +apply(simp) +apply(simp) +apply(drule_tac x="[(xa,ca)]\[(aa,cb)]\[(b,cb)]\[(x,ca)]\c" in spec) +apply(drule_tac x="[(xa,ca)]\[(aa,cb)]\[(b,cb)]\[(x,ca)]\Q" in spec) +apply(drule mp) +apply(rule conjI) +apply(auto simp add: calc_atm fresh_prod fresh_atm)[1] +apply(rule conjI) +apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(drule_tac pi="[(x,ca)]" and X="\\" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(b,cb)]" and X="\\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(aa,cb)]" and X="\\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(xa,ca)]" and X="\\" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(xa,ca)]" and X="\\" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt) +apply(drule_tac pi="[(aa,cb)]" and X="\\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt) +apply(drule_tac pi="[(b,cb)]" and X="\\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(x,ca)]" and X="\\" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(perm_simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt) +apply(auto simp add: calc_atm fresh_prod fresh_atm)[1] +apply(generate_fresh "name") +apply(generate_fresh "coname") +apply(drule_tac a="ca" and z="c" in alpha_name_coname) +apply(simp add: fresh_left calc_atm fresh_prod fresh_atm) +apply(simp add: fresh_left calc_atm fresh_prod fresh_atm) +apply(auto)[1] +apply(simp) +apply(drule_tac x="[(a,ba)]\[(xa,c)]\[(ba,ca)]\[(b,ca)]\[(x,c)]\z" in spec) +apply(drule_tac x="[(a,ba)]\[(xa,c)]\[(ba,ca)]\[(b,ca)]\[(x,c)]\P" in spec) +apply(drule mp) +apply(rule conjI) +apply(auto simp add: calc_atm fresh_prod fresh_atm)[1] +apply(rule conjI) +apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(drule_tac pi="[(x,c)]" and X="\(ty2)\" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(b,ca)]" and X="\(ty2)\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(ba,ca)]" and X="\(ty2)\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(xa,c)]" and X="\(ty2)\" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(a,ba)]" and X="\(ty2)\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(a,ba)]" and X="\(ty1)\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt) +apply(drule_tac pi="[(xa,c)]" and X="\(ty1)\" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt) +apply(drule_tac pi="[(ba,ca)]" and X="\(ty1)\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt) +apply(drule_tac pi="[(b,ca)]" and X="\(ty1)\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(x,c)]" and X="\(ty1)\" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(perm_simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt) +apply(auto simp add: calc_atm fresh_prod fresh_atm)[1] +apply(generate_fresh "name") +apply(generate_fresh "coname") +apply(drule_tac a="cb" and z="ca" in alpha_name_coname) +apply(simp add: fresh_left calc_atm fresh_prod fresh_atm) +apply(simp add: fresh_left calc_atm fresh_prod fresh_atm) +apply(auto)[1] +apply(simp) +apply(drule_tac x="[(a,ba)]\[(xa,ca)]\[(ba,cb)]\[(b,cb)]\[(x,ca)]\c" in spec) +apply(drule_tac x="[(a,ba)]\[(xa,ca)]\[(ba,cb)]\[(b,cb)]\[(x,ca)]\Q" in spec) +apply(drule mp) +apply(rule conjI) +apply(auto simp add: calc_atm fresh_prod fresh_atm)[1] +apply(rule conjI) +apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(drule_tac pi="[(x,ca)]" and X="\\" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(b,cb)]" and X="\\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(ba,cb)]" and X="\\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(xa,ca)]" and X="\\" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(a,ba)]" and X="\\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(a,ba)]" and X="\\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt) +apply(drule_tac pi="[(xa,ca)]" and X="\\" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt) +apply(drule_tac pi="[(ba,cb)]" and X="\\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt) +apply(drule_tac pi="[(b,cb)]" and X="\\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(x,ca)]" and X="\\" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(perm_simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt) +apply(auto simp add: calc_atm fresh_prod fresh_atm)[1] +apply(case_tac "a=aa") +apply(simp) +apply(generate_fresh "name") +apply(generate_fresh "coname") +apply(drule_tac a="ca" and z="c" in alpha_name_coname) +apply(simp add: fresh_left calc_atm fresh_prod fresh_atm) +apply(simp add: fresh_left calc_atm fresh_prod fresh_atm) +apply(auto)[1] +apply(simp) +apply(drule_tac x="[(aa,ba)]\[(xa,c)]\[(ba,ca)]\[(b,ca)]\[(x,c)]\z" in spec) +apply(drule_tac x="[(aa,ba)]\[(xa,c)]\[(ba,ca)]\[(b,ca)]\[(x,c)]\P" in spec) +apply(drule mp) +apply(rule conjI) +apply(auto simp add: calc_atm fresh_prod fresh_atm)[1] +apply(rule conjI) +apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(drule_tac pi="[(x,c)]" and X="\(ty2)\" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(b,ca)]" and X="\(ty2)\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(ba,ca)]" and X="\(ty2)\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(xa,c)]" and X="\(ty2)\" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(aa,ba)]" and X="\(ty2)\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(aa,ba)]" and X="\(ty1)\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt) +apply(drule_tac pi="[(xa,c)]" and X="\(ty1)\" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt) +apply(drule_tac pi="[(ba,ca)]" and X="\(ty1)\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt) +apply(drule_tac pi="[(b,ca)]" and X="\(ty1)\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(x,c)]" and X="\(ty1)\" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(perm_simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt) +apply(auto simp add: calc_atm fresh_prod fresh_atm)[1] +apply(simp) +apply(case_tac "ba=aa") +apply(simp) +apply(generate_fresh "name") +apply(generate_fresh "coname") +apply(drule_tac a="ca" and z="c" in alpha_name_coname) +apply(simp add: fresh_left calc_atm fresh_prod fresh_atm) +apply(simp add: fresh_left calc_atm fresh_prod fresh_atm) +apply(auto)[1] +apply(simp) +apply(drule_tac x="[(a,aa)]\[(xa,c)]\[(a,ca)]\[(b,ca)]\[(x,c)]\z" in spec) +apply(drule_tac x="[(a,aa)]\[(xa,c)]\[(a,ca)]\[(b,ca)]\[(x,c)]\P" in spec) +apply(drule mp) +apply(rule conjI) +apply(auto simp add: calc_atm fresh_prod fresh_atm)[1] +apply(rule conjI) +apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(drule_tac pi="[(x,c)]" and X="\(ty2)\" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(b,ca)]" and X="\(ty2)\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(a,ca)]" and X="\(ty2)\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(xa,c)]" and X="\(ty2)\" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(a,aa)]" and X="\(ty2)\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(a,aa)]" and X="\(ty1)\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt) +apply(drule_tac pi="[(xa,c)]" and X="\(ty1)\" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt) +apply(drule_tac pi="[(a,ca)]" and X="\(ty1)\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt) +apply(drule_tac pi="[(b,ca)]" and X="\(ty1)\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(x,c)]" and X="\(ty1)\" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(perm_simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt) +apply(auto simp add: calc_atm fresh_prod fresh_atm)[1] +apply(simp) +apply(generate_fresh "name") +apply(generate_fresh "coname") +apply(drule_tac a="ca" and z="c" in alpha_name_coname) +apply(simp add: fresh_left calc_atm fresh_prod fresh_atm) +apply(simp add: fresh_left calc_atm fresh_prod fresh_atm) +apply(auto)[1] +apply(simp) +apply(drule_tac x="[(a,ba)]\[(xa,c)]\[(aa,ca)]\[(b,ca)]\[(x,c)]\z" in spec) +apply(drule_tac x="[(a,ba)]\[(xa,c)]\[(aa,ca)]\[(b,ca)]\[(x,c)]\P" in spec) +apply(drule mp) +apply(rule conjI) +apply(auto simp add: calc_atm fresh_prod fresh_atm)[1] +apply(rule conjI) +apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(drule_tac pi="[(x,c)]" and X="\(ty2)\" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(b,ca)]" and X="\(ty2)\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(aa,ca)]" and X="\(ty2)\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(xa,c)]" and X="\(ty2)\" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(a,ba)]" and X="\(ty2)\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(a,ba)]" and X="\(ty1)\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt) +apply(drule_tac pi="[(xa,c)]" and X="\(ty1)\" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt) +apply(drule_tac pi="[(aa,ca)]" and X="\(ty1)\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt) +apply(drule_tac pi="[(b,ca)]" and X="\(ty1)\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(x,c)]" and X="\(ty1)\" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(perm_simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt) +apply(auto simp add: calc_atm fresh_prod fresh_atm)[1] +apply(case_tac "a=aa") +apply(simp) +apply(generate_fresh "name") +apply(generate_fresh "coname") +apply(drule_tac a="cb" and z="ca" in alpha_name_coname) +apply(simp add: fresh_left calc_atm fresh_prod fresh_atm) +apply(simp add: fresh_left calc_atm fresh_prod fresh_atm) +apply(auto)[1] +apply(simp) +apply(drule_tac x="[(aa,ba)]\[(xa,ca)]\[(ba,cb)]\[(b,cb)]\[(x,ca)]\c" in spec) +apply(drule_tac x="[(aa,ba)]\[(xa,ca)]\[(ba,cb)]\[(b,cb)]\[(x,ca)]\Q" in spec) +apply(drule mp) +apply(rule conjI) +apply(auto simp add: calc_atm fresh_prod fresh_atm)[1] +apply(rule conjI) +apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(drule_tac pi="[(x,ca)]" and X="\\" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(b,cb)]" and X="\\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(ba,cb)]" and X="\\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(xa,ca)]" and X="\\" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(aa,ba)]" and X="\\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(aa,ba)]" and X="\\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt) +apply(drule_tac pi="[(xa,ca)]" and X="\\" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt) +apply(drule_tac pi="[(ba,cb)]" and X="\\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt) +apply(drule_tac pi="[(b,cb)]" and X="\\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(x,ca)]" and X="\\" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(perm_simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt) +apply(auto simp add: calc_atm fresh_prod fresh_atm)[1] +apply(simp) +apply(case_tac "ba=aa") +apply(simp) +apply(generate_fresh "name") +apply(generate_fresh "coname") +apply(drule_tac a="cb" and z="ca" in alpha_name_coname) +apply(simp add: fresh_left calc_atm fresh_prod fresh_atm) +apply(simp add: fresh_left calc_atm fresh_prod fresh_atm) +apply(auto)[1] +apply(simp) +apply(drule_tac x="[(a,aa)]\[(xa,ca)]\[(a,cb)]\[(b,cb)]\[(x,ca)]\c" in spec) +apply(drule_tac x="[(a,aa)]\[(xa,ca)]\[(a,cb)]\[(b,cb)]\[(x,ca)]\Q" in spec) +apply(drule mp) +apply(rule conjI) +apply(auto simp add: calc_atm fresh_prod fresh_atm)[1] +apply(rule conjI) +apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(drule_tac pi="[(x,ca)]" and X="\\" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(b,cb)]" and X="\\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(a,cb)]" and X="\\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(xa,ca)]" and X="\\" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(a,aa)]" and X="\\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(a,aa)]" and X="\\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt) +apply(drule_tac pi="[(xa,ca)]" and X="\\" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt) +apply(drule_tac pi="[(a,cb)]" and X="\\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt) +apply(drule_tac pi="[(b,cb)]" and X="\\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(x,ca)]" and X="\\" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(perm_simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt) +apply(auto simp add: calc_atm fresh_prod fresh_atm)[1] +apply(simp) +apply(generate_fresh "name") +apply(generate_fresh "coname") +apply(drule_tac a="cb" and z="ca" in alpha_name_coname) +apply(simp add: fresh_left calc_atm fresh_prod fresh_atm) +apply(simp add: fresh_left calc_atm fresh_prod fresh_atm) +apply(auto)[1] +apply(simp) +apply(drule_tac x="[(a,ba)]\[(xa,ca)]\[(aa,cb)]\[(b,cb)]\[(x,ca)]\c" in spec) +apply(drule_tac x="[(a,ba)]\[(xa,ca)]\[(aa,cb)]\[(b,cb)]\[(x,ca)]\Q" in spec) +apply(drule mp) +apply(rule conjI) +apply(auto simp add: calc_atm fresh_prod fresh_atm)[1] +apply(rule conjI) +apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(drule_tac pi="[(x,ca)]" and X="\\" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(b,cb)]" and X="\\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(aa,cb)]" and X="\\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(xa,ca)]" and X="\\" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(a,ba)]" and X="\\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(a,ba)]" and X="\\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt) +apply(drule_tac pi="[(xa,ca)]" and X="\\" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt) +apply(drule_tac pi="[(aa,cb)]" and X="\\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt) +apply(drule_tac pi="[(b,cb)]" and X="\\" in pt_set_bij2[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: CAND_eqvt_name CAND_eqvt_coname) +apply(drule_tac pi="[(x,ca)]" and X="\\" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(perm_simp add: CAND_eqvt_name CAND_eqvt_coname csubst_eqvt nsubst_eqvt) +apply(auto simp add: calc_atm fresh_prod fresh_atm)[1] +done + +text {* Main lemma 1 *} + +lemma AXIOMS_imply_SNa: + shows ":M \ AXIOMSc B \ SNa M" + and "(x):M \ AXIOMSn B \ SNa M" +apply - +apply(auto simp add: AXIOMSn_def AXIOMSc_def ntrm.inject ctrm.inject alpha) +apply(rule Ax_in_SNa)+ +done + +lemma BINDING_imply_SNa: + shows ":M \ BINDINGc B (\(B)\) \ SNa M" + and "(x):M \ BINDINGn B (\\) \ SNa M" +apply - +apply(auto simp add: BINDINGn_def BINDINGc_def ntrm.inject ctrm.inject alpha) +apply(drule_tac x="x" in spec) +apply(drule_tac x="Ax x a" in spec) +apply(drule mp) +apply(rule Ax_in_CANDs) +apply(drule a_star_preserves_SNa) +apply(rule subst_with_ax2) +apply(simp add: crename_id) +apply(drule_tac x="x" in spec) +apply(drule_tac x="Ax x aa" in spec) +apply(drule mp) +apply(rule Ax_in_CANDs) +apply(drule a_star_preserves_SNa) +apply(rule subst_with_ax2) +apply(simp add: crename_id SNa_eqvt) +apply(drule_tac x="a" in spec) +apply(drule_tac x="Ax x a" in spec) +apply(drule mp) +apply(rule Ax_in_CANDs) +apply(drule a_star_preserves_SNa) +apply(rule subst_with_ax1) +apply(simp add: nrename_id) +apply(drule_tac x="a" in spec) +apply(drule_tac x="Ax xa a" in spec) +apply(drule mp) +apply(rule Ax_in_CANDs) +apply(drule a_star_preserves_SNa) +apply(rule subst_with_ax1) +apply(simp add: nrename_id SNa_eqvt) +done + +lemma CANDs_imply_SNa: + shows ":M \ \\ \ SNa M" + and "(x):M \ \(B)\ \ SNa M" +proof(induct B arbitrary: a x M rule: ty.weak_induct) + case (PR X) + { case 1 + have ":M \ \\" by fact + then have ":M \ NEGc (PR X) (\(PR X)\)" by simp + then have ":M \ AXIOMSc (PR X) \ BINDINGc (PR X) (\(PR X)\)" by simp + moreover + { assume ":M \ AXIOMSc (PR X)" + then have "SNa M" by (simp add: AXIOMS_imply_SNa) + } + moreover + { assume ":M \ BINDINGc (PR X) (\(PR X)\)" + then have "SNa M" by (simp add: BINDING_imply_SNa) + } + ultimately show "SNa M" by blast + next + case 2 + have "(x):M \ (\(PR X)\)" by fact + then have "(x):M \ NEGn (PR X) (\\)" using NEG_simp by blast + then have "(x):M \ AXIOMSn (PR X) \ BINDINGn (PR X) (\\)" by simp + moreover + { assume "(x):M \ AXIOMSn (PR X)" + then have "SNa M" by (simp add: AXIOMS_imply_SNa) + } + moreover + { assume "(x):M \ BINDINGn (PR X) (\\)" + then have "SNa M" by (simp only: BINDING_imply_SNa) + } + ultimately show "SNa M" by blast + } +next + case (NOT B) + have ih1: "\a M. :M \ \\ \ SNa M" by fact + have ih2: "\x M. (x):M \ \(B)\ \ SNa M" by fact + { case 1 + have ":M \ (\\)" by fact + then have ":M \ NEGc (NOT B) (\(NOT B)\)" by simp + then have ":M \ AXIOMSc (NOT B) \ BINDINGc (NOT B) (\(NOT B)\) \ NOTRIGHT (NOT B) (\(B)\)" by simp + moreover + { assume ":M \ AXIOMSc (NOT B)" + then have "SNa M" by (simp add: AXIOMS_imply_SNa) + } + moreover + { assume ":M \ BINDINGc (NOT B) (\(NOT B)\)" + then have "SNa M" by (simp only: BINDING_imply_SNa) + } + moreover + { assume ":M \ NOTRIGHT (NOT B) (\(B)\)" + then obtain x' M' where eq: "M = NotR (x').M' a" and "(x'):M' \ (\(B)\)" + using NOTRIGHT_elim by blast + then have "SNa M'" using ih2 by blast + then have "SNa M" using eq by (simp add: NotR_in_SNa) + } + ultimately show "SNa M" by blast + next + case 2 + have "(x):M \ (\(NOT B)\)" by fact + then have "(x):M \ NEGn (NOT B) (\\)" using NEG_simp by blast + then have "(x):M \ AXIOMSn (NOT B) \ BINDINGn (NOT B) (\\) \ NOTLEFT (NOT B) (\\)" + by (simp only: NEGn.simps) + moreover + { assume "(x):M \ AXIOMSn (NOT B)" + then have "SNa M" by (simp add: AXIOMS_imply_SNa) + } + moreover + { assume "(x):M \ BINDINGn (NOT B) (\\)" + then have "SNa M" by (simp only: BINDING_imply_SNa) + } + moreover + { assume "(x):M \ NOTLEFT (NOT B) (\\)" + then obtain a' M' where eq: "M = NotL .M' x" and ":M' \ (\\)" + using NOTLEFT_elim by blast + then have "SNa M'" using ih1 by blast + then have "SNa M" using eq by (simp add: NotL_in_SNa) + } + ultimately show "SNa M" by blast + } +next + case (AND A B) + have ih1: "\a M. :M \ \\ \ SNa M" by fact + have ih2: "\x M. (x):M \ \(A)\ \ SNa M" by fact + have ih3: "\a M. :M \ \\ \ SNa M" by fact + have ih4: "\x M. (x):M \ \(B)\ \ SNa M" by fact + { case 1 + have ":M \ (\\)" by fact + then have ":M \ NEGc (A AND B) (\(A AND B)\)" by simp + then have ":M \ AXIOMSc (A AND B) \ BINDINGc (A AND B) (\(A AND B)\) + \ ANDRIGHT (A AND B) (\\) (\\)" by simp + moreover + { assume ":M \ AXIOMSc (A AND B)" + then have "SNa M" by (simp add: AXIOMS_imply_SNa) + } + moreover + { assume ":M \ BINDINGc (A AND B) (\(A AND B)\)" + then have "SNa M" by (simp only: BINDING_imply_SNa) + } + moreover + { assume ":M \ ANDRIGHT (A AND B) (\\) (\\)" + then obtain a' M' b' N' where eq: "M = AndR .M' .N' a" + and ":M' \ (\\)" and ":N' \ (\\)" + by (erule_tac ANDRIGHT_elim, blast) + then have "SNa M'" and "SNa N'" using ih1 ih3 by blast+ + then have "SNa M" using eq by (simp add: AndR_in_SNa) + } + ultimately show "SNa M" by blast + next + case 2 + have "(x):M \ (\(A AND B)\)" by fact + then have "(x):M \ NEGn (A AND B) (\\)" using NEG_simp by blast + then have "(x):M \ AXIOMSn (A AND B) \ BINDINGn (A AND B) (\\) + \ ANDLEFT1 (A AND B) (\(A)\) \ ANDLEFT2 (A AND B) (\(B)\)" + by (simp only: NEGn.simps) + moreover + { assume "(x):M \ AXIOMSn (A AND B)" + then have "SNa M" by (simp add: AXIOMS_imply_SNa) + } + moreover + { assume "(x):M \ BINDINGn (A AND B) (\\)" + then have "SNa M" by (simp only: BINDING_imply_SNa) + } + moreover + { assume "(x):M \ ANDLEFT1 (A AND B) (\(A)\)" + then obtain x' M' where eq: "M = AndL1 (x').M' x" and "(x'):M' \ (\(A)\)" + using ANDLEFT1_elim by blast + then have "SNa M'" using ih2 by blast + then have "SNa M" using eq by (simp add: AndL1_in_SNa) + } + moreover + { assume "(x):M \ ANDLEFT2 (A AND B) (\(B)\)" + then obtain x' M' where eq: "M = AndL2 (x').M' x" and "(x'):M' \ (\(B)\)" + using ANDLEFT2_elim by blast + then have "SNa M'" using ih4 by blast + then have "SNa M" using eq by (simp add: AndL2_in_SNa) + } + ultimately show "SNa M" by blast + } +next + case (OR A B) + have ih1: "\a M. :M \ \\ \ SNa M" by fact + have ih2: "\x M. (x):M \ \(A)\ \ SNa M" by fact + have ih3: "\a M. :M \ \\ \ SNa M" by fact + have ih4: "\x M. (x):M \ \(B)\ \ SNa M" by fact + { case 1 + have ":M \ (\\)" by fact + then have ":M \ NEGc (A OR B) (\(A OR B)\)" by simp + then have ":M \ AXIOMSc (A OR B) \ BINDINGc (A OR B) (\(A OR B)\) + \ ORRIGHT1 (A OR B) (\\) \ ORRIGHT2 (A OR B) (\\)" by simp + moreover + { assume ":M \ AXIOMSc (A OR B)" + then have "SNa M" by (simp add: AXIOMS_imply_SNa) + } + moreover + { assume ":M \ BINDINGc (A OR B) (\(A OR B)\)" + then have "SNa M" by (simp only: BINDING_imply_SNa) + } + moreover + { assume ":M \ ORRIGHT1 (A OR B) (\\)" + then obtain a' M' where eq: "M = OrR1 .M' a" + and ":M' \ (\\)" + by (erule_tac ORRIGHT1_elim, blast) + then have "SNa M'" using ih1 by blast + then have "SNa M" using eq by (simp add: OrR1_in_SNa) + } + moreover + { assume ":M \ ORRIGHT2 (A OR B) (\\)" + then obtain a' M' where eq: "M = OrR2 .M' a" and ":M' \ (\\)" + using ORRIGHT2_elim by blast + then have "SNa M'" using ih3 by blast + then have "SNa M" using eq by (simp add: OrR2_in_SNa) + } + ultimately show "SNa M" by blast + next + case 2 + have "(x):M \ (\(A OR B)\)" by fact + then have "(x):M \ NEGn (A OR B) (\\)" using NEG_simp by blast + then have "(x):M \ AXIOMSn (A OR B) \ BINDINGn (A OR B) (\\) + \ ORLEFT (A OR B) (\(A)\) (\(B)\)" + by (simp only: NEGn.simps) + moreover + { assume "(x):M \ AXIOMSn (A OR B)" + then have "SNa M" by (simp add: AXIOMS_imply_SNa) + } + moreover + { assume "(x):M \ BINDINGn (A OR B) (\\)" + then have "SNa M" by (simp only: BINDING_imply_SNa) + } + moreover + { assume "(x):M \ ORLEFT (A OR B) (\(A)\) (\(B)\)" + then obtain x' M' y' N' where eq: "M = OrL (x').M' (y').N' x" + and "(x'):M' \ (\(A)\)" and "(y'):N' \ (\(B)\)" + by (erule_tac ORLEFT_elim, blast) + then have "SNa M'" and "SNa N'" using ih2 ih4 by blast+ + then have "SNa M" using eq by (simp add: OrL_in_SNa) + } + ultimately show "SNa M" by blast + } +next + case (IMP A B) + have ih1: "\a M. :M \ \\ \ SNa M" by fact + have ih2: "\x M. (x):M \ \(A)\ \ SNa M" by fact + have ih3: "\a M. :M \ \\ \ SNa M" by fact + have ih4: "\x M. (x):M \ \(B)\ \ SNa M" by fact + { case 1 + have ":M \ (\\)" by fact + then have ":M \ NEGc (A IMP B) (\(A IMP B)\)" by simp + then have ":M \ AXIOMSc (A IMP B) \ BINDINGc (A IMP B) (\(A IMP B)\) + \ IMPRIGHT (A IMP B) (\(A)\) (\\) (\(B)\) (\\)" by simp + moreover + { assume ":M \ AXIOMSc (A IMP B)" + then have "SNa M" by (simp add: AXIOMS_imply_SNa) + } + moreover + { assume ":M \ BINDINGc (A IMP B) (\(A IMP B)\)" + then have "SNa M" by (simp only: BINDING_imply_SNa) + } + moreover + { assume ":M \ IMPRIGHT (A IMP B) (\(A)\) (\\) (\(B)\) (\\)" + then obtain x' a' M' where eq: "M = ImpR (x')..M' a" + and imp: "\z P. x'\(z,P) \ (z):P \ \(B)\ \ (x'):(M'{a':=(z).P}) \ \(A)\" + by (erule_tac IMPRIGHT_elim, blast) + obtain z::"name" where fs: "z\x'" by (rule_tac exists_fresh, rule fin_supp, blast) + have "(z):Ax z a'\ \(B)\" by (simp add: Ax_in_CANDs) + with imp fs have "(x'):(M'{a':=(z).Ax z a'}) \ \(A)\" by (simp add: fresh_prod fresh_atm) + then have "SNa (M'{a':=(z).Ax z a'})" using ih2 by blast + moreover + have "M'{a':=(z).Ax z a'} \\<^isub>a* M'[a'\c>a']" by (simp add: subst_with_ax2) + ultimately have "SNa (M'[a'\c>a'])" by (simp add: a_star_preserves_SNa) + then have "SNa M'" by (simp add: crename_id) + then have "SNa M" using eq by (simp add: ImpR_in_SNa) + } + ultimately show "SNa M" by blast + next + case 2 + have "(x):M \ (\(A IMP B)\)" by fact + then have "(x):M \ NEGn (A IMP B) (\\)" using NEG_simp by blast + then have "(x):M \ AXIOMSn (A IMP B) \ BINDINGn (A IMP B) (\\) + \ IMPLEFT (A IMP B) (\\) (\(B)\)" + by (simp only: NEGn.simps) + moreover + { assume "(x):M \ AXIOMSn (A IMP B)" + then have "SNa M" by (simp add: AXIOMS_imply_SNa) + } + moreover + { assume "(x):M \ BINDINGn (A IMP B) (\\)" + then have "SNa M" by (simp only: BINDING_imply_SNa) + } + moreover + { assume "(x):M \ IMPLEFT (A IMP B) (\\) (\(B)\)" + then obtain a' M' y' N' where eq: "M = ImpL .M' (y').N' x" + and ":M' \ (\\)" and "(y'):N' \ (\(B)\)" + by (erule_tac IMPLEFT_elim, blast) + then have "SNa M'" and "SNa N'" using ih1 ih4 by blast+ + then have "SNa M" using eq by (simp add: ImpL_in_SNa) + } + ultimately show "SNa M" by blast + } +qed + +text {* Main lemma 2 *} + +lemma AXIOMS_preserved: + shows ":M \ AXIOMSc B \ M \\<^isub>a* M' \ :M' \ AXIOMSc B" + and "(x):M \ AXIOMSn B \ M \\<^isub>a* M' \ (x):M' \ AXIOMSn B" + apply(simp_all add: AXIOMSc_def AXIOMSn_def) + apply(auto simp add: ntrm.inject ctrm.inject alpha) + apply(drule ax_do_not_a_star_reduce) + apply(auto) + apply(drule ax_do_not_a_star_reduce) + apply(auto) + apply(drule ax_do_not_a_star_reduce) + apply(auto) + apply(drule ax_do_not_a_star_reduce) + apply(auto) + done + +lemma BINDING_preserved: + shows ":M \ BINDINGc B (\(B)\) \ M \\<^isub>a* M' \ :M' \ BINDINGc B (\(B)\)" + and "(x):M \ BINDINGn B (\\) \ M \\<^isub>a* M' \ (x):M' \ BINDINGn B (\\)" +proof - + assume red: "M \\<^isub>a* M'" + assume asm: ":M \ BINDINGc B (\(B)\)" + { + fix x::"name" and P::"trm" + from asm have "((x):P) \ (\(B)\) \ SNa (M{a:=(x).P})" by (simp add: BINDINGc_elim) + moreover + have "M{a:=(x).P} \\<^isub>a* M'{a:=(x).P}" using red by (simp add: a_star_subst2) + ultimately + have "((x):P) \ (\(B)\) \ SNa (M'{a:=(x).P})" by (simp add: a_star_preserves_SNa) + } + then show ":M' \ BINDINGc B (\(B)\)" by (auto simp add: BINDINGc_def) +next + assume red: "M \\<^isub>a* M'" + assume asm: "(x):M \ BINDINGn B (\\)" + { + fix c::"coname" and P::"trm" + from asm have "(:P) \ (\\) \ SNa (M{x:=.P})" by (simp add: BINDINGn_elim) + moreover + have "M{x:=.P} \\<^isub>a* M'{x:=.P}" using red by (simp add: a_star_subst1) + ultimately + have "(:P) \ (\\) \ SNa (M'{x:=.P})" by (simp add: a_star_preserves_SNa) + } + then show "(x):M' \ BINDINGn B (\\)" by (auto simp add: BINDINGn_def) +qed + +lemma CANDs_preserved: + shows ":M \ \\ \ M \\<^isub>a* M' \ :M' \ \\" + and "(x):M \ \(B)\ \ M \\<^isub>a* M' \ (x):M' \ \(B)\" +proof(nominal_induct B arbitrary: a x M M' rule: ty.induct) + case (PR X) + { case 1 + have asm: "M \\<^isub>a* M'" by fact + have ":M \ \\" by fact + then have ":M \ NEGc (PR X) (\(PR X)\)" by simp + then have ":M \ AXIOMSc (PR X) \ BINDINGc (PR X) (\(PR X)\)" by simp + moreover + { assume ":M \ AXIOMSc (PR X)" + then have ":M' \ AXIOMSc (PR X)" using asm by (simp only: AXIOMS_preserved) + } + moreover + { assume ":M \ BINDINGc (PR X) (\(PR X)\)" + then have ":M' \ BINDINGc (PR X) (\(PR X)\)" using asm by (simp add: BINDING_preserved) + } + ultimately have ":M' \ AXIOMSc (PR X) \ BINDINGc (PR X) (\(PR X)\)" by blast + then have ":M' \ NEGc (PR X) (\(PR X)\)" by simp + then show ":M' \ (\\)" using NEG_simp by blast + next + case 2 + have asm: "M \\<^isub>a* M'" by fact + have "(x):M \ \(PR X)\" by fact + then have "(x):M \ NEGn (PR X) (\\)" using NEG_simp by blast + then have "(x):M \ AXIOMSn (PR X) \ BINDINGn (PR X) (\\)" by simp + moreover + { assume "(x):M \ AXIOMSn (PR X)" + then have "(x):M' \ AXIOMSn (PR X)" using asm by (simp only: AXIOMS_preserved) + } + moreover + { assume "(x):M \ BINDINGn (PR X) (\\)" + then have "(x):M' \ BINDINGn (PR X) (\\)" using asm by (simp only: BINDING_preserved) + } + ultimately have "(x):M' \ AXIOMSn (PR X) \ BINDINGn (PR X) (\\)" by blast + then have "(x):M' \ NEGn (PR X) (\\)" by simp + then show "(x):M' \ (\(PR X)\)" using NEG_simp by blast + } +next + case (IMP A B) + have ih1: "\a M M'. \:M \ \\; M \\<^isub>a* M'\ \ :M' \ \\" by fact + have ih2: "\x M M'. \(x):M \ \(A)\; M \\<^isub>a* M'\ \ (x):M' \ \(A)\" by fact + have ih3: "\a M M'. \:M \ \\; M \\<^isub>a* M'\ \ :M' \ \\" by fact + have ih4: "\x M M'. \(x):M \ \(B)\; M \\<^isub>a* M'\ \ (x):M' \ \(B)\" by fact + { case 1 + have asm: "M \\<^isub>a* M'" by fact + have ":M \ \\" by fact + then have ":M \ NEGc (A IMP B) (\(A IMP B)\)" by simp + then have ":M \ AXIOMSc (A IMP B) \ BINDINGc (A IMP B) (\(A IMP B)\) + \ IMPRIGHT (A IMP B) (\(A)\) (\\) (\(B)\) (\\)" by simp + moreover + { assume ":M \ AXIOMSc (A IMP B)" + then have ":M' \ AXIOMSc (A IMP B)" using asm by (simp only: AXIOMS_preserved) + } + moreover + { assume ":M \ BINDINGc (A IMP B) (\(A IMP B)\)" + then have ":M' \ BINDINGc (A IMP B) (\(A IMP B)\)" using asm by (simp only: BINDING_preserved) + } + moreover + { assume ":M \ IMPRIGHT (A IMP B) (\(A)\) (\\) (\(B)\) (\\)" + then obtain x' a' N' where eq: "M = ImpR (x')..N' a" and fic: "fic (ImpR (x')..N' a) a" + and imp1: "\z P. x'\(z,P) \ (z):P \ \(B)\ \ (x'):(N'{a':=(z).P}) \ \(A)\" + and imp2: "\c Q. a'\(c,Q) \ :Q \ \\ \ :(N'{x':=.Q}) \ \\" + using IMPRIGHT_elim by blast + from eq asm obtain N'' where eq': "M' = ImpR (x')..N'' a" and red: "N' \\<^isub>a* N''" + using a_star_redu_ImpR_elim by (blast) + from imp1 have "\z P. x'\(z,P) \ (z):P \ \(B)\ \ (x'):(N''{a':=(z).P}) \ \(A)\" using red ih2 + apply(auto) + apply(drule_tac x="z" in spec) + apply(drule_tac x="P" in spec) + apply(simp) + apply(drule_tac a_star_subst2) + apply(blast) + done + moreover + from imp2 have "\c Q. a'\(c,Q) \ :Q \ \\ \ :(N''{x':=.Q}) \ \\" using red ih3 + apply(auto) + apply(drule_tac x="c" in spec) + apply(drule_tac x="Q" in spec) + apply(simp) + apply(drule_tac a_star_subst1) + apply(blast) + done + moreover + from fic have "fic M' a" using eq asm by (simp add: fic_a_star_reduce) + ultimately have ":M' \ IMPRIGHT (A IMP B) (\(A)\) (\\) (\(B)\) (\\)" using eq' by auto + } + ultimately have ":M' \ AXIOMSc (A IMP B) \ BINDINGc (A IMP B) (\(A IMP B)\) + \ IMPRIGHT (A IMP B) (\(A)\) (\\) (\(B)\) (\\)" by blast + then have ":M' \ NEGc (A IMP B) (\(A IMP B)\)" by simp + then show ":M' \ (\\)" using NEG_simp by blast + next + case 2 + have asm: "M \\<^isub>a* M'" by fact + have "(x):M \ \(A IMP B)\" by fact + then have "(x):M \ NEGn (A IMP B) (\\)" using NEG_simp by blast + then have "(x):M \ AXIOMSn (A IMP B) \ BINDINGn (A IMP B) (\\) + \ IMPLEFT (A IMP B) (\\) (\(B)\)" by simp + moreover + { assume "(x):M \ AXIOMSn (A IMP B)" + then have "(x):M' \ AXIOMSn (A IMP B)" using asm by (simp only: AXIOMS_preserved) + } + moreover + { assume "(x):M \ BINDINGn (A IMP B) (\\)" + then have "(x):M' \ BINDINGn (A IMP B) (\\)" using asm by (simp only: BINDING_preserved) + } + moreover + { assume "(x):M \ IMPLEFT (A IMP B) (\\) (\(B)\)" + then obtain a' T' y' N' where eq: "M = ImpL .T' (y').N' x" + and fin: "fin (ImpL .T' (y').N' x) x" + and imp1: ":T' \ \\" and imp2: "(y'):N' \ \(B)\" + by (erule_tac IMPLEFT_elim, blast) + from eq asm obtain T'' N'' where eq': "M' = ImpL .T'' (y').N'' x" + and red1: "T' \\<^isub>a* T''" and red2: "N' \\<^isub>a* N''" + using a_star_redu_ImpL_elim by blast + from fin have "fin M' x" using eq asm by (simp add: fin_a_star_reduce) + moreover + from imp1 red1 have ":T'' \ \\" using ih1 by simp + moreover + from imp2 red2 have "(y'):N'' \ \(B)\" using ih4 by simp + ultimately have "(x):M' \ IMPLEFT (A IMP B) (\\) (\(B)\)" using eq' by (simp, blast) + } + ultimately have "(x):M' \ AXIOMSn (A IMP B) \ BINDINGn (A IMP B) (\\) + \ IMPLEFT (A IMP B) (\\) (\(B)\)" by blast + then have "(x):M' \ NEGn (A IMP B) (\\)" by simp + then show "(x):M' \ (\(A IMP B)\)" using NEG_simp by blast + } +next + case (AND A B) + have ih1: "\a M M'. \:M \ \\; M \\<^isub>a* M'\ \ :M' \ \\" by fact + have ih2: "\x M M'. \(x):M \ \(A)\; M \\<^isub>a* M'\ \ (x):M' \ \(A)\" by fact + have ih3: "\a M M'. \:M \ \\; M \\<^isub>a* M'\ \ :M' \ \\" by fact + have ih4: "\x M M'. \(x):M \ \(B)\; M \\<^isub>a* M'\ \ (x):M' \ \(B)\" by fact + { case 1 + have asm: "M \\<^isub>a* M'" by fact + have ":M \ \\" by fact + then have ":M \ NEGc (A AND B) (\(A AND B)\)" by simp + then have ":M \ AXIOMSc (A AND B) \ BINDINGc (A AND B) (\(A AND B)\) + \ ANDRIGHT (A AND B) (\\) (\\)" by simp + moreover + { assume ":M \ AXIOMSc (A AND B)" + then have ":M' \ AXIOMSc (A AND B)" using asm by (simp only: AXIOMS_preserved) + } + moreover + { assume ":M \ BINDINGc (A AND B) (\(A AND B)\)" + then have ":M' \ BINDINGc (A AND B) (\(A AND B)\)" using asm by (simp only: BINDING_preserved) + } + moreover + { assume ":M \ ANDRIGHT (A AND B) (\\) (\\)" + then obtain a' T' b' N' where eq: "M = AndR .T' .N' a" + and fic: "fic (AndR .T' .N' a) a" + and imp1: ":T' \ \\" and imp2: ":N' \ \\" + using ANDRIGHT_elim by blast + from eq asm obtain T'' N'' where eq': "M' = AndR .T'' .N'' a" + and red1: "T' \\<^isub>a* T''" and red2: "N' \\<^isub>a* N''" + using a_star_redu_AndR_elim by blast + from fic have "fic M' a" using eq asm by (simp add: fic_a_star_reduce) + moreover + from imp1 red1 have ":T'' \ \\" using ih1 by simp + moreover + from imp2 red2 have ":N'' \ \\" using ih3 by simp + ultimately have ":M' \ ANDRIGHT (A AND B) (\\) (\\)" using eq' by (simp, blast) + } + ultimately have ":M' \ AXIOMSc (A AND B) \ BINDINGc (A AND B) (\(A AND B)\) + \ ANDRIGHT (A AND B) (\\) (\\)" by blast + then have ":M' \ NEGc (A AND B) (\(A AND B)\)" by simp + then show ":M' \ (\\)" using NEG_simp by blast + next + case 2 + have asm: "M \\<^isub>a* M'" by fact + have "(x):M \ \(A AND B)\" by fact + then have "(x):M \ NEGn (A AND B) (\\)" using NEG_simp by blast + then have "(x):M \ AXIOMSn (A AND B) \ BINDINGn (A AND B) (\\) + \ ANDLEFT1 (A AND B) (\(A)\) \ ANDLEFT2 (A AND B) (\(B)\)" by simp + moreover + { assume "(x):M \ AXIOMSn (A AND B)" + then have "(x):M' \ AXIOMSn (A AND B)" using asm by (simp only: AXIOMS_preserved) + } + moreover + { assume "(x):M \ BINDINGn (A AND B) (\\)" + then have "(x):M' \ BINDINGn (A AND B) (\\)" using asm by (simp only: BINDING_preserved) + } + moreover + { assume "(x):M \ ANDLEFT1 (A AND B) (\(A)\)" + then obtain y' N' where eq: "M = AndL1 (y').N' x" + and fin: "fin (AndL1 (y').N' x) x" and imp: "(y'):N' \ \(A)\" + by (erule_tac ANDLEFT1_elim, blast) + from eq asm obtain N'' where eq': "M' = AndL1 (y').N'' x" and red1: "N' \\<^isub>a* N''" + using a_star_redu_AndL1_elim by blast + from fin have "fin M' x" using eq asm by (simp add: fin_a_star_reduce) + moreover + from imp red1 have "(y'):N'' \ \(A)\" using ih2 by simp + ultimately have "(x):M' \ ANDLEFT1 (A AND B) (\(A)\)" using eq' by (simp, blast) + } + moreover + { assume "(x):M \ ANDLEFT2 (A AND B) (\(B)\)" + then obtain y' N' where eq: "M = AndL2 (y').N' x" + and fin: "fin (AndL2 (y').N' x) x" and imp: "(y'):N' \ \(B)\" + by (erule_tac ANDLEFT2_elim, blast) + from eq asm obtain N'' where eq': "M' = AndL2 (y').N'' x" and red1: "N' \\<^isub>a* N''" + using a_star_redu_AndL2_elim by blast + from fin have "fin M' x" using eq asm by (simp add: fin_a_star_reduce) + moreover + from imp red1 have "(y'):N'' \ \(B)\" using ih4 by simp + ultimately have "(x):M' \ ANDLEFT2 (A AND B) (\(B)\)" using eq' by (simp, blast) + } + ultimately have "(x):M' \ AXIOMSn (A AND B) \ BINDINGn (A AND B) (\\) + \ ANDLEFT1 (A AND B) (\(A)\) \ ANDLEFT2 (A AND B) (\(B)\)" by blast + then have "(x):M' \ NEGn (A AND B) (\\)" by simp + then show "(x):M' \ (\(A AND B)\)" using NEG_simp by blast + } +next + case (OR A B) + have ih1: "\a M M'. \:M \ \\; M \\<^isub>a* M'\ \ :M' \ \\" by fact + have ih2: "\x M M'. \(x):M \ \(A)\; M \\<^isub>a* M'\ \ (x):M' \ \(A)\" by fact + have ih3: "\a M M'. \:M \ \\; M \\<^isub>a* M'\ \ :M' \ \\" by fact + have ih4: "\x M M'. \(x):M \ \(B)\; M \\<^isub>a* M'\ \ (x):M' \ \(B)\" by fact + { case 1 + have asm: "M \\<^isub>a* M'" by fact + have ":M \ \\" by fact + then have ":M \ NEGc (A OR B) (\(A OR B)\)" by simp + then have ":M \ AXIOMSc (A OR B) \ BINDINGc (A OR B) (\(A OR B)\) + \ ORRIGHT1 (A OR B) (\\) \ ORRIGHT2 (A OR B) (\\)" by simp + moreover + { assume ":M \ AXIOMSc (A OR B)" + then have ":M' \ AXIOMSc (A OR B)" using asm by (simp only: AXIOMS_preserved) + } + moreover + { assume ":M \ BINDINGc (A OR B) (\(A OR B)\)" + then have ":M' \ BINDINGc (A OR B) (\(A OR B)\)" using asm by (simp only: BINDING_preserved) + } + moreover + { assume ":M \ ORRIGHT1 (A OR B) (\\)" + then obtain a' N' where eq: "M = OrR1 .N' a" + and fic: "fic (OrR1 .N' a) a" and imp1: ":N' \ \\" + using ORRIGHT1_elim by blast + from eq asm obtain N'' where eq': "M' = OrR1 .N'' a" and red1: "N' \\<^isub>a* N''" + using a_star_redu_OrR1_elim by blast + from fic have "fic M' a" using eq asm by (simp add: fic_a_star_reduce) + moreover + from imp1 red1 have ":N'' \ \\" using ih1 by simp + ultimately have ":M' \ ORRIGHT1 (A OR B) (\\)" using eq' by (simp, blast) + } + moreover + { assume ":M \ ORRIGHT2 (A OR B) (\\)" + then obtain a' N' where eq: "M = OrR2 .N' a" + and fic: "fic (OrR2 .N' a) a" and imp1: ":N' \ \\" + using ORRIGHT2_elim by blast + from eq asm obtain N'' where eq': "M' = OrR2 .N'' a" and red1: "N' \\<^isub>a* N''" + using a_star_redu_OrR2_elim by blast + from fic have "fic M' a" using eq asm by (simp add: fic_a_star_reduce) + moreover + from imp1 red1 have ":N'' \ \\" using ih3 by simp + ultimately have ":M' \ ORRIGHT2 (A OR B) (\\)" using eq' by (simp, blast) + } + ultimately have ":M' \ AXIOMSc (A OR B) \ BINDINGc (A OR B) (\(A OR B)\) + \ ORRIGHT1 (A OR B) (\\) \ ORRIGHT2 (A OR B) (\\)" by blast + then have ":M' \ NEGc (A OR B) (\(A OR B)\)" by simp + then show ":M' \ (\\)" using NEG_simp by blast + next + case 2 + have asm: "M \\<^isub>a* M'" by fact + have "(x):M \ \(A OR B)\" by fact + then have "(x):M \ NEGn (A OR B) (\\)" using NEG_simp by blast + then have "(x):M \ AXIOMSn (A OR B) \ BINDINGn (A OR B) (\\) + \ ORLEFT (A OR B) (\(A)\) (\(B)\)" by simp + moreover + { assume "(x):M \ AXIOMSn (A OR B)" + then have "(x):M' \ AXIOMSn (A OR B)" using asm by (simp only: AXIOMS_preserved) + } + moreover + { assume "(x):M \ BINDINGn (A OR B) (\\)" + then have "(x):M' \ BINDINGn (A OR B) (\\)" using asm by (simp only: BINDING_preserved) + } + moreover + { assume "(x):M \ ORLEFT (A OR B) (\(A)\) (\(B)\)" + then obtain y' T' z' N' where eq: "M = OrL (y').T' (z').N' x" + and fin: "fin (OrL (y').T' (z').N' x) x" + and imp1: "(y'):T' \ \(A)\" and imp2: "(z'):N' \ \(B)\" + by (erule_tac ORLEFT_elim, blast) + from eq asm obtain T'' N'' where eq': "M' = OrL (y').T'' (z').N'' x" + and red1: "T' \\<^isub>a* T''" and red2: "N' \\<^isub>a* N''" + using a_star_redu_OrL_elim by blast + from fin have "fin M' x" using eq asm by (simp add: fin_a_star_reduce) + moreover + from imp1 red1 have "(y'):T'' \ \(A)\" using ih2 by simp + moreover + from imp2 red2 have "(z'):N'' \ \(B)\" using ih4 by simp + ultimately have "(x):M' \ ORLEFT (A OR B) (\(A)\) (\(B)\)" using eq' by (simp, blast) + } + ultimately have "(x):M' \ AXIOMSn (A OR B) \ BINDINGn (A OR B) (\\) + \ ORLEFT (A OR B) (\(A)\) (\(B)\)" by blast + then have "(x):M' \ NEGn (A OR B) (\\)" by simp + then show "(x):M' \ (\(A OR B)\)" using NEG_simp by blast + } +next + case (NOT A) + have ih1: "\a M M'. \:M \ \\; M \\<^isub>a* M'\ \ :M' \ \\" by fact + have ih2: "\x M M'. \(x):M \ \(A)\; M \\<^isub>a* M'\ \ (x):M' \ \(A)\" by fact + { case 1 + have asm: "M \\<^isub>a* M'" by fact + have ":M \ \\" by fact + then have ":M \ NEGc (NOT A) (\(NOT A)\)" by simp + then have ":M \ AXIOMSc (NOT A) \ BINDINGc (NOT A) (\(NOT A)\) + \ NOTRIGHT (NOT A) (\(A)\)" by simp + moreover + { assume ":M \ AXIOMSc (NOT A)" + then have ":M' \ AXIOMSc (NOT A)" using asm by (simp only: AXIOMS_preserved) + } + moreover + { assume ":M \ BINDINGc (NOT A) (\(NOT A)\)" + then have ":M' \ BINDINGc (NOT A) (\(NOT A)\)" using asm by (simp only: BINDING_preserved) + } + moreover + { assume ":M \ NOTRIGHT (NOT A) (\(A)\)" + then obtain y' N' where eq: "M = NotR (y').N' a" + and fic: "fic (NotR (y').N' a) a" and imp: "(y'):N' \ \(A)\" + using NOTRIGHT_elim by blast + from eq asm obtain N'' where eq': "M' = NotR (y').N'' a" and red: "N' \\<^isub>a* N''" + using a_star_redu_NotR_elim by blast + from fic have "fic M' a" using eq asm by (simp add: fic_a_star_reduce) + moreover + from imp red have "(y'):N'' \ \(A)\" using ih2 by simp + ultimately have ":M' \ NOTRIGHT (NOT A) (\(A)\)" using eq' by (simp, blast) + } + ultimately have ":M' \ AXIOMSc (NOT A) \ BINDINGc (NOT A) (\(NOT A)\) + \ NOTRIGHT (NOT A) (\(A)\)" by blast + then have ":M' \ NEGc (NOT A) (\(NOT A)\)" by simp + then show ":M' \ (\\)" using NEG_simp by blast + next + case 2 + have asm: "M \\<^isub>a* M'" by fact + have "(x):M \ \(NOT A)\" by fact + then have "(x):M \ NEGn (NOT A) (\\)" using NEG_simp by blast + then have "(x):M \ AXIOMSn (NOT A) \ BINDINGn (NOT A) (\\) + \ NOTLEFT (NOT A) (\\)" by simp + moreover + { assume "(x):M \ AXIOMSn (NOT A)" + then have "(x):M' \ AXIOMSn (NOT A)" using asm by (simp only: AXIOMS_preserved) + } + moreover + { assume "(x):M \ BINDINGn (NOT A) (\\)" + then have "(x):M' \ BINDINGn (NOT A) (\\)" using asm by (simp only: BINDING_preserved) + } + moreover + { assume "(x):M \ NOTLEFT (NOT A) (\\)" + then obtain a' N' where eq: "M = NotL .N' x" + and fin: "fin (NotL .N' x) x" and imp: ":N' \ \\" + by (erule_tac NOTLEFT_elim, blast) + from eq asm obtain N'' where eq': "M' = NotL .N'' x" and red1: "N' \\<^isub>a* N''" + using a_star_redu_NotL_elim by blast + from fin have "fin M' x" using eq asm by (simp add: fin_a_star_reduce) + moreover + from imp red1 have ":N'' \ \\" using ih1 by simp + ultimately have "(x):M' \ NOTLEFT (NOT A) (\\)" using eq' by (simp, blast) + } + ultimately have "(x):M' \ AXIOMSn (NOT A) \ BINDINGn (NOT A) (\\) + \ NOTLEFT (NOT A) (\\)" by blast + then have "(x):M' \ NEGn (NOT A) (\\)" by simp + then show "(x):M' \ (\(NOT A)\)" using NEG_simp by blast + } +qed + +lemma CANDs_preserved_single: + shows ":M \ \\ \ M \\<^isub>a M' \ :M' \ \\" + and "(x):M \ \(B)\ \ M \\<^isub>a M' \ (x):M' \ \(B)\" +by (auto simp add: a_starI CANDs_preserved) + +lemma fic_CANDS: + assumes a: "\fic M a" + and b: ":M \ \\" + shows ":M \ AXIOMSc B \ :M \ BINDINGc B (\(B)\)" +using a b +apply(nominal_induct B rule: ty.induct) +apply(simp) +apply(simp) +apply(erule disjE) +apply(simp) +apply(erule disjE) +apply(simp) +apply(auto simp add: ctrm.inject)[1] +apply(simp add: alpha) +apply(erule disjE) +apply(simp) +apply(auto simp add: calc_atm)[1] +apply(drule_tac pi="[(a,aa)]" in fic.eqvt(2)) +apply(simp add: calc_atm) +apply(simp) +apply(erule disjE) +apply(simp) +apply(erule disjE) +apply(simp) +apply(auto simp add: ctrm.inject)[1] +apply(simp add: alpha) +apply(erule disjE) +apply(simp) +apply(erule conjE)+ +apply(simp) +apply(drule_tac pi="[(a,c)]" in fic.eqvt(2)) +apply(simp add: calc_atm) +apply(simp) +apply(erule disjE) +apply(simp) +apply(erule disjE) +apply(simp) +apply(auto simp add: ctrm.inject)[1] +apply(simp add: alpha) +apply(erule disjE) +apply(simp) +apply(erule conjE)+ +apply(simp) +apply(drule_tac pi="[(a,b)]" in fic.eqvt(2)) +apply(simp add: calc_atm) +apply(simp add: alpha) +apply(erule disjE) +apply(simp) +apply(erule conjE)+ +apply(simp) +apply(drule_tac pi="[(a,b)]" in fic.eqvt(2)) +apply(simp add: calc_atm) +apply(simp) +apply(erule disjE) +apply(simp) +apply(erule disjE) +apply(simp) +apply(auto simp add: ctrm.inject)[1] +apply(simp add: alpha) +apply(erule disjE) +apply(simp) +apply(erule conjE)+ +apply(simp) +apply(drule_tac pi="[(a,b)]" in fic.eqvt(2)) +apply(simp add: calc_atm) +done + +lemma fin_CANDS_aux: + assumes a: "\fin M x" + and b: "(x):M \ (NEGn B (\\))" + shows "(x):M \ AXIOMSn B \ (x):M \ BINDINGn B (\\)" +using a b +apply(nominal_induct B rule: ty.induct) +apply(simp) +apply(simp) +apply(erule disjE) +apply(simp) +apply(erule disjE) +apply(simp) +apply(auto simp add: ntrm.inject)[1] +apply(simp add: alpha) +apply(erule disjE) +apply(simp) +apply(auto simp add: calc_atm)[1] +apply(drule_tac pi="[(x,xa)]" in fin.eqvt(1)) +apply(simp add: calc_atm) +apply(simp) +apply(erule disjE) +apply(simp) +apply(erule disjE) +apply(simp) +apply(auto simp add: ntrm.inject)[1] +apply(simp add: alpha) +apply(erule disjE) +apply(simp) +apply(erule conjE)+ +apply(simp) +apply(drule_tac pi="[(x,y)]" in fin.eqvt(1)) +apply(simp add: calc_atm) +apply(simp add: alpha) +apply(erule disjE) +apply(simp) +apply(erule conjE)+ +apply(simp) +apply(drule_tac pi="[(x,y)]" in fin.eqvt(1)) +apply(simp add: calc_atm) +apply(simp) +apply(erule disjE) +apply(simp) +apply(erule disjE) +apply(simp) +apply(auto simp add: ntrm.inject)[1] +apply(simp add: alpha) +apply(erule disjE) +apply(simp) +apply(erule conjE)+ +apply(simp) +apply(drule_tac pi="[(x,z)]" in fin.eqvt(1)) +apply(simp add: calc_atm) +apply(simp) +apply(erule disjE) +apply(simp) +apply(erule disjE) +apply(simp) +apply(auto simp add: ntrm.inject)[1] +apply(simp add: alpha) +apply(erule disjE) +apply(simp) +apply(erule conjE)+ +apply(simp) +apply(drule_tac pi="[(x,y)]" in fin.eqvt(1)) +apply(simp add: calc_atm) +done + +lemma fin_CANDS: + assumes a: "\fin M x" + and b: "(x):M \ (\(B)\)" + shows "(x):M \ AXIOMSn B \ (x):M \ BINDINGn B (\\)" +apply(rule fin_CANDS_aux) +apply(rule a) +apply(rule NEG_elim) +apply(rule b) +done + +lemma BINDING_implies_CAND: + shows ":M \ BINDINGc B (\(B)\) \ :M \ (\\)" + and "(x):N \ BINDINGn B (\\) \ (x):N \ (\(B)\)" +apply - +apply(nominal_induct B rule: ty.induct) +apply(auto) +apply(rule NEG_intro) +apply(nominal_induct B rule: ty.induct) +apply(auto) +done + +text {* 3rd Main Lemma *} + +lemma Cut_a_redu_elim: + assumes a: "Cut .M (x).N \\<^isub>a R" + shows "(\M'. R = Cut .M' (x).N \ M \\<^isub>a M') \ + (\N'. R = Cut .M (x).N' \ N \\<^isub>a N') \ + (Cut .M (x).N \\<^isub>c R) \ + (Cut .M (x).N \\<^isub>l R)" +using a +apply(erule_tac a_redu.cases) +apply(simp_all) +apply(simp_all add: trm.inject) +apply(rule disjI1) +apply(auto simp add: alpha)[1] +apply(rule_tac x="[(a,aa)]\M'" in exI) +apply(perm_simp add: fresh_left calc_atm a_redu.eqvt fresh_a_redu) +apply(rule_tac x="[(a,aa)]\M'" in exI) +apply(perm_simp add: fresh_left calc_atm a_redu.eqvt fresh_a_redu) +apply(rule disjI2) +apply(rule disjI1) +apply(auto simp add: alpha)[1] +apply(rule_tac x="[(x,xa)]\N'" in exI) +apply(perm_simp add: fresh_left calc_atm a_redu.eqvt fresh_a_redu) +apply(rule_tac x="[(x,xa)]\N'" in exI) +apply(perm_simp add: fresh_left calc_atm a_redu.eqvt fresh_a_redu) +done + +lemma Cut_c_redu_elim: + assumes a: "Cut .M (x).N \\<^isub>c R" + shows "(R = M{a:=(x).N} \ \fic M a) \ + (R = N{x:=.M} \ \fin N x)" +using a +apply(erule_tac c_redu.cases) +apply(simp_all) +apply(simp_all add: trm.inject) +apply(rule disjI1) +apply(auto simp add: alpha)[1] +apply(simp add: subst_rename fresh_atm) +apply(simp add: subst_rename fresh_atm) +apply(drule_tac pi="[(a,aa)]" in fic.eqvt(2)) +apply(perm_simp) +apply(simp add: subst_rename fresh_atm fresh_prod) +apply(drule_tac pi="[(a,aa)]" in fic.eqvt(2)) +apply(perm_simp) +apply(rule disjI2) +apply(auto simp add: alpha)[1] +apply(simp add: subst_rename fresh_atm) +apply(drule_tac pi="[(x,xa)]" in fin.eqvt(1)) +apply(perm_simp) +apply(simp add: subst_rename fresh_atm fresh_prod) +apply(simp add: subst_rename fresh_atm fresh_prod) +apply(drule_tac pi="[(x,xa)]" in fin.eqvt(1)) +apply(perm_simp) +done + +lemma not_fic_crename_aux: + assumes a: "fic M c" "c\(a,b)" + shows "fic (M[a\c>b]) c" +using a +apply(nominal_induct M avoiding: c a b rule: trm.induct) +apply(auto dest!: fic_elims intro!: fic.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh) +done + +lemma not_fic_crename: + assumes a: "\(fic (M[a\c>b]) c)" "c\(a,b)" + shows "\(fic M c)" +using a +apply(auto dest: not_fic_crename_aux) +done + +lemma not_fin_crename_aux: + assumes a: "fin M y" + shows "fin (M[a\c>b]) y" +using a +apply(nominal_induct M avoiding: a b rule: trm.induct) +apply(auto dest!: fin_elims intro!: fin.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh) +done + +lemma not_fin_crename: + assumes a: "\(fin (M[a\c>b]) y)" + shows "\(fin M y)" +using a +apply(auto dest: not_fin_crename_aux) +done + +lemma crename_fresh_interesting1: + fixes c::"coname" + assumes a: "c\(M[a\c>b])" "c\(a,b)" + shows "c\M" +using a +apply(nominal_induct M avoiding: c a b rule: trm.induct) +apply(auto split: if_splits simp add: abs_fresh) +done + +lemma crename_fresh_interesting2: + fixes x::"name" + assumes a: "x\(M[a\c>b])" + shows "x\M" +using a +apply(nominal_induct M avoiding: x a b rule: trm.induct) +apply(auto split: if_splits simp add: abs_fresh abs_supp fin_supp fresh_atm) +done + + +lemma fic_crename: + assumes a: "fic (M[a\c>b]) c" "c\(a,b)" + shows "fic M c" +using a +apply(nominal_induct M avoiding: c a b rule: trm.induct) +apply(auto dest!: fic_elims intro!: fic.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh + split: if_splits) +apply(auto dest: crename_fresh_interesting1 simp add: fresh_prod fresh_atm) +done + +lemma fin_crename: + assumes a: "fin (M[a\c>b]) x" + shows "fin M x" +using a +apply(nominal_induct M avoiding: x a b rule: trm.induct) +apply(auto dest!: fin_elims intro!: fin.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh + split: if_splits) +apply(auto dest: crename_fresh_interesting2 simp add: fresh_prod fresh_atm) +done + +lemma crename_Cut: + assumes a: "R[a\c>b] = Cut .M (x).N" "c\(a,b,N,R)" "x\(M,R)" + shows "\M' N'. R = Cut .M' (x).N' \ M'[a\c>b] = M \ N'[a\c>b] = N \ c\N' \ x\M'" +using a +apply(nominal_induct R avoiding: a b c x M N rule: trm.induct) +apply(auto split: if_splits) +apply(simp add: trm.inject) +apply(auto simp add: alpha) +apply(rule_tac x="[(name,x)]\trm2" in exI) +apply(perm_simp) +apply(simp add: abs_fresh fresh_left calc_atm fresh_prod) +apply(drule sym) +apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst]) +apply(simp add: eqvts calc_atm) +apply(rule_tac x="[(coname,c)]\trm1" in exI) +apply(perm_simp) +apply(simp add: abs_fresh fresh_left calc_atm fresh_prod) +apply(drule sym) +apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst]) +apply(simp add: eqvts calc_atm) +apply(auto simp add: fresh_atm)[1] +apply(rule_tac x="[(coname,c)]\trm1" in exI) +apply(perm_simp) +apply(simp add: abs_fresh fresh_left calc_atm fresh_prod) +apply(rule_tac x="[(name,x)]\trm2" in exI) +apply(perm_simp) +apply(simp add: abs_fresh fresh_left calc_atm fresh_prod) +apply(drule sym) +apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst]) +apply(simp add: eqvts calc_atm) +apply(auto simp add: fresh_atm)[1] +apply(drule sym) +apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst]) +apply(simp add: eqvts calc_atm) +apply(drule sym) +apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst]) +apply(simp add: eqvts calc_atm) +done + +lemma crename_NotR: + assumes a: "R[a\c>b] = NotR (x).N c" "x\R" "c\(a,b)" + shows "\N'. (R = NotR (x).N' c) \ N'[a\c>b] = N" +using a +apply(nominal_induct R avoiding: a b c x N rule: trm.induct) +apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject) +apply(rule_tac x="[(name,x)]\trm" in exI) +apply(perm_simp) +apply(simp add: abs_fresh fresh_left calc_atm fresh_prod) +apply(drule sym) +apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst]) +apply(simp add: eqvts calc_atm) +done + +lemma crename_NotR': + assumes a: "R[a\c>b] = NotR (x).N c" "x\R" "c\a" + shows "(\N'. (R = NotR (x).N' c) \ N'[a\c>b] = N) \ (\N'. (R = NotR (x).N' a) \ b=c \ N'[a\c>b] = N)" +using a +apply(nominal_induct R avoiding: a b c x N rule: trm.induct) +apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_fresh alpha trm.inject) +apply(rule_tac x="[(name,x)]\trm" in exI) +apply(perm_simp) +apply(simp add: abs_fresh fresh_left calc_atm fresh_prod) +apply(drule sym) +apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst]) +apply(simp add: eqvts calc_atm) +apply(rule_tac x="[(name,x)]\trm" in exI) +apply(perm_simp) +apply(simp add: abs_fresh fresh_left calc_atm fresh_prod) +apply(drule sym) +apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst]) +apply(simp add: eqvts calc_atm) +done + +lemma crename_NotR_aux: + assumes a: "R[a\c>b] = NotR (x).N c" + shows "(a=c \ a=b) \ (a\c)" +using a +apply(nominal_induct R avoiding: a b c x N rule: trm.induct) +apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject) +done + +lemma crename_NotL: + assumes a: "R[a\c>b] = NotL .N y" "c\(R,a,b)" + shows "\N'. (R = NotL .N' y) \ N'[a\c>b] = N" +using a +apply(nominal_induct R avoiding: a b c y N rule: trm.induct) +apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject) +apply(rule_tac x="[(coname,c)]\trm" in exI) +apply(perm_simp) +apply(simp add: abs_fresh fresh_left calc_atm fresh_prod) +apply(drule sym) +apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst]) +apply(simp add: eqvts calc_atm) +done + +lemma crename_AndL1: + assumes a: "R[a\c>b] = AndL1 (x).N y" "x\R" + shows "\N'. (R = AndL1 (x).N' y) \ N'[a\c>b] = N" +using a +apply(nominal_induct R avoiding: a b x y N rule: trm.induct) +apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject) +apply(rule_tac x="[(name1,x)]\trm" in exI) +apply(perm_simp) +apply(simp add: abs_fresh fresh_left calc_atm fresh_prod) +apply(drule sym) +apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst]) +apply(simp add: eqvts calc_atm) +done + +lemma crename_AndL2: + assumes a: "R[a\c>b] = AndL2 (x).N y" "x\R" + shows "\N'. (R = AndL2 (x).N' y) \ N'[a\c>b] = N" +using a +apply(nominal_induct R avoiding: a b x y N rule: trm.induct) +apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject) +apply(rule_tac x="[(name1,x)]\trm" in exI) +apply(perm_simp) +apply(simp add: abs_fresh fresh_left calc_atm fresh_prod) +apply(drule sym) +apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst]) +apply(simp add: eqvts calc_atm) +done + +lemma crename_AndR_aux: + assumes a: "R[a\c>b] = AndR .M .N e" + shows "(a=e \ a=b) \ (a\e)" +using a +apply(nominal_induct R avoiding: a b c d e M N rule: trm.induct) +apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject) +done + +lemma crename_AndR: + assumes a: "R[a\c>b] = AndR .M .N e" "c\(a,b,d,e,N,R)" "d\(a,b,c,e,M,R)" "e\(a,b)" + shows "\M' N'. R = AndR .M' .N' e \ M'[a\c>b] = M \ N'[a\c>b] = N \ c\N' \ d\M'" +using a +apply(nominal_induct R avoiding: a b c d e M N rule: trm.induct) +apply(auto split: if_splits simp add: trm.inject alpha) +apply(simp add: fresh_atm fresh_prod) +apply(rule_tac x="[(coname2,d)]\trm2" in exI) +apply(perm_simp) +apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(rule_tac x="[(coname1,c)]\trm1" in exI) +apply(perm_simp) +apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(rule_tac x="[(coname1,c)]\trm1" in exI) +apply(perm_simp) +apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(rule_tac x="[(coname2,d)]\trm2" in exI) +apply(perm_simp) +apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(rule_tac x="[(coname1,c)]\trm1" in exI) +apply(perm_simp) +apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(rule_tac x="[(coname1,c)]\trm1" in exI) +apply(perm_simp) +apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(rule_tac x="[(coname2,d)]\trm2" in exI) +apply(perm_simp) +apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(drule sym) +apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst]) +apply(simp add: eqvts calc_atm) +apply(drule_tac s="trm2[a\c>b]" in sym) +apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst]) +apply(simp add: eqvts calc_atm) +done + +lemma crename_AndR': + assumes a: "R[a\c>b] = AndR .M .N e" "c\(a,b,d,e,N,R)" "d\(a,b,c,e,M,R)" "e\a" + shows "(\M' N'. R = AndR .M' .N' e \ M'[a\c>b] = M \ N'[a\c>b] = N \ c\N' \ d\M') \ + (\M' N'. R = AndR .M' .N' a \ b=e \ M'[a\c>b] = M \ N'[a\c>b] = N \ c\N' \ d\M')" +using a +apply(nominal_induct R avoiding: a b c d e M N rule: trm.induct) +apply(auto split: if_splits simp add: trm.inject alpha)[1] +apply(auto split: if_splits simp add: trm.inject alpha)[1] +apply(auto split: if_splits simp add: trm.inject alpha)[1] +apply(auto split: if_splits simp add: trm.inject alpha)[1] +apply(simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm trm.inject alpha)[1] +apply(case_tac "coname3=a") +apply(simp) +apply(rule_tac x="[(coname1,c)]\trm1" in exI) +apply(perm_simp) +apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(rule_tac x="[(coname2,d)]\trm2" in exI) +apply(perm_simp) +apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm trm.inject alpha split: if_splits)[1] +apply(drule sym) +apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst]) +apply(simp add: eqvts calc_atm) +apply(drule_tac s="trm2[a\c>e]" in sym) +apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst]) +apply(simp add: eqvts calc_atm) +apply(simp) +apply(rule_tac x="[(coname1,c)]\trm1" in exI) +apply(perm_simp) +apply(simp add: abs_fresh fresh_left calc_atm fresh_prod) +apply(rule_tac x="[(coname2,d)]\trm2" in exI) +apply(perm_simp) +apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm trm.inject alpha split: if_splits)[1] +apply(drule sym) +apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst]) +apply(simp add: eqvts calc_atm) +apply(drule_tac s="trm2[a\c>b]" in sym) +apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst]) +apply(simp add: eqvts calc_atm) +apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject) +done + +lemma crename_OrR1_aux: + assumes a: "R[a\c>b] = OrR1 .M e" + shows "(a=e \ a=b) \ (a\e)" +using a +apply(nominal_induct R avoiding: a b c e M rule: trm.induct) +apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject) +done + +lemma crename_OrR1: + assumes a: "R[a\c>b] = OrR1 .N d" "c\(R,a,b)" "d\(a,b)" + shows "\N'. (R = OrR1 .N' d) \ N'[a\c>b] = N" +using a +apply(nominal_induct R avoiding: a b c d N rule: trm.induct) +apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject) +apply(rule_tac x="[(coname1,c)]\trm" in exI) +apply(perm_simp) +apply(simp add: abs_fresh fresh_left calc_atm fresh_prod) +apply(drule sym) +apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst]) +apply(simp add: eqvts calc_atm) +done + +lemma crename_OrR1': + assumes a: "R[a\c>b] = OrR1 .N d" "c\(R,a,b)" "d\a" + shows "(\N'. (R = OrR1 .N' d) \ N'[a\c>b] = N) \ + (\N'. (R = OrR1 .N' a) \ b=d \ N'[a\c>b] = N)" +using a +apply(nominal_induct R avoiding: a b c d N rule: trm.induct) +apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject) +apply(rule_tac x="[(coname1,c)]\trm" in exI) +apply(perm_simp) +apply(simp add: abs_fresh fresh_left calc_atm fresh_prod) +apply(drule sym) +apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst]) +apply(simp add: eqvts calc_atm) +apply(rule_tac x="[(coname1,c)]\trm" in exI) +apply(perm_simp) +apply(simp add: abs_fresh fresh_left calc_atm fresh_prod) +apply(drule sym) +apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst]) +apply(simp add: eqvts calc_atm) +done + +lemma crename_OrR2_aux: + assumes a: "R[a\c>b] = OrR2 .M e" + shows "(a=e \ a=b) \ (a\e)" +using a +apply(nominal_induct R avoiding: a b c e M rule: trm.induct) +apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject) +done + +lemma crename_OrR2: + assumes a: "R[a\c>b] = OrR2 .N d" "c\(R,a,b)" "d\(a,b)" + shows "\N'. (R = OrR2 .N' d) \ N'[a\c>b] = N" +using a +apply(nominal_induct R avoiding: a b c d N rule: trm.induct) +apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject) +apply(rule_tac x="[(coname1,c)]\trm" in exI) +apply(perm_simp) +apply(simp add: abs_fresh fresh_left calc_atm fresh_prod) +apply(drule sym) +apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst]) +apply(simp add: eqvts calc_atm) +done + +lemma crename_OrR2': + assumes a: "R[a\c>b] = OrR2 .N d" "c\(R,a,b)" "d\a" + shows "(\N'. (R = OrR2 .N' d) \ N'[a\c>b] = N) \ + (\N'. (R = OrR2 .N' a) \ b=d \ N'[a\c>b] = N)" +using a +apply(nominal_induct R avoiding: a b c d N rule: trm.induct) +apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject) +apply(rule_tac x="[(coname1,c)]\trm" in exI) +apply(perm_simp) +apply(simp add: abs_fresh fresh_left calc_atm fresh_prod) +apply(drule sym) +apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst]) +apply(simp add: eqvts calc_atm) +apply(rule_tac x="[(coname1,c)]\trm" in exI) +apply(perm_simp) +apply(simp add: abs_fresh fresh_left calc_atm fresh_prod) +apply(drule sym) +apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst]) +apply(simp add: eqvts calc_atm) +done + +lemma crename_OrL: + assumes a: "R[a\c>b] = OrL (x).M (y).N z" "x\(y,z,N,R)" "y\(x,z,M,R)" + shows "\M' N'. R = OrL (x).M' (y).N' z \ M'[a\c>b] = M \ N'[a\c>b] = N \ x\N' \ y\M'" +using a +apply(nominal_induct R avoiding: a b x y z M N rule: trm.induct) +apply(auto split: if_splits simp add: trm.inject alpha) +apply(rule_tac x="[(name2,y)]\trm2" in exI) +apply(perm_simp) +apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(rule_tac x="[(name1,x)]\trm1" in exI) +apply(perm_simp) +apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(rule_tac x="[(name1,x)]\trm1" in exI) +apply(perm_simp) +apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(rule_tac x="[(name2,y)]\trm2" in exI) +apply(perm_simp) +apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(drule sym) +apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst]) +apply(simp add: eqvts calc_atm) +apply(drule_tac s="trm2[a\c>b]" in sym) +apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst]) +apply(simp add: eqvts calc_atm) +done + +lemma crename_ImpL: + assumes a: "R[a\c>b] = ImpL .M (y).N z" "c\(a,b,N,R)" "y\(z,M,R)" + shows "\M' N'. R = ImpL .M' (y).N' z \ M'[a\c>b] = M \ N'[a\c>b] = N \ c\N' \ y\M'" +using a +apply(nominal_induct R avoiding: a b c y z M N rule: trm.induct) +apply(auto split: if_splits simp add: trm.inject alpha) +apply(rule_tac x="[(name1,y)]\trm2" in exI) +apply(perm_simp) +apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(rule_tac x="[(coname,c)]\trm1" in exI) +apply(perm_simp) +apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(rule_tac x="[(coname,c)]\trm1" in exI) +apply(perm_simp) +apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(rule_tac x="[(name1,y)]\trm2" in exI) +apply(perm_simp) +apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(drule sym) +apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst]) +apply(simp add: eqvts calc_atm) +apply(drule_tac s="trm2[a\c>b]" in sym) +apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst]) +apply(simp add: eqvts calc_atm) +done + +lemma crename_ImpR_aux: + assumes a: "R[a\c>b] = ImpR (x)..M e" + shows "(a=e \ a=b) \ (a\e)" +using a +apply(nominal_induct R avoiding: x a b c e M rule: trm.induct) +apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject) +done + +lemma crename_ImpR: + assumes a: "R[a\c>b] = ImpR (x)..N d" "c\(R,a,b)" "d\(a,b)" "x\R" + shows "\N'. (R = ImpR (x)..N' d) \ N'[a\c>b] = N" +using a +apply(nominal_induct R avoiding: a b x c d N rule: trm.induct) +apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_perm alpha abs_fresh trm.inject) +apply(rule_tac x="[(name,x)]\trm" in exI) +apply(perm_simp) +apply(simp add: abs_fresh fresh_left calc_atm fresh_prod) +apply(simp add: calc_atm) +apply(rule_tac x="[(name,x)]\[(coname1, c)]\trm" in exI) +apply(perm_simp) +apply(simp add: abs_supp fin_supp abs_fresh fresh_left calc_atm fresh_prod) +apply(drule sym) +apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst]) +apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst]) +apply(simp add: eqvts calc_atm) +done + +lemma crename_ImpR': + assumes a: "R[a\c>b] = ImpR (x)..N d" "c\(R,a,b)" "x\R" "d\a" + shows "(\N'. (R = ImpR (x)..N' d) \ N'[a\c>b] = N) \ + (\N'. (R = ImpR (x)..N' a) \ b=d \ N'[a\c>b] = N)" +using a +apply(nominal_induct R avoiding: x a b c d N rule: trm.induct) +apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject abs_perm calc_atm) +apply(rule_tac x="[(name,x)]\[(coname1,c)]\trm" in exI) +apply(perm_simp) +apply(simp add: abs_fresh fresh_left calc_atm fresh_prod abs_supp fin_supp) +apply(drule sym) +apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst]) +apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst]) +apply(simp add: eqvts calc_atm) +apply(rule_tac x="[(name,x)]\[(coname1,c)]\trm" in exI) +apply(perm_simp) +apply(simp add: abs_fresh fresh_left calc_atm fresh_prod abs_supp fin_supp) +apply(drule sym) +apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst]) +apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst]) +apply(simp add: eqvts calc_atm) +done + +lemma crename_ax2: + assumes a: "N[a\c>b] = Ax x c" + shows "\d. N = Ax x d" +using a +apply(nominal_induct N avoiding: a b rule: trm.induct) +apply(auto split: if_splits) +apply(simp add: trm.inject) +done + +lemma crename_interesting1: + assumes a: "distinct [a,b,c]" + shows "M[a\c>c][c\c>b] = M[c\c>b][a\c>b]" +using a +apply(nominal_induct M avoiding: a c b rule: trm.induct) +apply(auto simp add: rename_fresh simp add: trm.inject alpha) +apply(blast) +apply(rotate_tac 12) +apply(drule_tac x="a" in meta_spec) +apply(rotate_tac 15) +apply(drule_tac x="c" in meta_spec) +apply(rotate_tac 15) +apply(drule_tac x="b" in meta_spec) +apply(blast) +apply(blast) +apply(blast) +done + +lemma crename_interesting2: + assumes a: "a\c" "a\d" "a\b" "c\d" "b\c" + shows "M[a\c>b][c\c>d] = M[c\c>d][a\c>b]" +using a +apply(nominal_induct M avoiding: a c b d rule: trm.induct) +apply(auto simp add: rename_fresh simp add: trm.inject alpha) +done + +lemma crename_interesting3: + shows "M[a\c>c][x\n>y] = M[x\n>y][a\c>c]" +apply(nominal_induct M avoiding: a c x y rule: trm.induct) +apply(auto simp add: rename_fresh simp add: trm.inject alpha) +done + +lemma crename_credu: + assumes a: "(M[a\c>b]) \\<^isub>c M'" + shows "\M0. M0[a\c>b]=M' \ M \\<^isub>c M0" +using a +apply(nominal_induct M\"M[a\c>b]" M' avoiding: M a b rule: c_redu.strong_induct) +apply(drule sym) +apply(drule crename_Cut) +apply(simp) +apply(simp) +apply(auto) +apply(rule_tac x="M'{a:=(x).N'}" in exI) +apply(rule conjI) +apply(simp add: fresh_atm abs_fresh subst_comm fresh_prod) +apply(rule c_redu.intros) +apply(auto dest: not_fic_crename)[1] +apply(simp add: abs_fresh) +apply(simp add: abs_fresh) +apply(drule sym) +apply(drule crename_Cut) +apply(simp) +apply(simp) +apply(auto) +apply(rule_tac x="N'{x:=.M'}" in exI) +apply(rule conjI) +apply(simp add: fresh_atm abs_fresh subst_comm fresh_prod) +apply(rule c_redu.intros) +apply(auto dest: not_fin_crename)[1] +apply(simp add: abs_fresh) +apply(simp add: abs_fresh) +done + +lemma crename_lredu: + assumes a: "(M[a\c>b]) \\<^isub>l M'" + shows "\M0. M0[a\c>b]=M' \ M \\<^isub>l M0" +using a +apply(nominal_induct M\"M[a\c>b]" M' avoiding: M a b rule: l_redu.strong_induct) +apply(drule sym) +apply(drule crename_Cut) +apply(simp add: fresh_prod fresh_atm) +apply(simp) +apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1] +apply(case_tac "aa=ba") +apply(simp add: crename_id) +apply(rule l_redu.intros) +apply(simp) +apply(simp add: fresh_atm) +apply(assumption) +apply(frule crename_ax2) +apply(auto)[1] +apply(case_tac "d=aa") +apply(simp add: trm.inject) +apply(rule_tac x="M'[a\c>aa]" in exI) +apply(rule conjI) +apply(rule crename_interesting1) +apply(simp) +apply(rule l_redu.intros) +apply(simp) +apply(simp add: fresh_atm) +apply(auto dest: fic_crename simp add: fresh_prod fresh_atm)[1] +apply(simp add: trm.inject) +apply(rule_tac x="M'[a\c>b]" in exI) +apply(rule conjI) +apply(rule crename_interesting2) +apply(simp) +apply(simp) +apply(simp) +apply(simp) +apply(simp) +apply(rule l_redu.intros) +apply(simp) +apply(simp add: fresh_atm) +apply(auto dest: fic_crename simp add: fresh_prod fresh_atm)[1] +apply(drule sym) +apply(drule crename_Cut) +apply(simp add: fresh_prod fresh_atm) +apply(simp add: fresh_prod fresh_atm) +apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1] +apply(case_tac "aa=b") +apply(simp add: crename_id) +apply(rule l_redu.intros) +apply(simp) +apply(simp add: fresh_atm) +apply(assumption) +apply(frule crename_ax2) +apply(auto)[1] +apply(case_tac "d=aa") +apply(simp add: trm.inject) +apply(simp add: trm.inject) +apply(rule_tac x="N'[x\n>y]" in exI) +apply(rule conjI) +apply(rule sym) +apply(rule crename_interesting3) +apply(rule l_redu.intros) +apply(simp) +apply(simp add: fresh_atm) +apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1] +(* LNot *) +apply(drule sym) +apply(drule crename_Cut) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1] +apply(drule crename_NotR) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1] +apply(drule crename_NotL) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1] +apply(rule_tac x="Cut .N'b (x).N'a" in exI) +apply(simp add: fresh_atm)[1] +apply(rule l_redu.intros) +apply(auto simp add: fresh_prod intro: crename_fresh_interesting2)[1] +apply(auto simp add: fresh_atm fresh_prod intro: crename_fresh_interesting2)[1] +apply(auto simp add: fresh_atm fresh_prod intro: crename_fresh_interesting1)[1] +apply(auto simp add: fresh_atm fresh_prod intro: crename_fresh_interesting1)[1] +apply(simp add: fresh_atm) +apply(simp add: fresh_atm) +(* LAnd1 *) +apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1] +apply(drule sym) +apply(drule crename_Cut) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(auto)[1] +apply(drule crename_AndR) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1] +apply(drule crename_AndL1) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1] +apply(rule_tac x="Cut .M'a (x).N'a" in exI) +apply(simp add: fresh_atm)[1] +apply(rule l_redu.intros) +apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: crename_fresh_interesting1)[1] +apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting2)[1] +apply(auto simp add: fresh_atm fresh_prod intro: crename_fresh_interesting1)[1] +apply(auto simp add: fresh_atm fresh_prod intro: crename_fresh_interesting1)[1] +apply(simp add: fresh_atm) +apply(simp add: fresh_atm) +(* LAnd2 *) +apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1] +apply(drule sym) +apply(drule crename_Cut) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(auto)[1] +apply(drule crename_AndR) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1] +apply(drule crename_AndL2) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1] +apply(rule_tac x="Cut .N'b (x).N'a" in exI) +apply(simp add: fresh_atm)[1] +apply(rule l_redu.intros) +apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: crename_fresh_interesting1)[1] +apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting2)[1] +apply(auto simp add: fresh_atm fresh_prod intro: crename_fresh_interesting1)[1] +apply(auto simp add: fresh_atm fresh_prod intro: crename_fresh_interesting1)[1] +apply(simp add: fresh_atm) +apply(simp add: fresh_atm) +(* LOr1 *) +apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1] +apply(drule sym) +apply(drule crename_Cut) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(auto)[1] +apply(drule crename_OrL) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1] +apply(drule crename_OrR1) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1] +apply(auto) +apply(rule_tac x="Cut .N' (x1).M'a" in exI) +apply(rule conjI) +apply(simp add: abs_fresh fresh_atm)[1] +apply(rule l_redu.intros) +apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: crename_fresh_interesting1)[1] +apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting2)[1] +apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting1)[1] +apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting1)[1] +apply(simp add: fresh_atm) +apply(simp add: fresh_atm) +(* LOr2 *) +apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1] +apply(drule sym) +apply(drule crename_Cut) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(auto)[1] +apply(drule crename_OrL) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1] +apply(drule crename_OrR2) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1] +apply(auto) +apply(rule_tac x="Cut .N' (x2).N'a" in exI) +apply(rule conjI) +apply(simp add: abs_fresh fresh_atm)[1] +apply(rule l_redu.intros) +apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: crename_fresh_interesting1)[1] +apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting2)[1] +apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting1)[1] +apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: crename_fresh_interesting1)[1] +apply(simp add: fresh_atm) +apply(simp add: fresh_atm) +(* ImpL *) +apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1] +apply(drule sym) +apply(drule crename_Cut) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(simp add: fresh_prod abs_fresh fresh_atm abs_supp fin_supp) +apply(auto)[1] +apply(drule crename_ImpL) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1] +apply(drule crename_ImpR) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(simp) +apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1] +apply(rule_tac x="Cut .(Cut .M'a (x).N') (y).N'a" in exI) +apply(rule conjI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(rule l_redu.intros) +apply(auto simp add: fresh_atm abs_fresh abs_supp fin_supp fresh_prod intro: crename_fresh_interesting2)[1] +apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: crename_fresh_interesting1)[1] +apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: crename_fresh_interesting2)[1] +apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: crename_fresh_interesting1)[1] +apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: crename_fresh_interesting1)[1] +apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: crename_fresh_interesting1)[1] +done + +lemma crename_aredu: + assumes a: "(M[a\c>b]) \\<^isub>a M'" "a\b" + shows "\M0. M0[a\c>b]=M' \ M \\<^isub>a M0" +using a +apply(nominal_induct M\"M[a\c>b]" M' avoiding: M a b rule: a_redu.strong_induct) +apply(simp) +apply(drule crename_lredu) +apply(blast) +apply(simp) +apply(drule crename_credu) +apply(blast) +(* Cut *) +apply(drule sym) +apply(drule crename_Cut) +apply(simp) +apply(simp) +apply(auto)[1] +apply(drule_tac x="M'a" in meta_spec) +apply(drule_tac x="aa" in meta_spec) +apply(drule_tac x="b" in meta_spec) +apply(auto)[1] +apply(rule_tac x="Cut .M0 (x).N'" in exI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(rule conjI) +apply(rule trans) +apply(rule crename.simps) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(drule crename_fresh_interesting2) +apply(simp add: fresh_a_redu) +apply(simp) +apply(auto)[1] +apply(drule sym) +apply(drule crename_Cut) +apply(simp) +apply(simp) +apply(auto)[1] +apply(drule_tac x="N'a" in meta_spec) +apply(drule_tac x="aa" in meta_spec) +apply(drule_tac x="b" in meta_spec) +apply(auto)[1] +apply(rule_tac x="Cut .M' (x).M0" in exI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(rule conjI) +apply(rule trans) +apply(rule crename.simps) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1] +apply(drule crename_fresh_interesting1) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(simp add: fresh_a_redu) +apply(simp) +apply(simp) +apply(auto)[1] +(* NotL *) +apply(drule sym) +apply(drule crename_NotL) +apply(simp) +apply(auto)[1] +apply(drule_tac x="N'" in meta_spec) +apply(drule_tac x="aa" in meta_spec) +apply(drule_tac x="b" in meta_spec) +apply(auto)[1] +apply(rule_tac x="NotL .M0 x" in exI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +(* NotR *) +apply(drule sym) +apply(frule crename_NotR_aux) +apply(erule disjE) +apply(auto)[1] +apply(drule crename_NotR') +apply(simp) +apply(simp add: fresh_atm) +apply(erule disjE) +apply(auto)[1] +apply(drule_tac x="N'" in meta_spec) +apply(drule_tac x="aa" in meta_spec) +apply(drule_tac x="b" in meta_spec) +apply(auto)[1] +apply(rule_tac x="NotR (x).M0 a" in exI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +apply(auto)[1] +apply(drule_tac x="N'" in meta_spec) +apply(drule_tac x="aa" in meta_spec) +apply(drule_tac x="a" in meta_spec) +apply(auto)[1] +apply(rule_tac x="NotR (x).M0 aa" in exI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +(* AndR *) +apply(drule sym) +apply(frule crename_AndR_aux) +apply(erule disjE) +apply(auto)[1] +apply(drule crename_AndR') +apply(simp add: fresh_prod fresh_atm) +apply(simp add: fresh_atm) +apply(simp add: fresh_atm) +apply(erule disjE) +apply(auto)[1] +apply(drule_tac x="M'a" in meta_spec) +apply(drule_tac x="aa" in meta_spec) +apply(drule_tac x="ba" in meta_spec) +apply(auto)[1] +apply(rule_tac x="AndR .M0 .N' c" in exI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +apply(rule trans) +apply(rule crename.simps) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto intro: fresh_a_redu)[1] +apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +apply(drule_tac x="M'a" in meta_spec) +apply(drule_tac x="aa" in meta_spec) +apply(drule_tac x="c" in meta_spec) +apply(auto)[1] +apply(rule_tac x="AndR .M0 .N' aa" in exI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +apply(rule trans) +apply(rule crename.simps) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto intro: fresh_a_redu)[1] +apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(drule sym) +apply(frule crename_AndR_aux) +apply(erule disjE) +apply(auto)[1] +apply(drule crename_AndR') +apply(simp add: fresh_prod fresh_atm) +apply(simp add: fresh_atm) +apply(simp add: fresh_atm) +apply(erule disjE) +apply(auto)[1] +apply(drule_tac x="N'a" in meta_spec) +apply(drule_tac x="aa" in meta_spec) +apply(drule_tac x="ba" in meta_spec) +apply(auto)[1] +apply(rule_tac x="AndR .M' .M0 c" in exI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +apply(rule trans) +apply(rule crename.simps) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1] +apply(auto intro: fresh_a_redu)[1] +apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +apply(drule_tac x="N'a" in meta_spec) +apply(drule_tac x="aa" in meta_spec) +apply(drule_tac x="c" in meta_spec) +apply(auto)[1] +apply(rule_tac x="AndR .M' .M0 aa" in exI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +apply(rule trans) +apply(rule crename.simps) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1] +apply(auto intro: fresh_a_redu)[1] +apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(simp) +(* AndL1 *) +apply(drule sym) +apply(drule crename_AndL1) +apply(simp) +apply(auto)[1] +apply(drule_tac x="N'" in meta_spec) +apply(drule_tac x="a" in meta_spec) +apply(drule_tac x="b" in meta_spec) +apply(auto)[1] +apply(rule_tac x="AndL1 (x).M0 y" in exI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +(* AndL2 *) +apply(drule sym) +apply(drule crename_AndL2) +apply(simp) +apply(auto)[1] +apply(drule_tac x="N'" in meta_spec) +apply(drule_tac x="a" in meta_spec) +apply(drule_tac x="b" in meta_spec) +apply(auto)[1] +apply(rule_tac x="AndL2 (x).M0 y" in exI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +(* OrL *) +apply(drule sym) +apply(drule crename_OrL) +apply(simp) +apply(auto simp add: fresh_atm fresh_prod)[1] +apply(auto simp add: fresh_atm fresh_prod)[1] +apply(auto)[1] +apply(drule_tac x="M'a" in meta_spec) +apply(drule_tac x="a" in meta_spec) +apply(drule_tac x="b" in meta_spec) +apply(auto)[1] +apply(rule_tac x="OrL (x).M0 (y).N' z" in exI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +apply(rule trans) +apply(rule crename.simps) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1] +apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto intro: fresh_a_redu)[1] +apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(simp) +apply(drule sym) +apply(drule crename_OrL) +apply(simp) +apply(auto simp add: fresh_atm fresh_prod)[1] +apply(auto simp add: fresh_atm fresh_prod)[1] +apply(auto)[1] +apply(drule_tac x="N'a" in meta_spec) +apply(drule_tac x="a" in meta_spec) +apply(drule_tac x="b" in meta_spec) +apply(auto)[1] +apply(rule_tac x="OrL (x).M' (y).M0 z" in exI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +apply(rule trans) +apply(rule crename.simps) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1] +apply(auto intro: fresh_a_redu)[1] +apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(simp) +apply(simp) +(* OrR1 *) +apply(drule sym) +apply(frule crename_OrR1_aux) +apply(erule disjE) +apply(auto)[1] +apply(drule crename_OrR1') +apply(simp) +apply(simp add: fresh_atm) +apply(erule disjE) +apply(auto)[1] +apply(drule_tac x="N'" in meta_spec) +apply(drule_tac x="aa" in meta_spec) +apply(drule_tac x="ba" in meta_spec) +apply(auto)[1] +apply(rule_tac x="OrR1 .M0 b" in exI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +apply(auto)[1] +apply(drule_tac x="N'" in meta_spec) +apply(drule_tac x="aa" in meta_spec) +apply(drule_tac x="b" in meta_spec) +apply(auto)[1] +apply(rule_tac x="OrR1 .M0 aa" in exI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +(* OrR2 *) +apply(drule sym) +apply(frule crename_OrR2_aux) +apply(erule disjE) +apply(auto)[1] +apply(drule crename_OrR2') +apply(simp) +apply(simp add: fresh_atm) +apply(erule disjE) +apply(auto)[1] +apply(drule_tac x="N'" in meta_spec) +apply(drule_tac x="aa" in meta_spec) +apply(drule_tac x="ba" in meta_spec) +apply(auto)[1] +apply(rule_tac x="OrR2 .M0 b" in exI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +apply(auto)[1] +apply(drule_tac x="N'" in meta_spec) +apply(drule_tac x="aa" in meta_spec) +apply(drule_tac x="b" in meta_spec) +apply(auto)[1] +apply(rule_tac x="OrR2 .M0 aa" in exI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +(* ImpL *) +apply(drule sym) +apply(drule crename_ImpL) +apply(simp) +apply(simp) +apply(auto)[1] +apply(drule_tac x="M'a" in meta_spec) +apply(drule_tac x="aa" in meta_spec) +apply(drule_tac x="b" in meta_spec) +apply(auto)[1] +apply(rule_tac x="ImpL .M0 (x).N' y" in exI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +apply(rule trans) +apply(rule crename.simps) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1] +apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto intro: fresh_a_redu)[1] +apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(drule sym) +apply(drule crename_ImpL) +apply(simp) +apply(simp) +apply(auto)[1] +apply(drule_tac x="N'a" in meta_spec) +apply(drule_tac x="aa" in meta_spec) +apply(drule_tac x="b" in meta_spec) +apply(auto)[1] +apply(rule_tac x="ImpL .M' (x).M0 y" in exI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +apply(rule trans) +apply(rule crename.simps) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1] +apply(auto intro: fresh_a_redu)[1] +apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(simp) +(* ImpR *) +apply(drule sym) +apply(frule crename_ImpR_aux) +apply(erule disjE) +apply(auto)[1] +apply(drule crename_ImpR') +apply(simp) +apply(simp add: fresh_atm) +apply(simp add: fresh_atm) +apply(erule disjE) +apply(auto)[1] +apply(drule_tac x="N'" in meta_spec) +apply(drule_tac x="aa" in meta_spec) +apply(drule_tac x="ba" in meta_spec) +apply(auto)[1] +apply(rule_tac x="ImpR (x)..M0 b" in exI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +apply(auto)[1] +apply(drule_tac x="N'" in meta_spec) +apply(drule_tac x="aa" in meta_spec) +apply(drule_tac x="b" in meta_spec) +apply(auto)[1] +apply(rule_tac x="ImpR (x)..M0 aa" in exI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +done + +lemma SNa_preserved_renaming1: + assumes a: "SNa M" + shows "SNa (M[a\c>b])" +using a +apply(induct rule: SNa_induct) +apply(case_tac "a=b") +apply(simp add: crename_id) +apply(rule SNaI) +apply(drule crename_aredu) +apply(blast)+ +done + +lemma nrename_interesting1: + assumes a: "distinct [x,y,z]" + shows "M[x\n>z][z\n>y] = M[z\n>y][x\n>y]" +using a +apply(nominal_induct M avoiding: x y z rule: trm.induct) +apply(auto simp add: rename_fresh simp add: trm.inject alpha) +apply(blast) +apply(blast) +apply(rotate_tac 12) +apply(drule_tac x="x" in meta_spec) +apply(rotate_tac 15) +apply(drule_tac x="y" in meta_spec) +apply(rotate_tac 15) +apply(drule_tac x="z" in meta_spec) +apply(blast) +apply(rotate_tac 11) +apply(drule_tac x="x" in meta_spec) +apply(rotate_tac 14) +apply(drule_tac x="y" in meta_spec) +apply(rotate_tac 14) +apply(drule_tac x="z" in meta_spec) +apply(blast) +done + +lemma nrename_interesting2: + assumes a: "x\z" "x\u" "x\y" "z\u" "y\z" + shows "M[x\n>y][z\n>u] = M[z\n>u][x\n>y]" +using a +apply(nominal_induct M avoiding: x y z u rule: trm.induct) +apply(auto simp add: rename_fresh simp add: trm.inject alpha) +done + +lemma not_fic_nrename_aux: + assumes a: "fic M c" + shows "fic (M[x\n>y]) c" +using a +apply(nominal_induct M avoiding: c x y rule: trm.induct) +apply(auto dest!: fic_elims intro!: fic.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh) +done + +lemma not_fic_nrename: + assumes a: "\(fic (M[x\n>y]) c)" + shows "\(fic M c)" +using a +apply(auto dest: not_fic_nrename_aux) +done + +lemma fin_nrename: + assumes a: "fin M z" "z\(x,y)" + shows "fin (M[x\n>y]) z" +using a +apply(nominal_induct M avoiding: x y z rule: trm.induct) +apply(auto dest!: fin_elims intro!: fin.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh + split: if_splits) +done + +lemma nrename_fresh_interesting1: + fixes z::"name" + assumes a: "z\(M[x\n>y])" "z\(x,y)" + shows "z\M" +using a +apply(nominal_induct M avoiding: x y z rule: trm.induct) +apply(auto split: if_splits simp add: abs_fresh abs_supp fin_supp) +done + +lemma nrename_fresh_interesting2: + fixes c::"coname" + assumes a: "c\(M[x\n>y])" + shows "c\M" +using a +apply(nominal_induct M avoiding: x y c rule: trm.induct) +apply(auto split: if_splits simp add: abs_fresh abs_supp fin_supp fresh_atm) +done + +lemma fin_nrename2: + assumes a: "fin (M[x\n>y]) z" "z\(x,y)" + shows "fin M z" +using a +apply(nominal_induct M avoiding: x y z rule: trm.induct) +apply(auto dest!: fin_elims intro!: fin.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh + split: if_splits) +apply(auto dest: nrename_fresh_interesting1 simp add: fresh_atm fresh_prod) +done + +lemma nrename_Cut: + assumes a: "R[x\n>y] = Cut .M (z).N" "c\(N,R)" "z\(x,y,M,R)" + shows "\M' N'. R = Cut .M' (z).N' \ M'[x\n>y] = M \ N'[x\n>y] = N \ c\N' \ z\M'" +using a +apply(nominal_induct R avoiding: c y x z M N rule: trm.induct) +apply(auto split: if_splits) +apply(simp add: trm.inject) +apply(auto simp add: alpha fresh_atm) +apply(rule_tac x="[(coname,c)]\trm1" in exI) +apply(perm_simp) +apply(simp add: abs_fresh fresh_left calc_atm fresh_prod) +apply(rule_tac x="[(name,z)]\trm2" in exI) +apply(perm_simp) +apply(simp add: abs_fresh fresh_left calc_atm fresh_prod) +apply(rule conjI) +apply(drule sym) +apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst]) +apply(simp add: eqvts calc_atm) +apply(auto simp add: fresh_atm)[1] +apply(drule sym) +apply(drule sym) +apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst]) +apply(simp add: eqvts calc_atm) +done + +lemma nrename_NotR: + assumes a: "R[x\n>y] = NotR (z).N c" "z\(R,x,y)" + shows "\N'. (R = NotR (z).N' c) \ N'[x\n>y] = N" +using a +apply(nominal_induct R avoiding: x y c z N rule: trm.induct) +apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject) +apply(rule_tac x="[(name,z)]\trm" in exI) +apply(perm_simp) +apply(simp add: abs_fresh fresh_left calc_atm fresh_prod) +apply(drule sym) +apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst]) +apply(simp add: eqvts calc_atm) +done + +lemma nrename_NotL: + assumes a: "R[x\n>y] = NotL .N z" "c\R" "z\(x,y)" + shows "\N'. (R = NotL .N' z) \ N'[x\n>y] = N" +using a +apply(nominal_induct R avoiding: x y c z N rule: trm.induct) +apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject) +apply(rule_tac x="[(coname,c)]\trm" in exI) +apply(perm_simp) +apply(simp add: abs_fresh fresh_left calc_atm fresh_prod) +apply(drule sym) +apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst]) +apply(simp add: eqvts calc_atm) +done + +lemma nrename_NotL': + assumes a: "R[x\n>y] = NotL .N u" "c\R" "x\y" + shows "(\N'. (R = NotL .N' u) \ N'[x\n>y] = N) \ (\N'. (R = NotL .N' x) \ y=u \ N'[x\n>y] = N)" +using a +apply(nominal_induct R avoiding: y u c x N rule: trm.induct) +apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_fresh alpha trm.inject) +apply(rule_tac x="[(coname,c)]\trm" in exI) +apply(perm_simp) +apply(simp add: abs_fresh fresh_left calc_atm fresh_prod) +apply(drule sym) +apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst]) +apply(simp add: eqvts calc_atm) +apply(rule_tac x="[(coname,c)]\trm" in exI) +apply(perm_simp) +apply(simp add: abs_fresh fresh_left calc_atm fresh_prod) +apply(drule sym) +apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst]) +apply(simp add: eqvts calc_atm) +done + +lemma nrename_NotL_aux: + assumes a: "R[x\n>y] = NotL .N u" + shows "(x=u \ x=y) \ (x\u)" +using a +apply(nominal_induct R avoiding: y u c x N rule: trm.induct) +apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject) +done + +lemma nrename_AndL1: + assumes a: "R[x\n>y] = AndL1 (z).N u" "z\(R,x,y)" "u\(x,y)" + shows "\N'. (R = AndL1 (z).N' u) \ N'[x\n>y] = N" +using a +apply(nominal_induct R avoiding: z u x y N rule: trm.induct) +apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject) +apply(rule_tac x="[(name1,z)]\trm" in exI) +apply(perm_simp) +apply(simp add: abs_fresh fresh_left calc_atm fresh_prod) +apply(drule sym) +apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst]) +apply(simp add: eqvts calc_atm) +done + +lemma nrename_AndL1': + assumes a: "R[x\n>y] = AndL1 (v).N u" "v\(R,u,x,y)" "x\y" + shows "(\N'. (R = AndL1 (v).N' u) \ N'[x\n>y] = N) \ (\N'. (R = AndL1 (v).N' x) \ y=u \ N'[x\n>y] = N)" +using a +apply(nominal_induct R avoiding: y u v x N rule: trm.induct) +apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_fresh alpha trm.inject) +apply(rule_tac x="[(name1,v)]\trm" in exI) +apply(perm_simp) +apply(simp add: abs_fresh fresh_left calc_atm fresh_prod) +apply(drule sym) +apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst]) +apply(simp add: eqvts calc_atm) +apply(rule_tac x="[(name1,v)]\trm" in exI) +apply(perm_simp) +apply(simp add: abs_fresh fresh_left calc_atm fresh_prod) +apply(drule sym) +apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst]) +apply(simp add: eqvts calc_atm) +done + +lemma nrename_AndL1_aux: + assumes a: "R[x\n>y] = AndL1 (v).N u" + shows "(x=u \ x=y) \ (x\u)" +using a +apply(nominal_induct R avoiding: y u v x N rule: trm.induct) +apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject) +done + +lemma nrename_AndL2: + assumes a: "R[x\n>y] = AndL2 (z).N u" "z\(R,x,y)" "u\(x,y)" + shows "\N'. (R = AndL2 (z).N' u) \ N'[x\n>y] = N" +using a +apply(nominal_induct R avoiding: z u x y N rule: trm.induct) +apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject) +apply(rule_tac x="[(name1,z)]\trm" in exI) +apply(perm_simp) +apply(simp add: abs_fresh fresh_left calc_atm fresh_prod) +apply(drule sym) +apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst]) +apply(simp add: eqvts calc_atm) +done + +lemma nrename_AndL2': + assumes a: "R[x\n>y] = AndL2 (v).N u" "v\(R,u,x,y)" "x\y" + shows "(\N'. (R = AndL2 (v).N' u) \ N'[x\n>y] = N) \ (\N'. (R = AndL2 (v).N' x) \ y=u \ N'[x\n>y] = N)" +using a +apply(nominal_induct R avoiding: y u v x N rule: trm.induct) +apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_fresh alpha trm.inject) +apply(rule_tac x="[(name1,v)]\trm" in exI) +apply(perm_simp) +apply(simp add: abs_fresh fresh_left calc_atm fresh_prod) +apply(drule sym) +apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst]) +apply(simp add: eqvts calc_atm) +apply(rule_tac x="[(name1,v)]\trm" in exI) +apply(perm_simp) +apply(simp add: abs_fresh fresh_left calc_atm fresh_prod) +apply(drule sym) +apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst]) +apply(simp add: eqvts calc_atm) +done + +lemma nrename_AndL2_aux: + assumes a: "R[x\n>y] = AndL2 (v).N u" + shows "(x=u \ x=y) \ (x\u)" +using a +apply(nominal_induct R avoiding: y u v x N rule: trm.induct) +apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject) +done + +lemma nrename_AndR: + assumes a: "R[x\n>y] = AndR .M .N e" "c\(d,e,N,R)" "d\(c,e,M,R)" + shows "\M' N'. R = AndR .M' .N' e \ M'[x\n>y] = M \ N'[x\n>y] = N \ c\N' \ d\M'" +using a +apply(nominal_induct R avoiding: x y c d e M N rule: trm.induct) +apply(auto split: if_splits simp add: trm.inject alpha) +apply(simp add: fresh_atm fresh_prod) +apply(rule_tac x="[(coname1,c)]\trm1" in exI) +apply(perm_simp) +apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(rule_tac x="[(coname1,c)]\trm1" in exI) +apply(perm_simp) +apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(rule_tac x="[(coname2,d)]\trm2" in exI) +apply(perm_simp) +apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(drule sym) +apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst]) +apply(simp add: eqvts calc_atm) +apply(drule_tac s="trm2[x\n>y]" in sym) +apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst]) +apply(simp add: eqvts calc_atm) +done + +lemma nrename_OrR1: + assumes a: "R[x\n>y] = OrR1 .N d" "c\(R,d)" + shows "\N'. (R = OrR1 .N' d) \ N'[x\n>y] = N" +using a +apply(nominal_induct R avoiding: x y c d N rule: trm.induct) +apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject) +apply(rule_tac x="[(coname1,c)]\trm" in exI) +apply(perm_simp) +apply(simp add: abs_fresh fresh_left calc_atm fresh_prod) +apply(drule sym) +apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst]) +apply(simp add: eqvts calc_atm) +done + +lemma nrename_OrR2: + assumes a: "R[x\n>y] = OrR2 .N d" "c\(R,d)" + shows "\N'. (R = OrR2 .N' d) \ N'[x\n>y] = N" +using a +apply(nominal_induct R avoiding: x y c d N rule: trm.induct) +apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject) +apply(rule_tac x="[(coname1,c)]\trm" in exI) +apply(perm_simp) +apply(simp add: abs_fresh fresh_left calc_atm fresh_prod) +apply(drule sym) +apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst]) +apply(simp add: eqvts calc_atm) +done + +lemma nrename_OrL: + assumes a: "R[u\n>v] = OrL (x).M (y).N z" "x\(y,z,u,v,N,R)" "y\(x,z,u,v,M,R)" "z\(u,v)" + shows "\M' N'. R = OrL (x).M' (y).N' z \ M'[u\n>v] = M \ N'[u\n>v] = N \ x\N' \ y\M'" +using a +apply(nominal_induct R avoiding: u v x y z M N rule: trm.induct) +apply(auto split: if_splits simp add: trm.inject alpha fresh_prod fresh_atm) +apply(rule_tac x="[(name1,x)]\trm1" in exI) +apply(perm_simp) +apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(rule_tac x="[(name2,y)]\trm2" in exI) +apply(perm_simp) +apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(drule sym) +apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst]) +apply(simp add: eqvts calc_atm) +apply(drule_tac s="trm2[u\n>v]" in sym) +apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst]) +apply(simp add: eqvts calc_atm) +done + +lemma nrename_OrL': + assumes a: "R[x\n>y] = OrL (v).M (w).N u" "v\(R,N,u,x,y)" "w\(R,M,u,x,y)" "x\y" + shows "(\M' N'. (R = OrL (v).M' (w).N' u) \ M'[x\n>y] = M \ N'[x\n>y] = N) \ + (\M' N'. (R = OrL (v).M' (w).N' x) \ y=u \ M'[x\n>y] = M \ N'[x\n>y] = N)" +using a +apply(nominal_induct R avoiding: y x u v w M N rule: trm.induct) +apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_fresh alpha trm.inject) +apply(rule_tac x="[(name1,v)]\trm1" in exI) +apply(perm_simp) +apply(simp add: abs_fresh fresh_left calc_atm fresh_prod) +apply(rule_tac x="[(name2,w)]\trm2" in exI) +apply(perm_simp) +apply(simp add: abs_fresh fresh_left calc_atm fresh_prod) +apply(rule conjI) +apply(drule sym) +apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst]) +apply(simp add: eqvts calc_atm) +apply(drule_tac s="trm2[x\n>u]" in sym) +apply(drule_tac pt_bij1[OF pt_name_inst,OF at_name_inst]) +apply(simp add: eqvts calc_atm) +apply(rule_tac x="[(name1,v)]\trm1" in exI) +apply(perm_simp) +apply(simp add: abs_fresh fresh_left calc_atm fresh_prod) +apply(rule_tac x="[(name2,w)]\trm2" in exI) +apply(perm_simp) +apply(simp add: abs_fresh fresh_left calc_atm fresh_prod) +apply(rule conjI) +apply(drule sym) +apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst]) +apply(simp add: eqvts calc_atm) +apply(drule_tac s="trm2[x\n>y]" in sym) +apply(drule_tac pt_bij1[OF pt_name_inst,OF at_name_inst]) +apply(simp add: eqvts calc_atm) +done + +lemma nrename_OrL_aux: + assumes a: "R[x\n>y] = OrL (v).M (w).N u" + shows "(x=u \ x=y) \ (x\u)" +using a +apply(nominal_induct R avoiding: y x w u v M N rule: trm.induct) +apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject) +done + +lemma nrename_ImpL: + assumes a: "R[x\n>y] = ImpL .M (u).N z" "c\(N,R)" "u\(y,x,z,M,R)" "z\(x,y)" + shows "\M' N'. R = ImpL .M' (u).N' z \ M'[x\n>y] = M \ N'[x\n>y] = N \ c\N' \ u\M'" +using a +apply(nominal_induct R avoiding: u x c y z M N rule: trm.induct) +apply(auto split: if_splits simp add: trm.inject alpha fresh_prod fresh_atm) +apply(rule_tac x="[(coname,c)]\trm1" in exI) +apply(perm_simp) +apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(rule_tac x="[(name1,u)]\trm2" in exI) +apply(perm_simp) +apply(auto simp add: abs_fresh fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(drule sym) +apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst]) +apply(simp add: eqvts calc_atm) +apply(drule_tac s="trm2[x\n>y]" in sym) +apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst]) +apply(simp add: eqvts calc_atm fresh_prod fresh_atm) +done + +lemma nrename_ImpL': + assumes a: "R[x\n>y] = ImpL .M (w).N u" "c\(R,N)" "w\(R,M,u,x,y)" "x\y" + shows "(\M' N'. (R = ImpL .M' (w).N' u) \ M'[x\n>y] = M \ N'[x\n>y] = N) \ + (\M' N'. (R = ImpL .M' (w).N' x) \ y=u \ M'[x\n>y] = M \ N'[x\n>y] = N)" +using a +apply(nominal_induct R avoiding: y x u c w M N rule: trm.induct) +apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_fresh alpha trm.inject) +apply(rule_tac x="[(coname,c)]\trm1" in exI) +apply(perm_simp) +apply(simp add: abs_fresh fresh_left calc_atm fresh_prod) +apply(rule_tac x="[(name1,w)]\trm2" in exI) +apply(perm_simp) +apply(simp add: abs_fresh fresh_left calc_atm fresh_prod) +apply(rule conjI) +apply(drule sym) +apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst]) +apply(simp add: eqvts calc_atm) +apply(drule_tac s="trm2[x\n>u]" in sym) +apply(drule_tac pt_bij1[OF pt_name_inst,OF at_name_inst]) +apply(simp add: eqvts calc_atm) +apply(rule_tac x="[(coname,c)]\trm1" in exI) +apply(perm_simp) +apply(simp add: abs_fresh fresh_left calc_atm fresh_prod) +apply(rule_tac x="[(name1,w)]\trm2" in exI) +apply(perm_simp) +apply(simp add: abs_fresh fresh_left calc_atm fresh_prod) +apply(rule conjI) +apply(drule sym) +apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst]) +apply(simp add: eqvts calc_atm) +apply(drule_tac s="trm2[x\n>y]" in sym) +apply(drule_tac pt_bij1[OF pt_name_inst,OF at_name_inst]) +apply(simp add: eqvts calc_atm) +done + +lemma nrename_ImpL_aux: + assumes a: "R[x\n>y] = ImpL .M (w).N u" + shows "(x=u \ x=y) \ (x\u)" +using a +apply(nominal_induct R avoiding: y x w u c M N rule: trm.induct) +apply(auto split: if_splits simp add: fresh_prod fresh_atm alpha abs_fresh trm.inject) +done + +lemma nrename_ImpR: + assumes a: "R[u\n>v] = ImpR (x)..N d" "c\(R,d)" "x\(R,u,v)" + shows "\N'. (R = ImpR (x)..N' d) \ N'[u\n>v] = N" +using a +apply(nominal_induct R avoiding: u v x c d N rule: trm.induct) +apply(auto split: if_splits simp add: fresh_prod fresh_atm abs_perm alpha abs_fresh trm.inject) +apply(rule_tac x="[(name,x)]\trm" in exI) +apply(perm_simp) +apply(simp add: abs_fresh fresh_left calc_atm fresh_prod) +apply(simp add: calc_atm) +apply(rule_tac x="[(name,x)]\[(coname1, c)]\trm" in exI) +apply(perm_simp) +apply(simp add: abs_supp fin_supp abs_fresh fresh_left calc_atm fresh_prod) +apply(drule sym) +apply(drule pt_bij1[OF pt_coname_inst,OF at_coname_inst]) +apply(drule pt_bij1[OF pt_name_inst,OF at_name_inst]) +apply(simp add: eqvts calc_atm) +done + +lemma nrename_credu: + assumes a: "(M[x\n>y]) \\<^isub>c M'" + shows "\M0. M0[x\n>y]=M' \ M \\<^isub>c M0" +using a +apply(nominal_induct M\"M[x\n>y]" M' avoiding: M x y rule: c_redu.strong_induct) +apply(drule sym) +apply(drule nrename_Cut) +apply(simp) +apply(simp) +apply(auto) +apply(rule_tac x="M'{a:=(x).N'}" in exI) +apply(rule conjI) +apply(simp add: fresh_atm abs_fresh subst_comm fresh_prod) +apply(rule c_redu.intros) +apply(auto dest: not_fic_nrename)[1] +apply(simp add: abs_fresh) +apply(simp add: abs_fresh) +apply(drule sym) +apply(drule nrename_Cut) +apply(simp) +apply(simp) +apply(auto) +apply(rule_tac x="N'{x:=.M'}" in exI) +apply(rule conjI) +apply(simp add: fresh_atm abs_fresh subst_comm fresh_prod) +apply(rule c_redu.intros) +apply(auto) +apply(drule_tac x="xa" and y="y" in fin_nrename) +apply(auto simp add: fresh_prod abs_fresh) +done + +lemma nrename_ax2: + assumes a: "N[x\n>y] = Ax z c" + shows "\z. N = Ax z c" +using a +apply(nominal_induct N avoiding: x y rule: trm.induct) +apply(auto split: if_splits) +apply(simp add: trm.inject) +done + +lemma fic_nrename: + assumes a: "fic (M[x\n>y]) c" + shows "fic M c" +using a +apply(nominal_induct M avoiding: c x y rule: trm.induct) +apply(auto dest!: fic_elims intro!: fic.intros simp add: fresh_prod fresh_atm rename_fresh abs_fresh + split: if_splits) +apply(auto dest: nrename_fresh_interesting2 simp add: fresh_prod fresh_atm) +done + +lemma nrename_lredu: + assumes a: "(M[x\n>y]) \\<^isub>l M'" + shows "\M0. M0[x\n>y]=M' \ M \\<^isub>l M0" +using a +apply(nominal_induct M\"M[x\n>y]" M' avoiding: M x y rule: l_redu.strong_induct) +apply(drule sym) +apply(drule nrename_Cut) +apply(simp add: fresh_prod fresh_atm) +apply(simp) +apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1] +apply(case_tac "xa=y") +apply(simp add: nrename_id) +apply(rule l_redu.intros) +apply(simp) +apply(simp add: fresh_atm) +apply(assumption) +apply(frule nrename_ax2) +apply(auto)[1] +apply(case_tac "z=xa") +apply(simp add: trm.inject) +apply(simp) +apply(rule_tac x="M'[a\c>b]" in exI) +apply(rule conjI) +apply(rule crename_interesting3) +apply(rule l_redu.intros) +apply(simp) +apply(simp add: fresh_atm) +apply(auto dest: fic_nrename simp add: fresh_prod fresh_atm)[1] +apply(drule sym) +apply(drule nrename_Cut) +apply(simp add: fresh_prod fresh_atm) +apply(simp add: fresh_prod fresh_atm) +apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1] +apply(case_tac "xa=ya") +apply(simp add: nrename_id) +apply(rule l_redu.intros) +apply(simp) +apply(simp add: fresh_atm) +apply(assumption) +apply(frule nrename_ax2) +apply(auto)[1] +apply(case_tac "z=xa") +apply(simp add: trm.inject) +apply(rule_tac x="N'[x\n>xa]" in exI) +apply(rule conjI) +apply(rule nrename_interesting1) +apply(auto)[1] +apply(rule l_redu.intros) +apply(simp) +apply(simp add: fresh_atm) +apply(auto dest: fin_nrename2 simp add: fresh_prod fresh_atm)[1] +apply(simp add: trm.inject) +apply(rule_tac x="N'[x\n>y]" in exI) +apply(rule conjI) +apply(rule nrename_interesting2) +apply(simp_all) +apply(rule l_redu.intros) +apply(simp) +apply(simp add: fresh_atm) +apply(auto dest: fin_nrename2 simp add: fresh_prod fresh_atm)[1] +(* LNot *) +apply(drule sym) +apply(drule nrename_Cut) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1] +apply(drule nrename_NotR) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1] +apply(drule nrename_NotL) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1] +apply(rule_tac x="Cut .N'b (x).N'a" in exI) +apply(simp add: fresh_atm)[1] +apply(rule l_redu.intros) +apply(auto simp add: fresh_prod fresh_atm intro: nrename_fresh_interesting1)[1] +apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1] +apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting2)[1] +apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting2)[1] +apply(simp add: fresh_atm) +apply(simp add: fresh_atm) +(* LAnd1 *) +apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1] +apply(drule sym) +apply(drule nrename_Cut) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(auto)[1] +apply(drule nrename_AndR) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1] +apply(drule nrename_AndL1) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1] +apply(rule_tac x="Cut .M'a (x).N'b" in exI) +apply(simp add: fresh_atm)[1] +apply(rule l_redu.intros) +apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: nrename_fresh_interesting2)[1] +apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1] +apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1] +apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1] +apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1] +apply(simp add: fresh_atm) +(* LAnd2 *) +apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1] +apply(drule sym) +apply(drule nrename_Cut) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(auto)[1] +apply(drule nrename_AndR) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1] +apply(drule nrename_AndL2) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1] +apply(rule_tac x="Cut .N'a (x).N'b" in exI) +apply(simp add: fresh_atm)[1] +apply(rule l_redu.intros) +apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: nrename_fresh_interesting2)[1] +apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1] +apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1] +apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1] +apply(auto simp add: fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1] +apply(simp add: fresh_atm) +(* LOr1 *) +apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1] +apply(drule sym) +apply(drule nrename_Cut) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(auto)[1] +apply(drule nrename_OrL) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(simp add: fresh_prod fresh_atm) +apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1] +apply(drule nrename_OrR1) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1] +apply(rule_tac x="Cut .N' (x1).M'a" in exI) +apply(rule conjI) +apply(simp add: abs_fresh fresh_atm)[1] +apply(rule l_redu.intros) +apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: nrename_fresh_interesting2)[1] +apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1] +apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1] +apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1] +apply(simp add: fresh_atm) +apply(simp add: fresh_atm) +(* LOr2 *) +apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1] +apply(drule sym) +apply(drule nrename_Cut) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(auto)[1] +apply(drule nrename_OrL) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(simp add: fresh_prod fresh_atm) +apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1] +apply(drule nrename_OrR2) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1] +apply(rule_tac x="Cut .N' (x2).N'a" in exI) +apply(rule conjI) +apply(simp add: abs_fresh fresh_atm)[1] +apply(rule l_redu.intros) +apply(auto simp add: fresh_atm abs_fresh fresh_prod intro: nrename_fresh_interesting2)[1] +apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1] +apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1] +apply(auto simp add: abs_fresh fresh_atm fresh_prod intro: nrename_fresh_interesting1)[1] +apply(simp add: fresh_atm) +apply(simp add: fresh_atm) +(* ImpL *) +apply(auto dest: fin_crename simp add: fresh_prod fresh_atm)[1] +apply(drule sym) +apply(drule nrename_Cut) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(simp add: fresh_prod abs_fresh fresh_atm abs_supp fin_supp) +apply(auto)[1] +apply(drule nrename_ImpL) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(simp add: fresh_prod fresh_atm) +apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1] +apply(drule nrename_ImpR) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(simp add: fresh_prod abs_fresh fresh_atm) +apply(auto simp add: abs_fresh fresh_prod fresh_atm)[1] +apply(rule_tac x="Cut .(Cut .M'a (x).N') (y).N'a" in exI) +apply(rule conjI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(rule l_redu.intros) +apply(auto simp add: fresh_atm abs_fresh abs_supp fin_supp fresh_prod intro: nrename_fresh_interesting1)[1] +apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: nrename_fresh_interesting2)[1] +apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: nrename_fresh_interesting1)[1] +apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: nrename_fresh_interesting2)[1] +apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: nrename_fresh_interesting2)[1] +apply(auto simp add: abs_fresh fresh_atm abs_supp fin_supp fresh_prod intro: nrename_fresh_interesting2)[1] +done + +lemma nrename_aredu: + assumes a: "(M[x\n>y]) \\<^isub>a M'" "x\y" + shows "\M0. M0[x\n>y]=M' \ M \\<^isub>a M0" +using a +apply(nominal_induct M\"M[x\n>y]" M' avoiding: M x y rule: a_redu.strong_induct) +apply(simp) +apply(drule nrename_lredu) +apply(blast) +apply(simp) +apply(drule nrename_credu) +apply(blast) +(* Cut *) +apply(drule sym) +apply(drule nrename_Cut) +apply(simp) +apply(simp) +apply(auto)[1] +apply(drule_tac x="M'a" in meta_spec) +apply(drule_tac x="xa" in meta_spec) +apply(drule_tac x="y" in meta_spec) +apply(auto)[1] +apply(rule_tac x="Cut .M0 (x).N'" in exI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(rule conjI) +apply(rule trans) +apply(rule nrename.simps) +apply(drule nrename_fresh_interesting2) +apply(simp add: fresh_a_redu) +apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(drule nrename_fresh_interesting1) +apply(simp add: fresh_prod fresh_atm) +apply(simp add: fresh_a_redu) +apply(simp) +apply(auto)[1] +apply(drule sym) +apply(drule nrename_Cut) +apply(simp) +apply(simp) +apply(auto)[1] +apply(drule_tac x="N'a" in meta_spec) +apply(drule_tac x="xa" in meta_spec) +apply(drule_tac x="y" in meta_spec) +apply(auto)[1] +apply(rule_tac x="Cut .M' (x).M0" in exI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(rule conjI) +apply(rule trans) +apply(rule nrename.simps) +apply(simp add: fresh_a_redu) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1] +apply(simp) +apply(auto)[1] +(* NotL *) +apply(drule sym) +apply(frule nrename_NotL_aux) +apply(erule disjE) +apply(auto)[1] +apply(drule nrename_NotL') +apply(simp) +apply(simp add: fresh_atm) +apply(erule disjE) +apply(auto)[1] +apply(drule_tac x="N'" in meta_spec) +apply(drule_tac x="xa" in meta_spec) +apply(drule_tac x="y" in meta_spec) +apply(auto)[1] +apply(rule_tac x="NotL .M0 x" in exI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +apply(auto)[1] +apply(drule_tac x="N'" in meta_spec) +apply(drule_tac x="xa" in meta_spec) +apply(drule_tac x="x" in meta_spec) +apply(auto)[1] +apply(rule_tac x="NotL .M0 xa" in exI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +(* NotR *) +apply(drule sym) +apply(drule nrename_NotR) +apply(simp) +apply(auto)[1] +apply(drule_tac x="N'" in meta_spec) +apply(drule_tac x="xa" in meta_spec) +apply(drule_tac x="y" in meta_spec) +apply(auto)[1] +apply(rule_tac x="NotR (x).M0 a" in exI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +(* AndR *) +apply(drule sym) +apply(drule nrename_AndR) +apply(simp) +apply(auto simp add: fresh_atm fresh_prod)[1] +apply(auto simp add: fresh_atm fresh_prod)[1] +apply(auto)[1] +apply(drule_tac x="M'a" in meta_spec) +apply(drule_tac x="x" in meta_spec) +apply(drule_tac x="y" in meta_spec) +apply(auto)[1] +apply(rule_tac x="AndR .M0 .N' c" in exI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +apply(rule trans) +apply(rule nrename.simps) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1] +apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto intro: fresh_a_redu)[1] +apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(simp) +apply(drule sym) +apply(drule nrename_AndR) +apply(simp) +apply(auto simp add: fresh_atm fresh_prod)[1] +apply(auto simp add: fresh_atm fresh_prod)[1] +apply(auto)[1] +apply(drule_tac x="N'a" in meta_spec) +apply(drule_tac x="x" in meta_spec) +apply(drule_tac x="y" in meta_spec) +apply(auto)[1] +apply(rule_tac x="AndR .M' .M0 c" in exI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +apply(rule trans) +apply(rule nrename.simps) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1] +apply(auto intro: fresh_a_redu)[1] +apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(simp) +apply(simp) +(* AndL1 *) +apply(drule sym) +apply(frule nrename_AndL1_aux) +apply(erule disjE) +apply(auto)[1] +apply(drule nrename_AndL1') +apply(simp) +apply(simp add: fresh_atm) +apply(erule disjE) +apply(auto)[1] +apply(drule_tac x="N'" in meta_spec) +apply(drule_tac x="xa" in meta_spec) +apply(drule_tac x="ya" in meta_spec) +apply(auto)[1] +apply(rule_tac x="AndL1 (x).M0 y" in exI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +apply(auto)[1] +apply(drule_tac x="N'" in meta_spec) +apply(drule_tac x="xa" in meta_spec) +apply(drule_tac x="y" in meta_spec) +apply(auto)[1] +apply(rule_tac x="AndL1 (x).M0 xa" in exI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +(* AndL2 *) +apply(drule sym) +apply(frule nrename_AndL2_aux) +apply(erule disjE) +apply(auto)[1] +apply(drule nrename_AndL2') +apply(simp) +apply(simp add: fresh_atm) +apply(erule disjE) +apply(auto)[1] +apply(drule_tac x="N'" in meta_spec) +apply(drule_tac x="xa" in meta_spec) +apply(drule_tac x="ya" in meta_spec) +apply(auto)[1] +apply(rule_tac x="AndL2 (x).M0 y" in exI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +apply(auto)[1] +apply(drule_tac x="N'" in meta_spec) +apply(drule_tac x="xa" in meta_spec) +apply(drule_tac x="y" in meta_spec) +apply(auto)[1] +apply(rule_tac x="AndL2 (x).M0 xa" in exI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +(* OrL *) +apply(drule sym) +apply(frule nrename_OrL_aux) +apply(erule disjE) +apply(auto)[1] +apply(drule nrename_OrL') +apply(simp add: fresh_prod fresh_atm) +apply(simp add: fresh_atm) +apply(simp add: fresh_atm) +apply(erule disjE) +apply(auto)[1] +apply(drule_tac x="M'a" in meta_spec) +apply(drule_tac x="xa" in meta_spec) +apply(drule_tac x="ya" in meta_spec) +apply(auto)[1] +apply(rule_tac x="OrL (x).M0 (y).N' z" in exI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +apply(rule trans) +apply(rule nrename.simps) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto intro: fresh_a_redu)[1] +apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +apply(drule_tac x="M'a" in meta_spec) +apply(drule_tac x="xa" in meta_spec) +apply(drule_tac x="z" in meta_spec) +apply(auto)[1] +apply(rule_tac x="OrL (x).M0 (y).N' xa" in exI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +apply(rule trans) +apply(rule nrename.simps) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto intro: fresh_a_redu)[1] +apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(drule sym) +apply(frule nrename_OrL_aux) +apply(erule disjE) +apply(auto)[1] +apply(drule nrename_OrL') +apply(simp add: fresh_prod fresh_atm) +apply(simp add: fresh_atm) +apply(simp add: fresh_atm) +apply(erule disjE) +apply(auto)[1] +apply(drule_tac x="N'a" in meta_spec) +apply(drule_tac x="xa" in meta_spec) +apply(drule_tac x="ya" in meta_spec) +apply(auto)[1] +apply(rule_tac x="OrL (x).M' (y).M0 z" in exI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +apply(rule trans) +apply(rule nrename.simps) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1] +apply(auto intro: fresh_a_redu)[1] +apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +apply(drule_tac x="N'a" in meta_spec) +apply(drule_tac x="xa" in meta_spec) +apply(drule_tac x="z" in meta_spec) +apply(auto)[1] +apply(rule_tac x="OrL (x).M' (y).M0 xa" in exI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +apply(rule trans) +apply(rule nrename.simps) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm fresh_prod)[1] +apply(auto intro: fresh_a_redu)[1] +apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(simp) +apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1] +(* OrR1 *) +apply(drule sym) +apply(drule nrename_OrR1) +apply(simp) +apply(auto)[1] +apply(drule_tac x="N'" in meta_spec) +apply(drule_tac x="x" in meta_spec) +apply(drule_tac x="y" in meta_spec) +apply(auto)[1] +apply(rule_tac x="OrR1 .M0 b" in exI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +(* OrR2 *) +apply(drule sym) +apply(drule nrename_OrR2) +apply(simp) +apply(auto)[1] +apply(drule_tac x="N'" in meta_spec) +apply(drule_tac x="x" in meta_spec) +apply(drule_tac x="y" in meta_spec) +apply(auto)[1] +apply(rule_tac x="OrR2 .M0 b" in exI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +(* ImpL *) +apply(drule sym) +apply(frule nrename_ImpL_aux) +apply(erule disjE) +apply(auto)[1] +apply(drule nrename_ImpL') +apply(simp add: fresh_prod fresh_atm) +apply(simp add: fresh_atm) +apply(simp add: fresh_atm) +apply(erule disjE) +apply(auto)[1] +apply(drule_tac x="M'a" in meta_spec) +apply(drule_tac x="xa" in meta_spec) +apply(drule_tac x="ya" in meta_spec) +apply(auto)[1] +apply(rule_tac x="ImpL .M0 (x).N' y" in exI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +apply(rule trans) +apply(rule nrename.simps) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto intro: fresh_a_redu)[1] +apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +apply(drule_tac x="M'a" in meta_spec) +apply(drule_tac x="xa" in meta_spec) +apply(drule_tac x="y" in meta_spec) +apply(auto)[1] +apply(rule_tac x="ImpL .M0 (x).N' xa" in exI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +apply(rule trans) +apply(rule nrename.simps) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto intro: fresh_a_redu)[1] +apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(drule sym) +apply(frule nrename_ImpL_aux) +apply(erule disjE) +apply(auto)[1] +apply(drule nrename_ImpL') +apply(simp add: fresh_prod fresh_atm) +apply(simp add: fresh_atm) +apply(simp add: fresh_atm) +apply(erule disjE) +apply(auto)[1] +apply(drule_tac x="N'a" in meta_spec) +apply(drule_tac x="xa" in meta_spec) +apply(drule_tac x="ya" in meta_spec) +apply(auto)[1] +apply(rule_tac x="ImpL .M' (x).M0 y" in exI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +apply(rule trans) +apply(rule nrename.simps) +apply(auto intro: fresh_a_redu)[1] +apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +apply(drule_tac x="N'a" in meta_spec) +apply(drule_tac x="xa" in meta_spec) +apply(drule_tac x="y" in meta_spec) +apply(auto)[1] +apply(rule_tac x="ImpL .M' (x).M0 xa" in exI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +apply(rule trans) +apply(rule nrename.simps) +apply(auto intro: fresh_a_redu)[1] +apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(simp add: fresh_prod abs_fresh abs_supp fin_supp fresh_atm)[1] +(* ImpR *) +apply(drule sym) +apply(drule nrename_ImpR) +apply(simp) +apply(simp) +apply(auto)[1] +apply(drule_tac x="N'" in meta_spec) +apply(drule_tac x="xa" in meta_spec) +apply(drule_tac x="y" in meta_spec) +apply(auto)[1] +apply(rule_tac x="ImpR (x)..M0 b" in exI) +apply(simp add: abs_fresh abs_supp fin_supp fresh_atm)[1] +apply(auto)[1] +done + +lemma SNa_preserved_renaming2: + assumes a: "SNa N" + shows "SNa (N[x\n>y])" +using a +apply(induct rule: SNa_induct) +apply(case_tac "x=y") +apply(simp add: nrename_id) +apply(rule SNaI) +apply(drule nrename_aredu) +apply(blast)+ +done + +text {* helper-stuff to set up the induction *} + +abbreviation + SNa_set :: "trm set" +where + "SNa_set \ {M. SNa M}" + +abbreviation + A_Redu_set :: "(trm\trm) set" +where + "A_Redu_set \ {(N,M)| M N. M \\<^isub>a N}" + +lemma SNa_elim: + assumes a: "SNa M" + shows "(\M. (\N. M \\<^isub>a N \ P N)\ P M) \ P M" +using a +by (induct rule: SNa.induct) (blast) + +lemma wf_SNa_restricted: + shows "wf (A_Redu_set \ (UNIV <*> SNa_set))" +apply(unfold wf_def) +apply(intro strip) +apply(case_tac "SNa x") +apply(simp (no_asm_use)) +apply(drule_tac P="P" in SNa_elim) +apply(erule mp) +apply(blast) +(* other case *) +apply(drule_tac x="x" in spec) +apply(erule mp) +apply(fast) +done + +constdefs + SNa_Redu :: "(trm \ trm) set" + "SNa_Redu \ A_Redu_set \ (UNIV <*> SNa_set)" + +lemma wf_SNa_Redu: + shows "wf SNa_Redu" +apply(unfold SNa_Redu_def) +apply(rule wf_SNa_restricted) +done + +lemma wf_measure_triple: +shows "wf ((measure size) <*lex*> SNa_Redu <*lex*> SNa_Redu)" +by (auto intro: wf_SNa_Redu) + +lemma my_wf_induct_triple: + assumes a: " wf(r1 <*lex*> r2 <*lex*> r3)" + and b: "\x. \\y. ((fst y,fst (snd y),snd (snd y)),(fst x,fst (snd x), snd (snd x))) + \ (r1 <*lex*> r2 <*lex*> r3) \ P y\ \ P x" + shows "P x" +using a +apply(induct x rule: wf_induct_rule) +apply(rule b) +apply(simp) +done + +lemma my_wf_induct_triple': + assumes a: " wf(r1 <*lex*> r2 <*lex*> r3)" + and b: "\x1 x2 x3. \\y1 y2 y3. ((y1,y2,y3),(x1,x2,x3)) \ (r1 <*lex*> r2 <*lex*> r3) \ P (y1,y2,y3)\ + \ P (x1,x2,x3)" + shows "P (x1,x2,x3)" +apply(rule_tac my_wf_induct_triple[OF a]) +apply(case_tac x) +apply(simp) +apply(case_tac b) +apply(simp) +apply(rule b) +apply(blast) +done + +lemma my_wf_induct_triple'': + assumes a: " wf(r1 <*lex*> r2 <*lex*> r3)" + and b: "\x1 x2 x3. \\y1 y2 y3. ((y1,y2,y3),(x1,x2,x3)) \ (r1 <*lex*> r2 <*lex*> r3) \ P y1 y2 y3\ + \ P x1 x2 x3" + shows "P x1 x2 x3" +apply(rule_tac my_wf_induct_triple'[where P="\(x1,x2,x3). P x1 x2 x3", simplified]) +apply(rule a) +apply(rule b) +apply(auto) +done + +lemma excluded_m: + assumes a: ":M \ (\\)" "(x):N \ (\(B)\)" + shows "(:M \ BINDINGc B (\(B)\) \ (x):N \ BINDINGn B (\\)) + \\(:M \ BINDINGc B (\(B)\) \ (x):N \ BINDINGn B (\\))" +by (blast) + +lemma tricky_subst: + assumes a1: "b\(c,N)" + and a2: "z\(x,P)" + and a3: "M\Ax z b" + shows "(Cut .N (z).M){b:=(x).P} = Cut .N (z).(M{b:=(x).P})" +using a1 a2 a3 +apply - +apply(generate_fresh "coname") +apply(subgoal_tac "Cut .N (z).M = Cut .([(ca,c)]\N) (z).M") +apply(simp) +apply(rule trans) +apply(rule better_Cut_substc) +apply(simp) +apply(simp add: abs_fresh) +apply(simp) +apply(subgoal_tac "b\([(ca,c)]\N)") +apply(simp add: forget) +apply(simp add: trm.inject) +apply(simp add: fresh_left calc_atm fresh_prod fresh_atm) +apply(simp add: trm.inject) +apply(rule sym) +apply(simp add: alpha fresh_prod fresh_atm) +done + +text {* 3rd lemma *} + +lemma CUT_SNa_aux: + assumes a1: ":M \ (\\)" + and a2: "SNa M" + and a3: "(x):N \ (\(B)\)" + and a4: "SNa N" + shows "SNa (Cut .M (x).N)" +using a1 a2 a3 a4 +apply(induct B M N arbitrary: a x rule: my_wf_induct_triple''[OF wf_measure_triple]) +apply(rule SNaI) +apply(drule Cut_a_redu_elim) +apply(erule disjE) +(* left-inner reduction *) +apply(erule exE) +apply(erule conjE)+ +apply(simp) +apply(drule_tac x="x1" in meta_spec) +apply(drule_tac x="M'a" in meta_spec) +apply(drule_tac x="x3" in meta_spec) +apply(drule conjunct2) +apply(drule mp) +apply(rule conjI) +apply(simp) +apply(rule disjI1) +apply(simp add: SNa_Redu_def) +apply(drule_tac x="a" in spec) +apply(drule mp) +apply(simp add: CANDs_preserved_single) +apply(drule mp) +apply(simp add: a_preserves_SNa) +apply(drule_tac x="x" in spec) +apply(simp) +apply(erule disjE) +(* right-inner reduction *) +apply(erule exE) +apply(erule conjE)+ +apply(simp) +apply(drule_tac x="x1" in meta_spec) +apply(drule_tac x="x2" in meta_spec) +apply(drule_tac x="N'" in meta_spec) +apply(drule conjunct2) +apply(drule mp) +apply(rule conjI) +apply(simp) +apply(rule disjI2) +apply(simp add: SNa_Redu_def) +apply(drule_tac x="a" in spec) +apply(drule mp) +apply(simp add: CANDs_preserved_single) +apply(drule mp) +apply(assumption) +apply(drule_tac x="x" in spec) +apply(drule mp) +apply(simp add: CANDs_preserved_single) +apply(drule mp) +apply(simp add: a_preserves_SNa) +apply(assumption) +apply(erule disjE) +(******** c-reduction *********) +apply(drule Cut_c_redu_elim) +(* c-left reduction*) +apply(erule disjE) +apply(erule conjE) +apply(frule_tac B="x1" in fic_CANDS) +apply(simp) +apply(erule disjE) +(* in AXIOMSc *) +apply(simp add: AXIOMSc_def) +apply(erule exE)+ +apply(simp add: ctrm.inject) +apply(simp add: alpha) +apply(erule disjE) +apply(simp) +apply(rule impI) +apply(simp) +apply(subgoal_tac "fic (Ax y b) b")(*A*) +apply(simp) +(*A*) +apply(auto)[1] +apply(simp) +apply(rule impI) +apply(simp) +apply(subgoal_tac "fic (Ax ([(a,aa)]\y) a) a")(*B*) +apply(simp) +(*B*) +apply(auto)[1] +(* in BINDINGc *) +apply(simp) +apply(drule BINDINGc_elim) +apply(simp) +(* c-right reduction*) +apply(erule conjE) +apply(frule_tac B="x1" in fin_CANDS) +apply(simp) +apply(erule disjE) +(* in AXIOMSc *) +apply(simp add: AXIOMSn_def) +apply(erule exE)+ +apply(simp add: ntrm.inject) +apply(simp add: alpha) +apply(erule disjE) +apply(simp) +apply(rule impI) +apply(simp) +apply(subgoal_tac "fin (Ax xa b) xa")(*A*) +apply(simp) +(*A*) +apply(auto)[1] +apply(simp) +apply(rule impI) +apply(simp) +apply(subgoal_tac "fin (Ax x ([(x,xa)]\b)) x")(*B*) +apply(simp) +(*B*) +apply(auto)[1] +(* in BINDINGc *) +apply(simp) +apply(drule BINDINGn_elim) +apply(simp) +(*********** l-reductions ************) +apply(drule Cut_l_redu_elim) +apply(erule disjE) +(* ax1 *) +apply(erule exE) +apply(simp) +apply(simp add: SNa_preserved_renaming1) +apply(erule disjE) +(* ax2 *) +apply(erule exE) +apply(simp add: SNa_preserved_renaming2) +apply(erule disjE) +(* LNot *) +apply(erule exE)+ +apply(auto)[1] +apply(frule_tac excluded_m) +apply(assumption) +apply(erule disjE) +(* one of them in BINDING *) +apply(erule disjE) +apply(drule fin_elims) +apply(drule fic_elims) +apply(simp) +apply(drule BINDINGc_elim) +apply(drule_tac x="x" in spec) +apply(drule_tac x="NotL .N' x" in spec) +apply(simp) +apply(simp add: better_NotR_substc) +apply(generate_fresh "coname") +apply(subgoal_tac "fresh_fun (\a'. Cut .NotR (y).M'a a' (x).NotL .N' x) + = Cut .NotR (y).M'a c (x).NotL .N' x") +apply(simp) +apply(subgoal_tac "Cut .NotR (y).M'a c (x).NotL .N' x \\<^isub>a Cut .N' (y).M'a") +apply(simp only: a_preserves_SNa) +apply(rule al_redu) +apply(rule better_LNot_intro) +apply(simp) +apply(simp) +apply(fresh_fun_simp) +apply(fresh_fun_simp) +apply(simp) +(* other case of in BINDING *) +apply(drule fin_elims) +apply(drule fic_elims) +apply(simp) +apply(drule BINDINGn_elim) +apply(drule_tac x="a" in spec) +apply(drule_tac x="NotR (y).M'a a" in spec) +apply(simp) +apply(simp add: better_NotL_substn) +apply(generate_fresh "name") +apply(subgoal_tac "fresh_fun (\x'. Cut .NotR (y).M'a a (x').NotL .N' x') + = Cut .NotR (y).M'a a (c).NotL .N' c") +apply(simp) +apply(subgoal_tac "Cut .NotR (y).M'a a (c).NotL .N' c \\<^isub>a Cut .N' (y).M'a") +apply(simp only: a_preserves_SNa) +apply(rule al_redu) +apply(rule better_LNot_intro) +apply(simp) +apply(simp) +apply(fresh_fun_simp) +apply(fresh_fun_simp) +apply(simp) +(* none of them in BINDING *) +apply(simp) +apply(erule conjE) +apply(frule CAND_NotR_elim) +apply(assumption) +apply(erule exE) +apply(frule CAND_NotL_elim) +apply(assumption) +apply(erule exE) +apply(simp only: ty.inject) +apply(drule_tac x="B'" in meta_spec) +apply(drule_tac x="N'" in meta_spec) +apply(drule_tac x="M'a" in meta_spec) +apply(erule conjE)+ +apply(drule mp) +apply(simp) +apply(drule_tac x="b" in spec) +apply(rotate_tac 13) +apply(drule mp) +apply(assumption) +apply(rotate_tac 13) +apply(drule mp) +apply(simp add: CANDs_imply_SNa) +apply(drule_tac x="y" in spec) +apply(rotate_tac 13) +apply(drule mp) +apply(assumption) +apply(rotate_tac 13) +apply(drule mp) +apply(simp add: CANDs_imply_SNa) +apply(assumption) +(* LAnd1 case *) +apply(erule disjE) +apply(erule exE)+ +apply(auto)[1] +apply(frule_tac excluded_m) +apply(assumption) +apply(erule disjE) +(* one of them in BINDING *) +apply(erule disjE) +apply(drule fin_elims) +apply(drule fic_elims) +apply(simp) +apply(drule BINDINGc_elim) +apply(drule_tac x="x" in spec) +apply(drule_tac x="AndL1 (y).N' x" in spec) +apply(simp) +apply(simp add: better_AndR_substc) +apply(generate_fresh "coname") +apply(subgoal_tac "fresh_fun (\a'. Cut .AndR .M1 .M2 a' (x).AndL1 (y).N' x) + = Cut .AndR .M1 .M2 ca (x).AndL1 (y).N' x") +apply(simp) +apply(subgoal_tac "Cut .AndR .M1 .M2 ca (x).AndL1 (y).N' x \\<^isub>a Cut .M1 (y).N'") +apply(auto intro: a_preserves_SNa)[1] +apply(rule al_redu) +apply(rule better_LAnd1_intro) +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(simp) +apply(fresh_fun_simp) +apply(fresh_fun_simp) +apply(simp) +(* other case of in BINDING *) +apply(drule fin_elims) +apply(drule fic_elims) +apply(simp) +apply(drule BINDINGn_elim) +apply(drule_tac x="a" in spec) +apply(drule_tac x="AndR .M1 .M2 a" in spec) +apply(simp) +apply(simp add: better_AndL1_substn) +apply(generate_fresh "name") +apply(subgoal_tac "fresh_fun (\z'. Cut .AndR .M1 .M2 a (z').AndL1 (y).N' z') + = Cut .AndR .M1 .M2 a (ca).AndL1 (y).N' ca") +apply(simp) +apply(subgoal_tac "Cut .AndR .M1 .M2 a (ca).AndL1 (y).N' ca \\<^isub>a Cut .M1 (y).N'") +apply(auto intro: a_preserves_SNa)[1] +apply(rule al_redu) +apply(rule better_LAnd1_intro) +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(fresh_fun_simp) +apply(fresh_fun_simp) +apply(simp) +(* none of them in BINDING *) +apply(simp) +apply(erule conjE) +apply(frule CAND_AndR_elim) +apply(assumption) +apply(erule exE) +apply(frule CAND_AndL1_elim) +apply(assumption) +apply(erule exE)+ +apply(simp only: ty.inject) +apply(drule_tac x="B1" in meta_spec) +apply(drule_tac x="M1" in meta_spec) +apply(drule_tac x="N'" in meta_spec) +apply(erule conjE)+ +apply(drule mp) +apply(simp) +apply(drule_tac x="b" in spec) +apply(rotate_tac 14) +apply(drule mp) +apply(assumption) +apply(rotate_tac 14) +apply(drule mp) +apply(simp add: CANDs_imply_SNa) +apply(drule_tac x="y" in spec) +apply(rotate_tac 15) +apply(drule mp) +apply(assumption) +apply(rotate_tac 15) +apply(drule mp) +apply(simp add: CANDs_imply_SNa) +apply(assumption) +(* LAnd2 case *) +apply(erule disjE) +apply(erule exE)+ +apply(auto)[1] +apply(frule_tac excluded_m) +apply(assumption) +apply(erule disjE) +(* one of them in BINDING *) +apply(erule disjE) +apply(drule fin_elims) +apply(drule fic_elims) +apply(simp) +apply(drule BINDINGc_elim) +apply(drule_tac x="x" in spec) +apply(drule_tac x="AndL2 (y).N' x" in spec) +apply(simp) +apply(simp add: better_AndR_substc) +apply(generate_fresh "coname") +apply(subgoal_tac "fresh_fun (\a'. Cut .AndR .M1 .M2 a' (x).AndL2 (y).N' x) + = Cut .AndR .M1 .M2 ca (x).AndL2 (y).N' x") +apply(simp) +apply(subgoal_tac "Cut .AndR .M1 .M2 ca (x).AndL2 (y).N' x \\<^isub>a Cut .M2 (y).N'") +apply(auto intro: a_preserves_SNa)[1] +apply(rule al_redu) +apply(rule better_LAnd2_intro) +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(simp) +apply(fresh_fun_simp) +apply(fresh_fun_simp) +apply(simp) +(* other case of in BINDING *) +apply(drule fin_elims) +apply(drule fic_elims) +apply(simp) +apply(drule BINDINGn_elim) +apply(drule_tac x="a" in spec) +apply(drule_tac x="AndR .M1 .M2 a" in spec) +apply(simp) +apply(simp add: better_AndL2_substn) +apply(generate_fresh "name") +apply(subgoal_tac "fresh_fun (\z'. Cut .AndR .M1 .M2 a (z').AndL2 (y).N' z') + = Cut .AndR .M1 .M2 a (ca).AndL2 (y).N' ca") +apply(simp) +apply(subgoal_tac "Cut .AndR .M1 .M2 a (ca).AndL2 (y).N' ca \\<^isub>a Cut .M2 (y).N'") +apply(auto intro: a_preserves_SNa)[1] +apply(rule al_redu) +apply(rule better_LAnd2_intro) +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(fresh_fun_simp) +apply(fresh_fun_simp) +apply(simp) +(* none of them in BINDING *) +apply(simp) +apply(erule conjE) +apply(frule CAND_AndR_elim) +apply(assumption) +apply(erule exE) +apply(frule CAND_AndL2_elim) +apply(assumption) +apply(erule exE)+ +apply(simp only: ty.inject) +apply(drule_tac x="B2" in meta_spec) +apply(drule_tac x="M2" in meta_spec) +apply(drule_tac x="N'" in meta_spec) +apply(erule conjE)+ +apply(drule mp) +apply(simp) +apply(drule_tac x="c" in spec) +apply(rotate_tac 14) +apply(drule mp) +apply(assumption) +apply(rotate_tac 14) +apply(drule mp) +apply(simp add: CANDs_imply_SNa) +apply(drule_tac x="y" in spec) +apply(rotate_tac 15) +apply(drule mp) +apply(assumption) +apply(rotate_tac 15) +apply(drule mp) +apply(simp add: CANDs_imply_SNa) +apply(assumption) +(* LOr1 case *) +apply(erule disjE) +apply(erule exE)+ +apply(auto)[1] +apply(frule_tac excluded_m) +apply(assumption) +apply(erule disjE) +(* one of them in BINDING *) +apply(erule disjE) +apply(drule fin_elims) +apply(drule fic_elims) +apply(simp) +apply(drule BINDINGc_elim) +apply(drule_tac x="x" in spec) +apply(drule_tac x="OrL (z).M1 (y).M2 x" in spec) +apply(simp) +apply(simp add: better_OrR1_substc) +apply(generate_fresh "coname") +apply(subgoal_tac "fresh_fun (\a'. Cut .OrR1 .N' a' (x).OrL (z).M1 (y).M2 x) + = Cut .OrR1 .N' c (x).OrL (z).M1 (y).M2 x") +apply(simp) +apply(subgoal_tac "Cut .OrR1 .N' c (x).OrL (z).M1 (y).M2 x \\<^isub>a Cut .N' (z).M1") +apply(auto intro: a_preserves_SNa)[1] +apply(rule al_redu) +apply(rule better_LOr1_intro) +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(simp add: abs_fresh) +apply(fresh_fun_simp) +apply(fresh_fun_simp) +apply(simp) +(* other case of in BINDING *) +apply(drule fin_elims) +apply(drule fic_elims) +apply(simp) +apply(drule BINDINGn_elim) +apply(drule_tac x="a" in spec) +apply(drule_tac x="OrR1 .N' a" in spec) +apply(simp) +apply(simp add: better_OrL_substn) +apply(generate_fresh "name") +apply(subgoal_tac "fresh_fun (\z'. Cut .OrR1 .N' a (z').OrL (z).M1 (y).M2 z') + = Cut .OrR1 .N' a (c).OrL (z).M1 (y).M2 c") +apply(simp) +apply(subgoal_tac "Cut .OrR1 .N' a (c).OrL (z).M1 (y).M2 c \\<^isub>a Cut .N' (z).M1") +apply(auto intro: a_preserves_SNa)[1] +apply(rule al_redu) +apply(rule better_LOr1_intro) +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(fresh_fun_simp) +apply(fresh_fun_simp) +apply(simp) +(* none of them in BINDING *) +apply(simp) +apply(erule conjE) +apply(frule CAND_OrR1_elim) +apply(assumption) +apply(erule exE)+ +apply(frule CAND_OrL_elim) +apply(assumption) +apply(erule exE)+ +apply(simp only: ty.inject) +apply(drule_tac x="B1" in meta_spec) +apply(drule_tac x="N'" in meta_spec) +apply(drule_tac x="M1" in meta_spec) +apply(erule conjE)+ +apply(drule mp) +apply(simp) +apply(drule_tac x="b" in spec) +apply(rotate_tac 15) +apply(drule mp) +apply(assumption) +apply(rotate_tac 15) +apply(drule mp) +apply(simp add: CANDs_imply_SNa) +apply(drule_tac x="z" in spec) +apply(rotate_tac 15) +apply(drule mp) +apply(assumption) +apply(rotate_tac 15) +apply(drule mp) +apply(simp add: CANDs_imply_SNa) +apply(assumption) +(* LOr2 case *) +apply(erule disjE) +apply(erule exE)+ +apply(auto)[1] +apply(frule_tac excluded_m) +apply(assumption) +apply(erule disjE) +(* one of them in BINDING *) +apply(erule disjE) +apply(drule fin_elims) +apply(drule fic_elims) +apply(simp) +apply(drule BINDINGc_elim) +apply(drule_tac x="x" in spec) +apply(drule_tac x="OrL (z).M1 (y).M2 x" in spec) +apply(simp) +apply(simp add: better_OrR2_substc) +apply(generate_fresh "coname") +apply(subgoal_tac "fresh_fun (\a'. Cut .OrR2 .N' a' (x).OrL (z).M1 (y).M2 x) + = Cut .OrR2 .N' c (x).OrL (z).M1 (y).M2 x") +apply(simp) +apply(subgoal_tac "Cut .OrR2 .N' c (x).OrL (z).M1 (y).M2 x \\<^isub>a Cut .N' (y).M2") +apply(auto intro: a_preserves_SNa)[1] +apply(rule al_redu) +apply(rule better_LOr2_intro) +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(simp add: abs_fresh) +apply(fresh_fun_simp) +apply(fresh_fun_simp) +apply(simp) +(* other case of in BINDING *) +apply(drule fin_elims) +apply(drule fic_elims) +apply(simp) +apply(drule BINDINGn_elim) +apply(drule_tac x="a" in spec) +apply(drule_tac x="OrR2 .N' a" in spec) +apply(simp) +apply(simp add: better_OrL_substn) +apply(generate_fresh "name") +apply(subgoal_tac "fresh_fun (\z'. Cut .OrR2 .N' a (z').OrL (z).M1 (y).M2 z') + = Cut .OrR2 .N' a (c).OrL (z).M1 (y).M2 c") +apply(simp) +apply(subgoal_tac "Cut .OrR2 .N' a (c).OrL (z).M1 (y).M2 c \\<^isub>a Cut .N' (y).M2") +apply(auto intro: a_preserves_SNa)[1] +apply(rule al_redu) +apply(rule better_LOr2_intro) +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(fresh_fun_simp) +apply(fresh_fun_simp) +apply(simp) +(* none of them in BINDING *) +apply(simp) +apply(erule conjE) +apply(frule CAND_OrR2_elim) +apply(assumption) +apply(erule exE)+ +apply(frule CAND_OrL_elim) +apply(assumption) +apply(erule exE)+ +apply(simp only: ty.inject) +apply(drule_tac x="B2" in meta_spec) +apply(drule_tac x="N'" in meta_spec) +apply(drule_tac x="M2" in meta_spec) +apply(erule conjE)+ +apply(drule mp) +apply(simp) +apply(drule_tac x="b" in spec) +apply(rotate_tac 15) +apply(drule mp) +apply(assumption) +apply(rotate_tac 15) +apply(drule mp) +apply(simp add: CANDs_imply_SNa) +apply(drule_tac x="y" in spec) +apply(rotate_tac 15) +apply(drule mp) +apply(assumption) +apply(rotate_tac 15) +apply(drule mp) +apply(simp add: CANDs_imply_SNa) +apply(assumption) +(* LImp case *) +apply(erule exE)+ +apply(auto)[1] +apply(frule_tac excluded_m) +apply(assumption) +apply(erule disjE) +(* one of them in BINDING *) +apply(erule disjE) +apply(drule fin_elims) +apply(drule fic_elims) +apply(simp) +apply(drule BINDINGc_elim) +apply(drule_tac x="x" in spec) +apply(drule_tac x="ImpL .N1 (y).N2 x" in spec) +apply(simp) +apply(simp add: better_ImpR_substc) +apply(generate_fresh "coname") +apply(subgoal_tac "fresh_fun (\a'. Cut .ImpR (z)..M'a a' (x).ImpL .N1 (y).N2 x) + = Cut .ImpR (z)..M'a ca (x).ImpL .N1 (y).N2 x") +apply(simp) +apply(subgoal_tac "Cut .ImpR (z)..M'a ca (x).ImpL .N1 (y).N2 x \\<^isub>a + Cut .Cut .N1 (z).M'a (y).N2") +apply(auto intro: a_preserves_SNa)[1] +apply(rule al_redu) +apply(rule better_LImp_intro) +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(simp add: abs_fresh) +apply(simp) +apply(fresh_fun_simp) +apply(fresh_fun_simp) +apply(simp) +(* other case of in BINDING *) +apply(drule fin_elims) +apply(drule fic_elims) +apply(simp) +apply(drule BINDINGn_elim) +apply(drule_tac x="a" in spec) +apply(drule_tac x="ImpR (z)..M'a a" in spec) +apply(simp) +apply(simp add: better_ImpL_substn) +apply(generate_fresh "name") +apply(subgoal_tac "fresh_fun (\z'. Cut .ImpR (z)..M'a a (z').ImpL .N1 (y).N2 z') + = Cut .ImpR (z)..M'a a (ca).ImpL .N1 (y).N2 ca") +apply(simp) +apply(subgoal_tac "Cut .ImpR (z)..M'a a (ca).ImpL .N1 (y).N2 ca \\<^isub>a + Cut .Cut .N1 (z).M'a (y).N2") +apply(auto intro: a_preserves_SNa)[1] +apply(rule al_redu) +apply(rule better_LImp_intro) +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(simp) +apply(fresh_fun_simp) +apply(simp add: abs_fresh abs_supp fin_supp) +apply(fresh_fun_simp) +apply(simp add: abs_fresh abs_supp fin_supp) +apply(simp) +(* none of them in BINDING *) +apply(simp) +apply(erule conjE) +apply(frule CAND_ImpL_elim) +apply(assumption) +apply(erule exE)+ +apply(frule CAND_ImpR_elim) (* check here *) +apply(assumption) +apply(erule exE)+ +apply(erule conjE)+ +apply(simp only: ty.inject) +apply(erule conjE)+ +apply(case_tac "M'a=Ax z b") +(* case Ma = Ax z b *) +apply(rule_tac t="Cut .(Cut .N1 (z).M'a) (y).N2" and s="Cut .(M'a{z:=.N1}) (y).N2" in subst) +apply(simp) +apply(drule_tac x="c" in spec) +apply(drule_tac x="N1" in spec) +apply(drule mp) +apply(simp) +apply(drule_tac x="B2" in meta_spec) +apply(drule_tac x="M'a{z:=.N1}" in meta_spec) +apply(drule_tac x="N2" in meta_spec) +apply(drule conjunct1) +apply(drule mp) +apply(simp) +apply(rotate_tac 17) +apply(drule_tac x="b" in spec) +apply(drule mp) +apply(assumption) +apply(drule mp) +apply(simp add: CANDs_imply_SNa) +apply(rotate_tac 17) +apply(drule_tac x="y" in spec) +apply(drule mp) +apply(assumption) +apply(drule mp) +apply(simp add: CANDs_imply_SNa) +apply(assumption) +(* case Ma \ Ax z b *) +apply(subgoal_tac ":Cut .N1 (z).M'a \ \\") (* lemma *) +apply(frule_tac meta_spec) +apply(drule_tac x="B2" in meta_spec) +apply(drule_tac x="Cut .N1 (z).M'a" in meta_spec) +apply(drule_tac x="N2" in meta_spec) +apply(erule conjE)+ +apply(drule mp) +apply(simp) +apply(rotate_tac 20) +apply(drule_tac x="b" in spec) +apply(rotate_tac 20) +apply(drule mp) +apply(assumption) +apply(rotate_tac 20) +apply(drule mp) +apply(simp add: CANDs_imply_SNa) +apply(rotate_tac 20) +apply(drule_tac x="y" in spec) +apply(rotate_tac 20) +apply(drule mp) +apply(assumption) +apply(rotate_tac 20) +apply(drule mp) +apply(simp add: CANDs_imply_SNa) +apply(assumption) +(* lemma *) +apply(subgoal_tac ":Cut .N1 (z).M'a \ BINDINGc B2 (\(B2)\)") (* second lemma *) +apply(simp add: BINDING_implies_CAND) +(* second lemma *) +apply(simp (no_asm) add: BINDINGc_def) +apply(rule exI)+ +apply(rule conjI) +apply(rule refl) +apply(rule allI)+ +apply(rule impI) +apply(generate_fresh "name") +apply(rule_tac t="Cut .N1 (z).M'a" and s="Cut .N1 (ca).([(ca,z)]\M'a)" in subst) +apply(simp add: trm.inject alpha fresh_prod fresh_atm) +apply(rule_tac t="(Cut .N1 (ca).([(ca,z)]\M'a)){b:=(xa).P}" + and s="Cut .N1 (ca).(([(ca,z)]\M'a){b:=(xa).P})" in subst) +apply(rule sym) +apply(rule tricky_subst) +apply(simp) +apply(simp) +apply(clarify) +apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm) +apply(drule_tac x="B1" in meta_spec) +apply(drule_tac x="N1" in meta_spec) +apply(drule_tac x="([(ca,z)]\M'a){b:=(xa).P}" in meta_spec) +apply(drule conjunct1) +apply(drule mp) +apply(simp) +apply(rotate_tac 19) +apply(drule_tac x="c" in spec) +apply(drule mp) +apply(assumption) +apply(drule mp) +apply(simp add: CANDs_imply_SNa) +apply(rotate_tac 19) +apply(drule_tac x="ca" in spec) +apply(subgoal_tac "(ca):([(ca,z)]\M'a){b:=(xa).P} \ \(B1)\")(*A*) +apply(drule mp) +apply(assumption) +apply(drule mp) +apply(simp add: CANDs_imply_SNa) +apply(assumption) +(*A*) +apply(drule_tac x="[(ca,z)]\xa" in spec) +apply(drule_tac x="[(ca,z)]\P" in spec) +apply(rotate_tac 19) +apply(simp add: fresh_prod fresh_left) +apply(drule mp) +apply(rule conjI) +apply(auto simp add: calc_atm)[1] +apply(rule conjI) +apply(auto simp add: calc_atm)[1] +apply(drule_tac pi="[(ca,z)]" and x="(xa):P" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name) +apply(drule_tac pi="[(ca,z)]" and X="\(B1)\" in pt_set_bij2[OF pt_name_inst, OF at_name_inst]) +apply(simp add: CAND_eqvt_name csubst_eqvt) +apply(perm_simp) +apply(simp add: calc_atm) +done + + +(* HERE *) + + +(* parallel substitution *) + + +lemma CUT_SNa: + assumes a1: ":M \ (\\)" + and a2: "(x):N \ (\(B)\)" + shows "SNa (Cut .M (x).N)" +using a1 a2 +apply - +apply(rule CUT_SNa_aux[OF a1]) +apply(simp_all add: CANDs_imply_SNa) +done + + +fun + findn :: "(name\coname\trm) list\name\(coname\trm) option" +where + "findn [] x = None" +| "findn ((y,c,P)#\n) x = (if y=x then Some (c,P) else findn \n x)" + +lemma findn_eqvt[eqvt]: + fixes pi1::"name prm" + and pi2::"coname prm" + shows "(pi1\findn \n x) = findn (pi1\\n) (pi1\x)" + and "(pi2\findn \n x) = findn (pi2\\n) (pi2\x)" +apply(induct \n) +apply(auto simp add: perm_bij) +done + +lemma findn_fresh: + assumes a: "x\\n" + shows "findn \n x = None" +using a +apply(induct \n) +apply(auto simp add: fresh_list_cons fresh_atm fresh_prod) +done + +fun + findc :: "(coname\name\trm) list\coname\(name\trm) option" +where + "findc [] x = None" +| "findc ((c,y,P)#\c) a = (if a=c then Some (y,P) else findc \c a)" + +lemma findc_eqvt[eqvt]: + fixes pi1::"name prm" + and pi2::"coname prm" + shows "(pi1\findc \c a) = findc (pi1\\c) (pi1\a)" + and "(pi2\findc \c a) = findc (pi2\\c) (pi2\a)" +apply(induct \c) +apply(auto simp add: perm_bij) +done + +lemma findc_fresh: + assumes a: "a\\c" + shows "findc \c a = None" +using a +apply(induct \c) +apply(auto simp add: fresh_list_cons fresh_atm fresh_prod) +done + +abbreviation + nmaps :: "(name\coname\trm) list \ name \ (coname\trm) option \ bool" ("_ nmaps _ to _" [55,55,55] 55) +where + "\n nmaps x to P \ (findn \n x) = P" + +abbreviation + cmaps :: "(coname\name\trm) list \ coname \ (name\trm) option \ bool" ("_ cmaps _ to _" [55,55,55] 55) +where + "\c cmaps a to P \ (findc \c a) = P" + +lemma nmaps_fresh: + shows "\n nmaps x to Some (c,P) \ a\\n \ a\(c,P)" +apply(induct \n) +apply(auto simp add: fresh_list_cons fresh_prod fresh_atm) +apply(case_tac "aa=x") +apply(auto) +apply(case_tac "aa=x") +apply(auto) +done + +lemma cmaps_fresh: + shows "\c cmaps a to Some (y,P) \ x\\c \ x\(y,P)" +apply(induct \c) +apply(auto simp add: fresh_list_cons fresh_prod fresh_atm) +apply(case_tac "a=aa") +apply(auto) +apply(case_tac "a=aa") +apply(auto) +done + +lemma nmaps_false: + shows "\n nmaps x to Some (c,P) \ x\\n \ False" +apply(induct \n) +apply(auto simp add: fresh_list_cons fresh_prod fresh_atm) +done + +lemma cmaps_false: + shows "\c cmaps c to Some (x,P) \ c\\c \ False" +apply(induct \c) +apply(auto simp add: fresh_list_cons fresh_prod fresh_atm) +done + +fun + lookupa :: "name\coname\(coname\name\trm) list\trm" +where + "lookupa x a [] = Ax x a" +| "lookupa x a ((c,y,P)#\c) = (if a=c then Cut .Ax x a (y).P else lookupa x a \c)" + +lemma lookupa_eqvt[eqvt]: + fixes pi1::"name prm" + and pi2::"coname prm" + shows "(pi1\(lookupa x a \c)) = lookupa (pi1\x) (pi1\a) (pi1\\c)" + and "(pi2\(lookupa x a \c)) = lookupa (pi2\x) (pi2\a) (pi2\\c)" +apply - +apply(induct \c) +apply(auto simp add: eqvts) +apply(induct \c) +apply(auto simp add: eqvts) +done + +lemma lookupa_fire: + assumes a: "\c cmaps a to Some (y,P)" + shows "(lookupa x a \c) = Cut .Ax x a (y).P" +using a +apply(induct \c arbitrary: x a y P) +apply(auto) +done + +fun + lookupb :: "name\coname\(coname\name\trm) list\coname\trm\trm" +where + "lookupb x a [] c P = Cut .P (x).Ax x a" +| "lookupb x a ((d,y,N)#\c) c P = (if a=d then Cut .P (y).N else lookupb x a \c c P)" + +lemma lookupb_eqvt[eqvt]: + fixes pi1::"name prm" + and pi2::"coname prm" + shows "(pi1\(lookupb x a \c c P)) = lookupb (pi1\x) (pi1\a) (pi1\\c) (pi1\c) (pi1\P)" + and "(pi2\(lookupb x a \c c P)) = lookupb (pi2\x) (pi2\a) (pi2\\c) (pi2\c) (pi2\P)" +apply - +apply(induct \c) +apply(auto simp add: eqvts) +apply(induct \c) +apply(auto simp add: eqvts) +done + +fun + lookup :: "name\coname\(name\coname\trm) list\(coname\name\trm) list\trm" +where + "lookup x a [] \c = lookupa x a \c" +| "lookup x a ((y,c,P)#\n) \c = (if x=y then (lookupb x a \c c P) else lookup x a \n \c)" + +lemma lookup_eqvt[eqvt]: + fixes pi1::"name prm" + and pi2::"coname prm" + shows "(pi1\(lookup x a \n \c)) = lookup (pi1\x) (pi1\a) (pi1\\n) (pi1\\c)" + and "(pi2\(lookup x a \n \c)) = lookup (pi2\x) (pi2\a) (pi2\\n) (pi2\\c)" +apply - +apply(induct \n) +apply(auto simp add: eqvts) +apply(induct \n) +apply(auto simp add: eqvts) +done + +fun + lookupc :: "name\coname\(name\coname\trm) list\trm" +where + "lookupc x a [] = Ax x a" +| "lookupc x a ((y,c,P)#\n) = (if x=y then P[c\c>a] else lookupc x a \n)" + +lemma lookupc_eqvt[eqvt]: + fixes pi1::"name prm" + and pi2::"coname prm" + shows "(pi1\(lookupc x a \n)) = lookupc (pi1\x) (pi1\a) (pi1\\n)" + and "(pi2\(lookupc x a \n)) = lookupc (pi2\x) (pi2\a) (pi2\\n)" +apply - +apply(induct \n) +apply(auto simp add: eqvts) +apply(induct \n) +apply(auto simp add: eqvts) +done + +fun + lookupd :: "name\coname\(coname\name\trm) list\trm" +where + "lookupd x a [] = Ax x a" +| "lookupd x a ((c,y,P)#\c) = (if a=c then P[y\n>x] else lookupd x a \c)" + +lemma lookupd_eqvt[eqvt]: + fixes pi1::"name prm" + and pi2::"coname prm" + shows "(pi1\(lookupd x a \n)) = lookupd (pi1\x) (pi1\a) (pi1\\n)" + and "(pi2\(lookupd x a \n)) = lookupd (pi2\x) (pi2\a) (pi2\\n)" +apply - +apply(induct \n) +apply(auto simp add: eqvts) +apply(induct \n) +apply(auto simp add: eqvts) +done + +lemma lookupa_fresh: + assumes a: "a\\c" + shows "lookupa y a \c = Ax y a" +using a +apply(induct \c) +apply(auto simp add: fresh_prod fresh_list_cons fresh_atm) +done + +lemma lookupa_csubst: + assumes a: "a\\c" + shows "Cut .Ax y a (x).P = (lookupa y a \c){a:=(x).P}" +using a by (simp add: lookupa_fresh) + +lemma lookupa_freshness: + fixes a::"coname" + and x::"name" + shows "a\(\c,c) \ a\lookupa y c \c" + and "x\(\c,y) \ x\lookupa y c \c" +apply(induct \c) +apply(auto simp add: fresh_prod fresh_list_cons abs_fresh fresh_atm) +done + +lemma lookupa_unicity: + assumes a: "lookupa x a \c= Ax y b" "b\\c" "y\\c" + shows "x=y \ a=b" +using a +apply(induct \c) +apply(auto simp add: trm.inject fresh_list_cons fresh_prod fresh_atm) +apply(case_tac "a=aa") +apply(auto) +apply(case_tac "a=aa") +apply(auto) +done + +lemma lookupb_csubst: + assumes a: "a\(\c,c,N)" + shows "Cut .N (x).P = (lookupb y a \c c N){a:=(x).P}" +using a +apply(induct \c) +apply(auto simp add: fresh_list_cons fresh_atm fresh_prod) +apply(rule sym) +apply(generate_fresh "name") +apply(generate_fresh "coname") +apply(subgoal_tac "Cut .N (y).Ax y a = Cut .([(caa,c)]\N) (ca).Ax ca a") +apply(simp) +apply(rule trans) +apply(rule better_Cut_substc) +apply(simp) +apply(simp add: abs_fresh) +apply(simp) +apply(subgoal_tac "a\([(caa,c)]\N)") +apply(simp add: forget) +apply(simp add: trm.inject) +apply(simp add: fresh_left calc_atm fresh_prod fresh_atm) +apply(simp add: trm.inject) +apply(rule conjI) +apply(rule sym) +apply(simp add: alpha fresh_prod fresh_atm) +apply(simp add: alpha calc_atm fresh_prod fresh_atm) +done + +lemma lookupb_freshness: + fixes a::"coname" + and x::"name" + shows "a\(\c,c,b,P) \ a\lookupb y c \c b P" + and "x\(\c,y,P) \ x\lookupb y c \c b P" +apply(induct \c) +apply(auto simp add: fresh_prod fresh_list_cons abs_fresh fresh_atm) +done + +lemma lookupb_unicity: + assumes a: "lookupb x a \c c P = Ax y b" "b\(\c,c,P)" "y\\c" + shows "x=y \ a=b" +using a +apply(induct \c) +apply(auto simp add: fresh_list_cons fresh_prod fresh_atm) +apply(case_tac "a=aa") +apply(auto) +apply(case_tac "a=aa") +apply(auto) +done + +lemma lookupb_lookupa: + assumes a: "x\\c" + shows "lookupb x c \c a P = (lookupa x c \c){x:=.P}" +using a +apply(induct \c) +apply(auto simp add: fresh_list_cons fresh_prod) +apply(generate_fresh "coname") +apply(generate_fresh "name") +apply(subgoal_tac "Cut .Ax x c (aa).b = Cut .Ax x ca (caa).([(caa,aa)]\b)") +apply(simp) +apply(rule sym) +apply(rule trans) +apply(rule better_Cut_substn) +apply(simp add: abs_fresh) +apply(simp) +apply(simp) +apply(subgoal_tac "x\([(caa,aa)]\b)") +apply(simp add: forget) +apply(simp add: trm.inject) +apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(simp add: trm.inject) +apply(rule conjI) +apply(simp add: alpha calc_atm fresh_atm fresh_prod) +apply(rule sym) +apply(simp add: alpha calc_atm fresh_atm fresh_prod) +done + +lemma lookup_csubst: + assumes a: "a\(\n,\c)" + shows "lookup y c \n ((a,x,P)#\c) = (lookup y c \n \c){a:=(x).P}" +using a +apply(induct \n) +apply(auto simp add: fresh_prod fresh_list_cons) +apply(simp add: lookupa_csubst) +apply(simp add: lookupa_freshness forget fresh_atm fresh_prod) +apply(rule lookupb_csubst) +apply(simp) +apply(auto simp add: lookupb_freshness forget fresh_atm fresh_prod) +done + +lemma lookup_fresh: + assumes a: "x\(\n,\c)" + shows "lookup x c \n \c = lookupa x c \c" +using a +apply(induct \n) +apply(auto simp add: fresh_prod fresh_list_cons fresh_atm) +done + +lemma lookup_unicity: + assumes a: "lookup x a \n \c= Ax y b" "b\(\c,\n)" "y\(\c,\n)" + shows "x=y \ a=b" +using a +apply(induct \n) +apply(auto simp add: trm.inject fresh_list_cons fresh_prod fresh_atm) +apply(drule lookupa_unicity) +apply(simp)+ +apply(drule lookupa_unicity) +apply(simp)+ +apply(case_tac "x=aa") +apply(auto) +apply(drule lookupb_unicity) +apply(simp add: fresh_atm) +apply(simp) +apply(simp) +apply(case_tac "x=aa") +apply(auto) +apply(drule lookupb_unicity) +apply(simp add: fresh_atm) +apply(simp) +apply(simp) +done + +lemma lookup_freshness: + fixes a::"coname" + and x::"name" + shows "a\(c,\c,\n) \ a\lookup y c \n \c" + and "x\(y,\c,\n) \ x\lookup y c \n \c" +apply(induct \n) +apply(auto simp add: fresh_prod fresh_list_cons abs_fresh fresh_atm) +apply(simp add: fresh_atm fresh_prod lookupa_freshness) +apply(simp add: fresh_atm fresh_prod lookupa_freshness) +apply(simp add: fresh_atm fresh_prod lookupb_freshness) +apply(simp add: fresh_atm fresh_prod lookupb_freshness) +done + +lemma lookupc_freshness: + fixes a::"coname" + and x::"name" + shows "a\(\c,c) \ a\lookupc y c \c" + and "x\(\c,y) \ x\lookupc y c \c" +apply(induct \c) +apply(auto simp add: fresh_prod fresh_list_cons abs_fresh fresh_atm) +apply(rule rename_fresh) +apply(simp add: fresh_atm) +apply(rule rename_fresh) +apply(simp add: fresh_atm) +done + +lemma lookupc_fresh: + assumes a: "y\\n" + shows "lookupc y a \n = Ax y a" +using a +apply(induct \n) +apply(auto simp add: fresh_prod fresh_list_cons fresh_atm) +done + +lemma lookupc_nmaps: + assumes a: "\n nmaps x to Some (c,P)" + shows "lookupc x a \n = P[c\c>a]" +using a +apply(induct \n) +apply(auto) +done + +lemma lookupc_unicity: + assumes a: "lookupc y a \n = Ax x b" "x\\n" + shows "y=x" +using a +apply(induct \n) +apply(auto simp add: trm.inject fresh_list_cons fresh_prod) +apply(case_tac "y=aa") +apply(auto) +apply(subgoal_tac "x\(ba[aa\c>a])") +apply(simp add: fresh_atm) +apply(rule rename_fresh) +apply(simp) +done + +lemma lookupd_fresh: + assumes a: "a\\c" + shows "lookupd y a \c = Ax y a" +using a +apply(induct \c) +apply(auto simp add: fresh_prod fresh_list_cons fresh_atm) +done + +lemma lookupd_unicity: + assumes a: "lookupd y a \c = Ax y b" "b\\c" + shows "a=b" +using a +apply(induct \c) +apply(auto simp add: trm.inject fresh_list_cons fresh_prod) +apply(case_tac "a=aa") +apply(auto) +apply(subgoal_tac "b\(ba[aa\n>y])") +apply(simp add: fresh_atm) +apply(rule rename_fresh) +apply(simp) +done + +lemma lookupd_freshness: + fixes a::"coname" + and x::"name" + shows "a\(\c,c) \ a\lookupd y c \c" + and "x\(\c,y) \ x\lookupd y c \c" +apply(induct \c) +apply(auto simp add: fresh_prod fresh_list_cons abs_fresh fresh_atm) +apply(rule rename_fresh) +apply(simp add: fresh_atm) +apply(rule rename_fresh) +apply(simp add: fresh_atm) +done + +lemma lookupd_cmaps: + assumes a: "\c cmaps a to Some (x,P)" + shows "lookupd y a \c = P[x\n>y]" +using a +apply(induct \c) +apply(auto) +done + +consts + stn :: "trm\(name\coname\trm) list\trm" + stc :: "trm\(coname\name\trm) list\trm" + +nominal_primrec (freshness_context: "\n::(name\coname\trm)") + "stn (Ax x a) \n = lookupc x a \n" + "\a\(N,\n);x\(M,\n)\ \ stn (Cut .M (x).N) \n = (Cut .M (x).N)" + "x\\n \ stn (NotR (x).M a) \n = (NotR (x).M a)" + "a\\n \stn (NotL .M x) \n = (NotL .M x)" + "\a\(N,d,b,\n);b\(M,d,a,\n)\ \ stn (AndR .M .N d) \n = (AndR .M .N d)" + "x\(z,\n) \ stn (AndL1 (x).M z) \n = (AndL1 (x).M z)" + "x\(z,\n) \ stn (AndL2 (x).M z) \n = (AndL2 (x).M z)" + "a\(b,\n) \ stn (OrR1 .M b) \n = (OrR1 .M b)" + "a\(b,\n) \ stn (OrR2 .M b) \n = (OrR2 .M b)" + "\x\(N,z,u,\n);u\(M,z,x,\n)\ \ stn (OrL (x).M (u).N z) \n = (OrL (x).M (u).N z)" + "\a\(b,\n);x\\n\ \ stn (ImpR (x)..M b) \n = (ImpR (x)..M b)" + "\a\(N,\n);x\(M,z,\n)\ \ stn (ImpL .M (x).N z) \n = (ImpL .M (x).N z)" +apply(finite_guess)+ +apply(rule TrueI)+ +apply(simp add: abs_fresh abs_supp fin_supp)+ +apply(fresh_guess)+ +done + +nominal_primrec (freshness_context: "\c::(coname\name\trm)") + "stc (Ax x a) \c = lookupd x a \c" + "\a\(N,\c);x\(M,\c)\ \ stc (Cut .M (x).N) \c = (Cut .M (x).N)" + "x\\c \ stc (NotR (x).M a) \c = (NotR (x).M a)" + "a\\c \ stc (NotL .M x) \c = (NotL .M x)" + "\a\(N,d,b,\c);b\(M,d,a,\c)\ \ stc (AndR .M .N d) \c = (AndR .M .N d)" + "x\(z,\c) \ stc (AndL1 (x).M z) \c = (AndL1 (x).M z)" + "x\(z,\c) \ stc (AndL2 (x).M z) \c = (AndL2 (x).M z)" + "a\(b,\c) \ stc (OrR1 .M b) \c = (OrR1 .M b)" + "a\(b,\c) \ stc (OrR2 .M b) \c = (OrR2 .M b)" + "\x\(N,z,u,\c);u\(M,z,x,\c)\ \ stc (OrL (x).M (u).N z) \c = (OrL (x).M (u).N z)" + "\a\(b,\c);x\\c\ \ stc (ImpR (x)..M b) \c = (ImpR (x)..M b)" + "\a\(N,\c);x\(M,z,\c)\ \ stc (ImpL .M (x).N z) \c = (ImpL .M (x).N z)" +apply(finite_guess)+ +apply(rule TrueI)+ +apply(simp add: abs_fresh abs_supp fin_supp)+ +apply(fresh_guess)+ +done + +lemma stn_eqvt[eqvt]: + fixes pi1::"name prm" + and pi2::"coname prm" + shows "(pi1\(stn M \n)) = stn (pi1\M) (pi1\\n)" + and "(pi2\(stn M \n)) = stn (pi2\M) (pi2\\n)" +apply - +apply(nominal_induct M avoiding: \n rule: trm.induct) +apply(auto simp add: eqvts fresh_bij fresh_prod eq_bij fresh_atm) +apply(nominal_induct M avoiding: \n rule: trm.induct) +apply(auto simp add: eqvts fresh_bij fresh_prod eq_bij fresh_atm) +done + +lemma stc_eqvt[eqvt]: + fixes pi1::"name prm" + and pi2::"coname prm" + shows "(pi1\(stc M \c)) = stc (pi1\M) (pi1\\c)" + and "(pi2\(stc M \c)) = stc (pi2\M) (pi2\\c)" +apply - +apply(nominal_induct M avoiding: \c rule: trm.induct) +apply(auto simp add: eqvts fresh_bij fresh_prod eq_bij fresh_atm) +apply(nominal_induct M avoiding: \c rule: trm.induct) +apply(auto simp add: eqvts fresh_bij fresh_prod eq_bij fresh_atm) +done + +lemma stn_fresh: + fixes a::"coname" + and x::"name" + shows "a\(\n,M) \ a\stn M \n" + and "x\(\n,M) \ x\stn M \n" +apply(nominal_induct M avoiding: \n a x rule: trm.induct) +apply(auto simp add: abs_fresh fresh_prod fresh_atm) +apply(rule lookupc_freshness) +apply(simp add: fresh_atm) +apply(rule lookupc_freshness) +apply(simp add: fresh_atm) +done + +lemma stc_fresh: + fixes a::"coname" + and x::"name" + shows "a\(\c,M) \ a\stc M \c" + and "x\(\c,M) \ x\stc M \c" +apply(nominal_induct M avoiding: \c a x rule: trm.induct) +apply(auto simp add: abs_fresh fresh_prod fresh_atm) +apply(rule lookupd_freshness) +apply(simp add: fresh_atm) +apply(rule lookupd_freshness) +apply(simp add: fresh_atm) +done + +lemma option_case_eqvt1[eqvt_force]: + fixes pi1::"name prm" + and pi2::"coname prm" + and B::"(name\trm) option" + and r::"trm" + shows "(pi1\(case B of Some (x,P) \ s x P | None \ r)) = + (case (pi1\B) of Some (x,P) \ (pi1\s) x P | None \ (pi1\r))" + and "(pi2\(case B of Some (x,P) \ s x P| None \ r)) = + (case (pi2\B) of Some (x,P) \ (pi2\s) x P | None \ (pi2\r))" +apply(cases "B") +apply(auto) +apply(perm_simp) +apply(cases "B") +apply(auto) +apply(perm_simp) +done + +lemma option_case_eqvt2[eqvt_force]: + fixes pi1::"name prm" + and pi2::"coname prm" + and B::"(coname\trm) option" + and r::"trm" + shows "(pi1\(case B of Some (x,P) \ s x P | None \ r)) = + (case (pi1\B) of Some (x,P) \ (pi1\s) x P | None \ (pi1\r))" + and "(pi2\(case B of Some (x,P) \ s x P| None \ r)) = + (case (pi2\B) of Some (x,P) \ (pi2\s) x P | None \ (pi2\r))" +apply(cases "B") +apply(auto) +apply(perm_simp) +apply(cases "B") +apply(auto) +apply(perm_simp) +done + +consts + psubst :: "(name\coname\trm) list\(coname\name\trm) list\trm\trm" ("_,_<_>" [100,100,100] 100) + +nominal_primrec (freshness_context: "(\n::(name\coname\trm) list,\c::(coname\name\trm) list)") + "\n,\c = lookup x a \n \c" + "\a\(N,\n,\c);x\(M,\n,\c)\ \ \n,\c.M (x).N> = + Cut .(if \x. M=Ax x a then stn M \n else \n,\c) + (x).(if \a. N=Ax x a then stc N \c else \n,\c)" + "x\(\n,\c) \ \n,\c = + (case (findc \c a) of + Some (u,P) \ fresh_fun (\a'. Cut .NotR (x).(\n,\c) a' (u).P) + | None \ NotR (x).(\n,\c) a)" + "a\(\n,\c) \ \n,\c.M x> = + (case (findn \n x) of + Some (c,P) \ fresh_fun (\x'. Cut .P (x').(NotL .(\n,\c) x')) + | None \ NotL .(\n,\c) x)" + "\a\(N,c,\n,\c);b\(M,c,\n,\c);b\a\ \ (\n,\c.M .N c>) = + (case (findc \c c) of + Some (x,P) \ fresh_fun (\a'. Cut .(AndR .(\n,\c) .(\n,\c) a') (x).P) + | None \ AndR .(\n,\c) .(\n,\c) c)" + "x\(z,\n,\c) \ (\n,\c) = + (case (findn \n z) of + Some (c,P) \ fresh_fun (\z'. Cut .P (z').AndL1 (x).(\n,\c) z') + | None \ AndL1 (x).(\n,\c) z)" + "x\(z,\n,\c) \ (\n,\c) = + (case (findn \n z) of + Some (c,P) \ fresh_fun (\z'. Cut .P (z').AndL2 (x).(\n,\c) z') + | None \ AndL2 (x).(\n,\c) z)" + "\x\(N,z,\n,\c);u\(M,z,\n,\c);x\u\ \ (\n,\c) = + (case (findn \n z) of + Some (c,P) \ fresh_fun (\z'. Cut .P (z').OrL (x).(\n,\c) (u).(\n,\c) z') + | None \ OrL (x).(\n,\c) (u).(\n,\c) z)" + "a\(b,\n,\c) \ (\n,\c.M b>) = + (case (findc \c b) of + Some (x,P) \ fresh_fun (\a'. Cut .OrR1 .(\n,\c) a' (x).P) + | None \ OrR1 .(\n,\c) b)" + "a\(b,\n,\c) \ (\n,\c.M b>) = + (case (findc \c b) of + Some (x,P) \ fresh_fun (\a'. Cut .OrR2 .(\n,\c) a' (x).P) + | None \ OrR2 .(\n,\c) b)" + "\a\(b,\n,\c); x\(\n,\c)\ \ (\n,\c.M b>) = + (case (findc \c b) of + Some (z,P) \ fresh_fun (\a'. Cut .ImpR (x)..(\n,\c) a' (z).P) + | None \ ImpR (x)..(\n,\c) b)" + "\a\(N,\n,\c); x\(z,M,\n,\c)\ \ (\n,\c.M (x).N z>) = + (case (findn \n z) of + Some (c,P) \ fresh_fun (\z'. Cut .P (z').ImpL .(\n,\c) (x).(\n,\c) z') + | None \ ImpL .(\n,\c) (x).(\n,\c) z)" +apply(finite_guess)+ +apply(rule TrueI)+ +apply(simp add: abs_fresh stc_fresh) +apply(simp add: abs_fresh stn_fresh) +apply(case_tac "findc \c x3") +apply(simp add: abs_fresh) +apply(auto)[1] +apply(generate_fresh "coname") +apply(thin_tac "x1\?p") +apply(fresh_fun_simp) +apply(drule cmaps_fresh) +apply(auto simp add: fresh_prod)[1] +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(case_tac "findn \n x3") +apply(simp add: abs_fresh) +apply(auto)[1] +apply(generate_fresh "name") +apply(thin_tac "x1\?p") +apply(fresh_fun_simp) +apply(drule nmaps_fresh) +apply(auto simp add: fresh_prod)[1] +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(case_tac "findc \c x5") +apply(simp add: abs_fresh) +apply(auto)[1] +apply(generate_fresh "coname") +apply(thin_tac "x1\?p") +apply(thin_tac "x3\?p") +apply(fresh_fun_simp) +apply(drule cmaps_fresh) +apply(auto simp add: fresh_prod)[1] +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(case_tac "findc \c x5") +apply(simp add: abs_fresh) +apply(auto)[1] +apply(generate_fresh "coname") +apply(thin_tac "x1\?p") +apply(thin_tac "x3\?p") +apply(fresh_fun_simp) +apply(drule_tac x="x3" in cmaps_fresh) +apply(auto simp add: fresh_prod)[1] +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(case_tac "findn \n x3") +apply(simp add: abs_fresh) +apply(auto)[1] +apply(generate_fresh "name") +apply(thin_tac "x1\?p") +apply(fresh_fun_simp) +apply(drule nmaps_fresh) +apply(auto simp add: fresh_prod)[1] +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(case_tac "findn \n x3") +apply(simp add: abs_fresh) +apply(auto)[1] +apply(generate_fresh "name") +apply(thin_tac "x1\?p") +apply(fresh_fun_simp) +apply(drule nmaps_fresh) +apply(auto simp add: fresh_prod)[1] +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(case_tac "findc \c x3") +apply(simp add: abs_fresh) +apply(auto)[1] +apply(generate_fresh "coname") +apply(thin_tac "x1\?p") +apply(fresh_fun_simp) +apply(drule cmaps_fresh) +apply(auto simp add: fresh_prod)[1] +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(case_tac "findc \c x3") +apply(simp add: abs_fresh) +apply(auto)[1] +apply(generate_fresh "coname") +apply(thin_tac "x1\?p") +apply(fresh_fun_simp) +apply(drule cmaps_fresh) +apply(auto simp add: fresh_prod)[1] +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(case_tac "findn \n x5") +apply(simp add: abs_fresh) +apply(auto)[1] +apply(generate_fresh "name") +apply(thin_tac "x1\?p") +apply(thin_tac "x3\?p") +apply(fresh_fun_simp) +apply(drule nmaps_fresh) +apply(auto simp add: fresh_prod)[1] +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(case_tac "findn \n x5") +apply(simp add: abs_fresh) +apply(auto)[1] +apply(generate_fresh "name") +apply(thin_tac "x1\?p") +apply(thin_tac "x3\?p") +apply(fresh_fun_simp) +apply(drule_tac a="x3" in nmaps_fresh) +apply(auto simp add: fresh_prod)[1] +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(case_tac "findc \c x4") +apply(simp add: abs_fresh abs_supp fin_supp) +apply(auto)[1] +apply(generate_fresh "coname") +apply(thin_tac "x1\?p") +apply(thin_tac "x2\?p") +apply(fresh_fun_simp) +apply(drule cmaps_fresh) +apply(auto simp add: fresh_prod)[1] +apply(simp add: abs_fresh fresh_prod fresh_atm abs_supp fin_supp) +apply(case_tac "findc \c x4") +apply(simp add: abs_fresh abs_supp fin_supp) +apply(auto)[1] +apply(generate_fresh "coname") +apply(thin_tac "x1\?p") +apply(thin_tac "x2\?p") +apply(fresh_fun_simp) +apply(drule_tac x="x2" in cmaps_fresh) +apply(auto simp add: fresh_prod)[1] +apply(simp add: abs_fresh fresh_prod fresh_atm abs_supp fin_supp) +apply(case_tac "findn \n x5") +apply(simp add: abs_fresh) +apply(auto)[1] +apply(generate_fresh "name") +apply(thin_tac "x1\?p") +apply(thin_tac "x3\?p") +apply(fresh_fun_simp) +apply(drule nmaps_fresh) +apply(auto simp add: fresh_prod)[1] +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(case_tac "findn \n x5") +apply(simp add: abs_fresh) +apply(auto)[1] +apply(generate_fresh "name") +apply(thin_tac "x1\?p") +apply(thin_tac "x3\?p") +apply(fresh_fun_simp) +apply(drule_tac a="x3" in nmaps_fresh) +apply(auto simp add: fresh_prod)[1] +apply(simp add: abs_fresh fresh_prod fresh_atm) +apply(fresh_guess)+ +done + +lemma case_cong: + assumes a: "B1=B2" "x1=x2" "y1=y2" + shows "(case B1 of None \ x1 | Some (x,P) \ y1 x P) = (case B2 of None \ x2 | Some (x,P) \ y2 x P)" +using a +apply(auto) +done + +lemma find_maps: + shows "\c cmaps a to (findc \c a)" + and "\n nmaps x to (findn \n x)" +apply(auto) +done + +lemma psubst_eqvt[eqvt]: + fixes pi1::"name prm" + and pi2::"coname prm" + shows "pi1\(\n,\c) = (pi1\\n),(pi1\\c)<(pi1\M)>" + and "pi2\(\n,\c) = (pi2\\n),(pi2\\c)<(pi2\M)>" +apply(nominal_induct M avoiding: \n \c rule: trm.induct) +apply(auto simp add: eq_bij fresh_bij eqvts perm_pi_simp) +apply(rule case_cong) +apply(rule find_maps) +apply(simp) +apply(perm_simp add: eqvts) +apply(rule case_cong) +apply(rule find_maps) +apply(simp) +apply(perm_simp add: eqvts) +apply(rule case_cong) +apply(rule find_maps) +apply(simp) +apply(perm_simp add: eqvts) +apply(rule case_cong) +apply(rule find_maps) +apply(simp) +apply(perm_simp add: eqvts) +apply(rule case_cong) +apply(rule find_maps) +apply(simp) +apply(perm_simp add: eqvts) +apply(rule case_cong) +apply(rule find_maps) +apply(simp) +apply(perm_simp add: eqvts) +apply(rule case_cong) +apply(rule find_maps) +apply(simp) +apply(perm_simp add: eqvts) +apply(rule case_cong) +apply(rule find_maps) +apply(simp) +apply(perm_simp add: eqvts) +apply(rule case_cong) +apply(rule find_maps) +apply(simp) +apply(perm_simp add: eqvts) +apply(rule case_cong) +apply(rule find_maps) +apply(simp) +apply(perm_simp add: eqvts) +apply(rule case_cong) +apply(rule find_maps) +apply(simp) +apply(perm_simp add: eqvts) +apply(rule case_cong) +apply(rule find_maps) +apply(simp) +apply(perm_simp add: eqvts) +apply(rule case_cong) +apply(rule find_maps) +apply(simp) +apply(perm_simp add: eqvts) +apply(rule case_cong) +apply(rule find_maps) +apply(simp) +apply(perm_simp add: eqvts) +apply(rule case_cong) +apply(rule find_maps) +apply(simp) +apply(perm_simp add: eqvts) +apply(rule case_cong) +apply(rule find_maps) +apply(simp) +apply(perm_simp add: eqvts) +apply(rule case_cong) +apply(rule find_maps) +apply(simp) +apply(perm_simp add: eqvts) +apply(rule case_cong) +apply(rule find_maps) +apply(simp) +apply(perm_simp add: eqvts) +apply(rule case_cong) +apply(rule find_maps) +apply(simp) +apply(perm_simp add: eqvts) +apply(rule case_cong) +apply(rule find_maps) +apply(simp) +apply(perm_simp add: eqvts) +done + +lemma ax_psubst: + assumes a: "\n,\c = Ax x a" + and b: "a\(\n,\c)" "x\(\n,\c)" + shows "M = Ax x a" +using a b +apply(nominal_induct M avoiding: \n \c rule: trm.induct) +apply(auto) +apply(drule lookup_unicity) +apply(simp)+ +apply(case_tac "findc \c coname") +apply(simp) +apply(auto)[1] +apply(generate_fresh "coname") +apply(fresh_fun_simp) +apply(simp) +apply(case_tac "findn \n name") +apply(simp) +apply(auto)[1] +apply(generate_fresh "name") +apply(fresh_fun_simp) +apply(simp) +apply(case_tac "findc \c coname3") +apply(simp) +apply(auto)[1] +apply(generate_fresh "coname") +apply(fresh_fun_simp) +apply(simp) +apply(case_tac "findn \n name2") +apply(simp) +apply(auto)[1] +apply(generate_fresh "name") +apply(fresh_fun_simp) +apply(simp) +apply(case_tac "findn \n name2") +apply(simp) +apply(auto)[1] +apply(generate_fresh "name") +apply(fresh_fun_simp) +apply(simp) +apply(case_tac "findc \c coname2") +apply(simp) +apply(auto)[1] +apply(generate_fresh "coname") +apply(fresh_fun_simp) +apply(simp) +apply(case_tac "findc \c coname2") +apply(simp) +apply(auto)[1] +apply(generate_fresh "coname") +apply(fresh_fun_simp) +apply(simp) +apply(case_tac "findn \n name3") +apply(simp) +apply(auto)[1] +apply(generate_fresh "name") +apply(fresh_fun_simp) +apply(simp) +apply(case_tac "findc \c coname2") +apply(simp) +apply(auto)[1] +apply(generate_fresh "coname") +apply(fresh_fun_simp) +apply(simp) +apply(case_tac "findn \n name2") +apply(simp) +apply(auto)[1] +apply(generate_fresh "name") +apply(fresh_fun_simp) +apply(simp) +done + +lemma better_Cut_substc1: + assumes a: "a\(P,b)" "b\N" + shows "(Cut .M (x).N){b:=(y).P} = Cut .(M{b:=(y).P}) (x).N" +using a +apply - +apply(generate_fresh "coname") +apply(generate_fresh "name") +apply(subgoal_tac "Cut .M (x).N = Cut .([(c,a)]\M) (ca).([(ca,x)]\N)") +apply(simp) +apply(rule trans) +apply(rule better_Cut_substc) +apply(simp) +apply(simp add: abs_fresh) +apply(auto)[1] +apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm fresh_atm) +apply(subgoal_tac"b\([(ca,x)]\N)") +apply(simp add: forget) +apply(simp add: trm.inject) +apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1] +apply(perm_simp) +apply(simp add: fresh_left calc_atm) +apply(simp add: trm.inject) +apply(rule conjI) +apply(rule sym) +apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1] +apply(rule sym) +apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1] +done + +lemma better_Cut_substc2: + assumes a: "x\(y,P)" "b\(a,M)" "N\Ax x b" + shows "(Cut .M (x).N){b:=(y).P} = Cut .M (x).(N{b:=(y).P})" +using a +apply - +apply(generate_fresh "coname") +apply(generate_fresh "name") +apply(subgoal_tac "Cut .M (x).N = Cut .([(c,a)]\M) (ca).([(ca,x)]\N)") +apply(simp) +apply(rule trans) +apply(rule better_Cut_substc) +apply(simp) +apply(simp add: abs_fresh) +apply(auto)[1] +apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst]) +apply(simp add: calc_atm fresh_atm fresh_prod) +apply(subgoal_tac"b\([(c,a)]\M)") +apply(simp add: forget) +apply(simp add: trm.inject) +apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1] +apply(perm_simp) +apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(simp add: trm.inject) +apply(rule conjI) +apply(rule sym) +apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1] +apply(rule sym) +apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1] +done + +lemma better_Cut_substn1: + assumes a: "y\(x,N)" "a\(b,P)" "M\Ax y a" + shows "(Cut .M (x).N){y:=.P} = Cut .(M{y:=.P}) (x).N" +using a +apply - +apply(generate_fresh "coname") +apply(generate_fresh "name") +apply(subgoal_tac "Cut .M (x).N = Cut .([(c,a)]\M) (ca).([(ca,x)]\N)") +apply(simp) +apply(rule trans) +apply(rule better_Cut_substn) +apply(simp add: abs_fresh) +apply(simp add: abs_fresh) +apply(auto)[1] +apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm fresh_atm fresh_prod) +apply(subgoal_tac"y\([(ca,x)]\N)") +apply(simp add: forget) +apply(simp add: trm.inject) +apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1] +apply(perm_simp) +apply(auto simp add: fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(simp add: trm.inject) +apply(rule conjI) +apply(rule sym) +apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1] +apply(rule sym) +apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1] +done + +lemma better_Cut_substn2: + assumes a: "x\(P,y)" "y\M" + shows "(Cut .M (x).N){y:=.P} = Cut .M (x).(N{y:=.P})" +using a +apply - +apply(generate_fresh "coname") +apply(generate_fresh "name") +apply(subgoal_tac "Cut .M (x).N = Cut .([(c,a)]\M) (ca).([(ca,x)]\N)") +apply(simp) +apply(rule trans) +apply(rule better_Cut_substn) +apply(simp add: abs_fresh) +apply(simp add: abs_fresh) +apply(auto)[1] +apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst]) +apply(simp add: calc_atm fresh_atm) +apply(subgoal_tac"y\([(c,a)]\M)") +apply(simp add: forget) +apply(simp add: trm.inject) +apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1] +apply(perm_simp) +apply(simp add: fresh_left calc_atm) +apply(simp add: trm.inject) +apply(rule conjI) +apply(rule sym) +apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1] +apply(rule sym) +apply(simp add: alpha eqvts calc_atm fresh_prod fresh_atm subst_fresh)[1] +done + +lemma psubst_fresh_name: + fixes x::"name" + assumes a: "x\\n" "x\\c" "x\M" + shows "x\\n,\c" +using a +apply(nominal_induct M avoiding: x \n \c rule: trm.induct) +apply(simp add: lookup_freshness) +apply(auto simp add: abs_fresh)[1] +apply(simp add: lookupc_freshness) +apply(simp add: lookupc_freshness) +apply(simp add: lookupc_freshness) +apply(simp add: lookupd_freshness) +apply(simp add: lookupd_freshness) +apply(simp add: lookupc_freshness) +apply(simp) +apply(case_tac "findc \c coname") +apply(auto simp add: abs_fresh)[1] +apply(auto)[1] +apply(generate_fresh "coname") +apply(fresh_fun_simp) +apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh) +apply(simp) +apply(case_tac "findn \n name") +apply(auto simp add: abs_fresh)[1] +apply(auto)[1] +apply(generate_fresh "name") +apply(fresh_fun_simp) +apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh) +apply(simp) +apply(case_tac "findc \c coname3") +apply(auto simp add: abs_fresh)[1] +apply(auto)[1] +apply(generate_fresh "coname") +apply(fresh_fun_simp) +apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh) +apply(simp) +apply(case_tac "findn \n name2") +apply(auto simp add: abs_fresh)[1] +apply(auto)[1] +apply(generate_fresh "name") +apply(fresh_fun_simp) +apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh) +apply(simp) +apply(case_tac "findn \n name2") +apply(auto simp add: abs_fresh)[1] +apply(auto)[1] +apply(generate_fresh "name") +apply(fresh_fun_simp) +apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh) +apply(simp) +apply(case_tac "findc \c coname2") +apply(auto simp add: abs_fresh)[1] +apply(auto)[1] +apply(generate_fresh "coname") +apply(fresh_fun_simp) +apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh) +apply(simp) +apply(case_tac "findc \c coname2") +apply(auto simp add: abs_fresh)[1] +apply(auto)[1] +apply(generate_fresh "coname") +apply(fresh_fun_simp) +apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh) +apply(simp) +apply(case_tac "findn \n name3") +apply(auto simp add: abs_fresh)[1] +apply(auto)[1] +apply(generate_fresh "name") +apply(fresh_fun_simp) +apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh) +apply(simp) +apply(case_tac "findc \c coname2") +apply(auto simp add: abs_fresh abs_supp fin_supp)[1] +apply(auto)[1] +apply(generate_fresh "coname") +apply(fresh_fun_simp) +apply(simp add: abs_fresh abs_supp fin_supp fresh_prod fresh_atm cmaps_fresh) +apply(simp) +apply(case_tac "findn \n name2") +apply(auto simp add: abs_fresh)[1] +apply(auto)[1] +apply(generate_fresh "name") +apply(fresh_fun_simp) +apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh) +done + +lemma psubst_fresh_coname: + fixes a::"coname" + assumes a: "a\\n" "a\\c" "a\M" + shows "a\\n,\c" +using a +apply(nominal_induct M avoiding: a \n \c rule: trm.induct) +apply(simp add: lookup_freshness) +apply(auto simp add: abs_fresh)[1] +apply(simp add: lookupd_freshness) +apply(simp add: lookupd_freshness) +apply(simp add: lookupc_freshness) +apply(simp add: lookupd_freshness) +apply(simp add: lookupc_freshness) +apply(simp add: lookupd_freshness) +apply(simp) +apply(case_tac "findc \c coname") +apply(auto simp add: abs_fresh)[1] +apply(auto)[1] +apply(generate_fresh "coname") +apply(fresh_fun_simp) +apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh) +apply(simp) +apply(case_tac "findn \n name") +apply(auto simp add: abs_fresh)[1] +apply(auto)[1] +apply(generate_fresh "name") +apply(fresh_fun_simp) +apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh) +apply(simp) +apply(case_tac "findc \c coname3") +apply(auto simp add: abs_fresh)[1] +apply(auto)[1] +apply(generate_fresh "coname") +apply(fresh_fun_simp) +apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh) +apply(simp) +apply(case_tac "findn \n name2") +apply(auto simp add: abs_fresh)[1] +apply(auto)[1] +apply(generate_fresh "name") +apply(fresh_fun_simp) +apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh) +apply(simp) +apply(case_tac "findn \n name2") +apply(auto simp add: abs_fresh)[1] +apply(auto)[1] +apply(generate_fresh "name") +apply(fresh_fun_simp) +apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh) +apply(simp) +apply(case_tac "findc \c coname2") +apply(auto simp add: abs_fresh)[1] +apply(auto)[1] +apply(generate_fresh "coname") +apply(fresh_fun_simp) +apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh) +apply(simp) +apply(case_tac "findc \c coname2") +apply(auto simp add: abs_fresh)[1] +apply(auto)[1] +apply(generate_fresh "coname") +apply(fresh_fun_simp) +apply(simp add: abs_fresh fresh_prod fresh_atm cmaps_fresh) +apply(simp) +apply(case_tac "findn \n name3") +apply(auto simp add: abs_fresh)[1] +apply(auto)[1] +apply(generate_fresh "name") +apply(fresh_fun_simp) +apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh) +apply(simp) +apply(case_tac "findc \c coname2") +apply(auto simp add: abs_fresh abs_supp fin_supp)[1] +apply(auto)[1] +apply(generate_fresh "coname") +apply(fresh_fun_simp) +apply(simp add: abs_fresh abs_supp fin_supp fresh_prod fresh_atm cmaps_fresh) +apply(simp) +apply(case_tac "findn \n name2") +apply(auto simp add: abs_fresh)[1] +apply(auto)[1] +apply(generate_fresh "name") +apply(fresh_fun_simp) +apply(simp add: abs_fresh fresh_prod fresh_atm nmaps_fresh) +done + +lemma psubst_csubst: + assumes a: "a\(\n,\c)" + shows "\n,((a,x,P)#\c) = ((\n,\c){a:=(x).P})" +using a +apply(nominal_induct M avoiding: a x \n \c P rule: trm.induct) +apply(auto simp add: fresh_list_cons fresh_prod)[1] +apply(simp add: lookup_csubst) +apply(simp add: fresh_list_cons fresh_prod) +apply(auto)[1] +apply(rule sym) +apply(rule trans) +apply(rule better_Cut_substc) +apply(simp) +apply(simp add: abs_fresh fresh_atm) +apply(simp add: lookupd_fresh) +apply(subgoal_tac "a\lookupc xa coname \n") +apply(simp add: forget) +apply(simp add: trm.inject) +apply(rule sym) +apply(simp add: alpha nrename_swap fresh_atm) +apply(rule lookupc_freshness) +apply(simp add: fresh_atm) +apply(rule sym) +apply(rule trans) +apply(rule better_Cut_substc) +apply(simp) +apply(simp add: abs_fresh fresh_atm) +apply(simp) +apply(rule conjI) +apply(rule impI) +apply(simp add: lookupd_unicity) +apply(rule impI) +apply(subgoal_tac "a\lookupc xa coname \n") +apply(subgoal_tac "a\lookupd name aa \c") +apply(simp add: forget) +apply(rule lookupd_freshness) +apply(simp add: fresh_atm) +apply(rule lookupc_freshness) +apply(simp add: fresh_atm) +apply(rule sym) +apply(rule trans) +apply(rule better_Cut_substc) +apply(simp) +apply(simp add: abs_fresh fresh_atm) +apply(simp) +apply(rule conjI) +apply(rule impI) +apply(drule ax_psubst) +apply(simp) +apply(simp) +apply(blast) +apply(rule impI) +apply(subgoal_tac "a\lookupc xa coname \n") +apply(simp add: forget) +apply(rule lookupc_freshness) +apply(simp add: fresh_atm) +apply(rule sym) +apply(rule trans) +apply(rule better_Cut_substc) +apply(simp) +apply(simp add: abs_fresh fresh_atm) +apply(simp) +apply(rule conjI) +apply(rule impI) +apply(simp add: trm.inject) +apply(rule sym) +apply(simp add: alpha) +apply(simp add: alpha nrename_swap fresh_atm) +apply(simp add: lookupd_fresh) +apply(rule sym) +apply(rule trans) +apply(rule better_Cut_substc) +apply(simp) +apply(simp add: abs_fresh fresh_atm) +apply(simp) +apply(rule conjI) +apply(rule impI) +apply(simp add: lookupd_unicity) +apply(rule impI) +apply(subgoal_tac "a\lookupd name aa \c") +apply(simp add: forget) +apply(rule lookupd_freshness) +apply(simp add: fresh_atm) +apply(rule sym) +apply(rule trans) +apply(rule better_Cut_substc) +apply(simp) +apply(simp add: abs_fresh fresh_atm) +apply(simp) +apply(rule impI) +apply(drule ax_psubst) +apply(simp) +apply(simp) +apply(blast) +(* NotR *) +apply(simp) +apply(case_tac "findc \c coname") +apply(simp) +apply(auto simp add: fresh_list_cons fresh_prod)[1] +apply(simp) +apply(auto simp add: fresh_list_cons fresh_prod)[1] +apply(drule cmaps_false) +apply(assumption) +apply(simp) +apply(generate_fresh "coname") +apply(fresh_fun_simp) +apply(fresh_fun_simp) +apply(rule sym) +apply(rule trans) +apply(rule better_Cut_substc1) +apply(simp) +apply(simp add: cmaps_fresh) +apply(auto simp add: fresh_prod fresh_atm)[1] +(* NotL *) +apply(simp) +apply(case_tac "findn \n name") +apply(simp) +apply(auto simp add: fresh_list_cons fresh_prod)[1] +apply(simp) +apply(auto simp add: fresh_list_cons fresh_prod)[1] +apply(generate_fresh "name") +apply(fresh_fun_simp) +apply(fresh_fun_simp) +apply(drule_tac a="a" in nmaps_fresh) +apply(assumption) +apply(rule sym) +apply(rule trans) +apply(rule better_Cut_substc2) +apply(simp) +apply(simp) +apply(simp) +apply(simp) +(* AndR *) +apply(simp) +apply(case_tac "findc \c coname3") +apply(simp) +apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1] +apply(simp) +apply(auto simp add: fresh_list_cons fresh_prod)[1] +apply(drule cmaps_false) +apply(assumption) +apply(simp) +apply(generate_fresh "coname") +apply(fresh_fun_simp) +apply(fresh_fun_simp) +apply(rule sym) +apply(rule trans) +apply(rule better_Cut_substc1) +apply(simp) +apply(simp add: cmaps_fresh) +apply(auto simp add: psubst_fresh_coname fresh_prod fresh_atm)[1] +(* AndL1 *) +apply(simp) +apply(case_tac "findn \n name2") +apply(simp) +apply(auto simp add: fresh_list_cons fresh_prod)[1] +apply(simp) +apply(auto simp add: fresh_list_cons fresh_prod)[1] +apply(generate_fresh "name") +apply(fresh_fun_simp) +apply(fresh_fun_simp) +apply(drule_tac a="a" in nmaps_fresh) +apply(assumption) +apply(rule sym) +apply(rule trans) +apply(rule better_Cut_substc2) +apply(simp) +apply(simp) +apply(simp) +apply(auto simp add: psubst_fresh_coname fresh_prod fresh_atm)[1] +(* AndL2 *) +apply(simp) +apply(case_tac "findn \n name2") +apply(simp) +apply(auto simp add: fresh_list_cons fresh_prod)[1] +apply(simp) +apply(auto simp add: fresh_list_cons fresh_prod)[1] +apply(generate_fresh "name") +apply(fresh_fun_simp) +apply(fresh_fun_simp) +apply(drule_tac a="a" in nmaps_fresh) +apply(assumption) +apply(rule sym) +apply(rule trans) +apply(rule better_Cut_substc2) +apply(simp) +apply(simp) +apply(simp) +apply(auto simp add: psubst_fresh_coname fresh_prod fresh_atm)[1] +(* OrR1 *) +apply(simp) +apply(case_tac "findc \c coname2") +apply(simp) +apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1] +apply(simp) +apply(auto simp add: fresh_list_cons fresh_prod)[1] +apply(drule cmaps_false) +apply(assumption) +apply(simp) +apply(generate_fresh "coname") +apply(fresh_fun_simp) +apply(fresh_fun_simp) +apply(rule sym) +apply(rule trans) +apply(rule better_Cut_substc1) +apply(simp) +apply(simp add: cmaps_fresh) +apply(auto simp add: psubst_fresh_coname fresh_prod fresh_atm)[1] +(* OrR2 *) +apply(simp) +apply(case_tac "findc \c coname2") +apply(simp) +apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1] +apply(simp) +apply(auto simp add: fresh_list_cons fresh_prod)[1] +apply(drule cmaps_false) +apply(assumption) +apply(simp) +apply(generate_fresh "coname") +apply(fresh_fun_simp) +apply(fresh_fun_simp) +apply(rule sym) +apply(rule trans) +apply(rule better_Cut_substc1) +apply(simp) +apply(simp add: cmaps_fresh) +apply(auto simp add: psubst_fresh_coname fresh_prod fresh_atm)[1] +(* OrL *) +apply(simp) +apply(case_tac "findn \n name3") +apply(simp) +apply(auto simp add: fresh_list_cons psubst_fresh_name fresh_atm fresh_prod)[1] +apply(simp) +apply(auto simp add: fresh_list_cons fresh_prod)[1] +apply(generate_fresh "name") +apply(fresh_fun_simp) +apply(fresh_fun_simp) +apply(drule_tac a="a" in nmaps_fresh) +apply(assumption) +apply(rule sym) +apply(rule trans) +apply(rule better_Cut_substc2) +apply(simp) +apply(simp) +apply(simp) +apply(auto simp add: psubst_fresh_name fresh_prod fresh_atm)[1] +(* ImpR *) +apply(simp) +apply(case_tac "findc \c coname2") +apply(simp) +apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1] +apply(simp) +apply(auto simp add: fresh_list_cons fresh_prod)[1] +apply(drule cmaps_false) +apply(assumption) +apply(simp) +apply(generate_fresh "coname") +apply(fresh_fun_simp) +apply(fresh_fun_simp) +apply(rule sym) +apply(rule trans) +apply(rule better_Cut_substc1) +apply(simp) +apply(simp add: cmaps_fresh) +apply(auto simp add: psubst_fresh_coname fresh_prod fresh_atm)[1] +(* ImpL *) +apply(simp) +apply(case_tac "findn \n name2") +apply(simp) +apply(auto simp add: fresh_list_cons psubst_fresh_coname psubst_fresh_name fresh_atm fresh_prod)[1] +apply(simp) +apply(auto simp add: fresh_list_cons fresh_prod)[1] +apply(generate_fresh "name") +apply(fresh_fun_simp) +apply(fresh_fun_simp) +apply(simp add: abs_fresh subst_fresh) +apply(drule_tac a="a" in nmaps_fresh) +apply(assumption) +apply(rule sym) +apply(rule trans) +apply(rule better_Cut_substc2) +apply(simp) +apply(simp) +apply(simp) +apply(auto simp add: psubst_fresh_coname psubst_fresh_name fresh_prod fresh_atm)[1] +done + +lemma psubst_nsubst: + assumes a: "x\(\n,\c)" + shows "((x,a,P)#\n),\c = ((\n,\c){x:=.P})" +using a +apply(nominal_induct M avoiding: a x \n \c P rule: trm.induct) +apply(auto simp add: fresh_list_cons fresh_prod)[1] +apply(simp add: lookup_fresh) +apply(rule lookupb_lookupa) +apply(simp) +apply(rule sym) +apply(rule forget) +apply(rule lookup_freshness) +apply(simp add: fresh_atm) +apply(auto simp add: lookupc_freshness fresh_list_cons fresh_prod)[1] +apply(simp add: lookupc_fresh) +apply(rule sym) +apply(rule trans) +apply(rule better_Cut_substn) +apply(simp add: abs_fresh) +apply(simp add: abs_fresh fresh_atm) +apply(simp add: lookupd_fresh) +apply(subgoal_tac "x\lookupd name aa \c") +apply(simp add: forget) +apply(simp add: trm.inject) +apply(rule sym) +apply(simp add: alpha crename_swap fresh_atm) +apply(rule lookupd_freshness) +apply(simp add: fresh_atm) +apply(rule sym) +apply(rule trans) +apply(rule better_Cut_substn) +apply(simp add: abs_fresh) +apply(simp add: abs_fresh fresh_atm) +apply(simp) +apply(rule conjI) +apply(rule impI) +apply(simp add: lookupc_unicity) +apply(rule impI) +apply(subgoal_tac "x\lookupc xa coname \n") +apply(subgoal_tac "x\lookupd name aa \c") +apply(simp add: forget) +apply(rule lookupd_freshness) +apply(simp add: fresh_atm) +apply(rule lookupc_freshness) +apply(simp add: fresh_atm) +apply(rule sym) +apply(rule trans) +apply(rule better_Cut_substn) +apply(simp add: abs_fresh) +apply(simp add: abs_fresh fresh_atm) +apply(simp) +apply(rule conjI) +apply(rule impI) +apply(simp add: trm.inject) +apply(rule sym) +apply(simp add: alpha crename_swap fresh_atm) +apply(rule impI) +apply(simp add: lookupc_fresh) +apply(rule sym) +apply(rule trans) +apply(rule better_Cut_substn) +apply(simp add: abs_fresh) +apply(simp add: abs_fresh fresh_atm) +apply(simp) +apply(rule conjI) +apply(rule impI) +apply(simp add: lookupc_unicity) +apply(rule impI) +apply(subgoal_tac "x\lookupc xa coname \n") +apply(simp add: forget) +apply(rule lookupc_freshness) +apply(simp add: fresh_prod fresh_atm) +apply(rule sym) +apply(rule trans) +apply(rule better_Cut_substn) +apply(simp add: abs_fresh) +apply(simp add: abs_fresh fresh_atm) +apply(simp) +apply(rule conjI) +apply(rule impI) +apply(drule ax_psubst) +apply(simp) +apply(simp) +apply(simp) +apply(blast) +apply(rule impI) +apply(subgoal_tac "x\lookupd name aa \c") +apply(simp add: forget) +apply(rule lookupd_freshness) +apply(simp add: fresh_atm) +apply(rule sym) +apply(rule trans) +apply(rule better_Cut_substn) +apply(simp add: abs_fresh) +apply(simp add: abs_fresh fresh_atm) +apply(simp) +apply(rule impI) +apply(drule ax_psubst) +apply(simp) +apply(simp) +apply(blast) +(* NotR *) +apply(simp) +apply(case_tac "findc \c coname") +apply(simp) +apply(auto simp add: fresh_list_cons fresh_prod)[1] +apply(simp) +apply(auto simp add: fresh_list_cons fresh_prod)[1] +apply(generate_fresh "coname") +apply(fresh_fun_simp) +apply(fresh_fun_simp) +apply(rule sym) +apply(rule trans) +apply(rule better_Cut_substn1) +apply(simp add: cmaps_fresh) +apply(simp) +apply(simp) +apply(simp) +(* NotL *) +apply(simp) +apply(case_tac "findn \n name") +apply(simp) +apply(auto simp add: fresh_list_cons fresh_prod)[1] +apply(simp) +apply(auto simp add: fresh_list_cons fresh_prod)[1] +apply(drule nmaps_false) +apply(simp) +apply(simp) +apply(generate_fresh "name") +apply(fresh_fun_simp) +apply(fresh_fun_simp) +apply(rule sym) +apply(rule trans) +apply(rule better_Cut_substn2) +apply(simp) +apply(simp add: nmaps_fresh) +apply(simp add: fresh_prod fresh_atm) +(* AndR *) +apply(simp) +apply(case_tac "findc \c coname3") +apply(simp) +apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1] +apply(simp) +apply(auto simp add: fresh_list_cons fresh_prod)[1] +apply(generate_fresh "coname") +apply(fresh_fun_simp) +apply(fresh_fun_simp) +apply(rule sym) +apply(rule trans) +apply(rule better_Cut_substn1) +apply(simp add: cmaps_fresh) +apply(simp) +apply(simp) +apply(auto simp add: psubst_fresh_coname fresh_prod fresh_atm)[1] +(* AndL1 *) +apply(simp) +apply(case_tac "findn \n name2") +apply(simp) +apply(auto simp add: fresh_list_cons fresh_prod)[1] +apply(simp) +apply(auto simp add: fresh_list_cons fresh_prod)[1] +apply(drule nmaps_false) +apply(simp) +apply(simp) +apply(generate_fresh "name") +apply(fresh_fun_simp) +apply(fresh_fun_simp) +apply(rule sym) +apply(rule trans) +apply(rule better_Cut_substn2) +apply(simp) +apply(simp add: nmaps_fresh) +apply(auto simp add: psubst_fresh_coname fresh_prod fresh_atm)[1] +(* AndL2 *) +apply(simp) +apply(case_tac "findn \n name2") +apply(simp) +apply(auto simp add: fresh_list_cons fresh_prod)[1] +apply(simp) +apply(auto simp add: fresh_list_cons fresh_prod)[1] +apply(drule nmaps_false) +apply(simp) +apply(simp) +apply(generate_fresh "name") +apply(fresh_fun_simp) +apply(fresh_fun_simp) +apply(rule sym) +apply(rule trans) +apply(rule better_Cut_substn2) +apply(simp) +apply(simp add: nmaps_fresh) +apply(auto simp add: psubst_fresh_coname fresh_prod fresh_atm)[1] +(* OrR1 *) +apply(simp) +apply(case_tac "findc \c coname2") +apply(simp) +apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1] +apply(simp) +apply(auto simp add: fresh_list_cons fresh_prod)[1] +apply(generate_fresh "coname") +apply(fresh_fun_simp) +apply(fresh_fun_simp) +apply(rule sym) +apply(rule trans) +apply(rule better_Cut_substn1) +apply(simp add: cmaps_fresh) +apply(simp) +apply(simp) +apply(auto simp add: psubst_fresh_coname fresh_prod fresh_atm)[1] +(* OrR2 *) +apply(simp) +apply(case_tac "findc \c coname2") +apply(simp) +apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1] +apply(simp) +apply(auto simp add: fresh_list_cons fresh_prod)[1] +apply(generate_fresh "coname") +apply(fresh_fun_simp) +apply(fresh_fun_simp) +apply(rule sym) +apply(rule trans) +apply(rule better_Cut_substn1) +apply(simp add: cmaps_fresh) +apply(simp) +apply(simp) +apply(auto simp add: psubst_fresh_coname fresh_prod fresh_atm)[1] +(* OrL *) +apply(simp) +apply(case_tac "findn \n name3") +apply(simp) +apply(auto simp add: fresh_list_cons psubst_fresh_name fresh_atm fresh_prod)[1] +apply(simp) +apply(auto simp add: fresh_list_cons fresh_prod)[1] +apply(drule nmaps_false) +apply(simp) +apply(simp) +apply(generate_fresh "name") +apply(fresh_fun_simp) +apply(fresh_fun_simp) +apply(rule sym) +apply(rule trans) +apply(rule better_Cut_substn2) +apply(simp) +apply(simp add: nmaps_fresh) +apply(auto simp add: psubst_fresh_name fresh_prod fresh_atm)[1] +(* ImpR *) +apply(simp) +apply(case_tac "findc \c coname2") +apply(simp) +apply(auto simp add: psubst_fresh_coname fresh_list_cons fresh_prod fresh_atm)[1] +apply(simp) +apply(auto simp add: fresh_list_cons fresh_prod)[1] +apply(generate_fresh "coname") +apply(fresh_fun_simp) +apply(fresh_fun_simp) +apply(rule sym) +apply(rule trans) +apply(rule better_Cut_substn1) +apply(simp add: cmaps_fresh) +apply(simp) +apply(simp) +apply(auto simp add: psubst_fresh_coname fresh_prod fresh_atm)[1] +(* ImpL *) +apply(simp) +apply(case_tac "findn \n name2") +apply(simp) +apply(auto simp add: fresh_list_cons psubst_fresh_coname psubst_fresh_name fresh_atm fresh_prod)[1] +apply(simp) +apply(auto simp add: fresh_list_cons fresh_prod)[1] +apply(drule nmaps_false) +apply(simp) +apply(simp) +apply(generate_fresh "name") +apply(fresh_fun_simp) +apply(fresh_fun_simp) +apply(rule sym) +apply(rule trans) +apply(rule better_Cut_substn2) +apply(simp) +apply(simp add: nmaps_fresh) +apply(auto simp add: psubst_fresh_coname psubst_fresh_name fresh_prod fresh_atm)[1] +done + +definition + ncloses :: "(name\coname\trm) list\(name\ty) list \ bool" ("_ ncloses _" [55,55] 55) +where + "\n ncloses \ \ \x B. ((x,B) \ set \ \ (\c P. \n nmaps x to Some (c,P) \ :P \ (\\)))" + +definition + ccloses :: "(coname\name\trm) list\(coname\ty) list \ bool" ("_ ccloses _" [55,55] 55) +where + "\c ccloses \ \ \a B. ((a,B) \ set \ \ (\x P. \c cmaps a to Some (x,P) \ (x):P \ (\(B)\)))" + +lemma ncloses_elim: + assumes a: "(x,B) \ set \" + and b: "\n ncloses \" + shows "\c P. \n nmaps x to Some (c,P) \ :P \ (\\)" +using a b by (auto simp add: ncloses_def) + +lemma ccloses_elim: + assumes a: "(a,B) \ set \" + and b: "\c ccloses \" + shows "\x P. \c cmaps a to Some (x,P) \ (x):P \ (\(B)\)" +using a b by (auto simp add: ccloses_def) + +lemma ncloses_subset: + assumes a: "\n ncloses \" + and b: "set \' \ set \" + shows "\n ncloses \'" +using a b by (auto simp add: ncloses_def) + +lemma ccloses_subset: + assumes a: "\c ccloses \" + and b: "set \' \ set \" + shows "\c ccloses \'" +using a b by (auto simp add: ccloses_def) + +lemma validc_fresh: + fixes a::"coname" + and \::"(coname\ty) list" + assumes a: "a\\" + shows "\(\B. (a,B)\set \)" +using a +apply(induct \) +apply(auto simp add: fresh_list_cons fresh_prod fresh_atm) +done + +lemma validn_fresh: + fixes x::"name" + and \::"(name\ty) list" + assumes a: "x\\" + shows "\(\B. (x,B)\set \)" +using a +apply(induct \) +apply(auto simp add: fresh_list_cons fresh_prod fresh_atm) +done + +lemma ccloses_extend: + assumes a: "\c ccloses \" "a\\" "a\\c" "(x):P\\(B)\" + shows "(a,x,P)#\c ccloses (a,B)#\" +using a +apply(simp add: ccloses_def) +apply(drule validc_fresh) +apply(auto) +done + +lemma ncloses_extend: + assumes a: "\n ncloses \" "x\\" "x\\n" ":P\\\" + shows "(x,a,P)#\n ncloses (x,B)#\" +using a +apply(simp add: ncloses_def) +apply(drule validn_fresh) +apply(auto) +done + + text {* typing relation *} inductive2 typing :: "ctxtn \ trm \ ctxtc \ bool" ("_ \ _ \ _" [100,100,100] 100) where - "\validn \;validc \; (x,B)\set \; (a,B)\set \\ \ \ \ Ax x a \ \" -| "\x\\; ((x,B)#\) \ M \ \\ \ \ \ NotR (x).M a \ ((a,NOT B)#\)" -| "\a\\; \ \ M \ ((a,B)#\)\ \ ((x,NOT B)#\) \ NotL .M x \ \" -| "\x\\; ((x,B1)#\) \ M \ \\ \ ((y,B1 AND B2)#\) \ AndL1 (x).M y \ \" -| "\x\\; ((x,B2)#\) \ M \ \\ \ ((y,B1 AND B2)#\) \ AndL2 (x).M y \ \" -| "\a\\;b\\; \ \ M \ ((a,B)#\); \ \ N \ ((b,C)#\)\ \ \ \ AndR .M .N c \ ((c,B AND C)#\)" -| "\x\\;y\\; ((x,B)#\) \ M \ \; ((y,C)#\) \ N \ \\ \ ((z,B OR C)#\) \ OrL (x).M (y).N z \ \" -| "\a\\; \ \ M \ ((a,B1)#\)\ \ \ \ OrR1 .M b \ ((b,B1 OR B2)#\)" -| "\a\\; \ \ M \ ((a,B2)#\)\ \ \ \ OrR2 .M b \ ((b,B1 OR B2)#\)" -| "\a\\;x\\; \ \ M \ ((a,B)#\); ((x,C)#\) \ N \ \\ \ ((y,B IMPL C)#\) \ ImpL .M (x).N y \ \" -| "\a\\;x\\; ((x,B)#\) \ M \ ((a,C)#\)\ \ \ \ ImpR (x)..M b \ ((b,B IMPL C)#\)" -| "\a\\1;x\\2; validn(\1@\2); validc (\1@\2); \1 \ M \ ((a,B)#\1); ((x,B)#\2) \ N \ \2\ - \ (\1@\2) \ Cut .M (x).N \ (\1@\2)" - -text {* relations about freshly introducing a name or coname *} - -inductive2 - fin :: "trm \ name \ bool" -where - "fin (Ax x a,x)" -| "x\M \ fin (NotL .M x,x)" -| "y\[x].M \ fin (AndL1 (x).M y,y)" -| "y\[x].M \ fin (AndL2 (x).M y,y)" -| "\z\[x].M;z\[y].N\ \ fin (OrL (x).M (y).N z,z)" -| "\y\M;y\[x].N\ \ fin (ImpL .M (x).N y,y)" - -inductive2 - fic :: "trm \ coname \ bool" -where - "fic (Ax x a,a)" -| "a\M \ fic (NotR (x).M a,a)" -| "\c\[a].M;c\[b].N\ \ fic (AndR .M .N c,c)" -| "b\[a].M \ fic (OrR1 .M b,b)" -| "b\[a].M \ fic (OrR2 .M b,b)" -| "\b\[a].M\ \ fic (ImpR (x)..M b,b)" - -text {* cut-reductions *} - -inductive2 - red :: "trm \ trm \ bool" ("_ \\<^isub>c _" [100,100] 100) -where - "fic (M,a) \ Cut .M (x).(Ax x b) \\<^isub>c M[a\c>b]" -| "fin (M,x) \ Cut .(Ax y a) (x).M \\<^isub>c M[x\n>y]" -| "\fic (NotR (x).M a,a); fin (NotL .N y,y)\ \ - Cut .(NotR (x).M a) (y).(NotL .N y) \\<^isub>c Cut .N (x).M" -| "\fic (AndR .M1 .M2 b,b); fin (AndL1 (x).N y,y)\ \ - Cut .(AndR .M1 .M2 b) (y).(AndL1 (x).N y) \\<^isub>c Cut .M1 (x).N" -| "\fic (AndR .M1 .M2 b,b); fin (AndL2 (x).N y,y)\ \ - Cut .(AndR .M1 .M2 b) (y).(AndL2 (x).N y) \\<^isub>c Cut .M2 (x).N" -| "\fic (AndR .M1 .M2 b,b); fin (AndL1 (x).N y,y)\ \ - Cut .(AndR .M1 .M2 b) (y).(AndL1 (x).N y) \\<^isub>c Cut .M1 (x).N" -| "\fic (OrR1 .M b,b); fin (OrL (x1).N1 (x2).N2 y,y)\ \ - Cut .(OrR1 .M b) (y).(OrL (x1).N1 (x2).N2 y) \\<^isub>c Cut .M (x1).N1" -| "\fic (OrR2 .M b,b); fin (OrL (x1).N1 (x2).N2 y,y)\ \ - Cut .(OrR2 .M b) (y).(OrL (x1).N1 (x2).N2 y) \\<^isub>c Cut .M (x2).N2" -| "\fin (ImpL .N (y).P z,z); fic (ImpR (x)..M b,b)\ \ - Cut .(ImpR (x)..M b) (z).(ImpL .N (y).P z) \\<^isub>c Cut .(Cut .N (x).M) (y).P" -| "\fin (ImpL .N (y).P z,z); fic (ImpR (x)..M b,b)\ \ - Cut .(ImpR (x)..M b) (z).(ImpL .N (y).P z) \\<^isub>c Cut .N (x).(Cut .M (y).P)" -| "\fic (M,a) \ Cut .M (x).N \\<^isub>c M[a:=(x).N]" -| "\fin (N,x) \ Cut .M (x).N \\<^isub>c N[x:=.M]" - -text {* PROOFS *} - -lemma validn_eqvt[eqvt]: - fixes pi1:: "name prm" - and pi2:: "coname prm" - assumes a: "validn \" - shows "validn (pi1\\)" and "validn (pi2\\)" -using a by (induct) (auto simp add: fresh_bij) - -lemma validc_eqvt[eqvt]: - fixes pi1:: "name prm" - and pi2:: "coname prm" - assumes a: "validc \" - shows "validc (pi1\\)" and "validc (pi2\\)" -using a by (induct) (auto simp add: fresh_bij) - -text {* Weakening Lemma *} - -abbreviation - "subn" :: "ctxtn \ ctxtn \ bool" (" _ \n _ " [80,80] 80) -where - "\1 \n \2 \ \x B. (x,B)\set \1 \ (x,B)\set \2" - -abbreviation - "subc" :: "ctxtc \ ctxtc \ bool" (" _ \c _ " [80,80] 80) -where - "\1 \c \2 \ \a B. (a,B)\set \1 \ (a,B)\set \2" - + TAx: "\validn \;validc \; (x,B)\set \; (a,B)\set \\ \ \ \ Ax x a \ \" +| TNotR: "\x\\; ((x,B)#\) \ M \ \; set \' = {(a,NOT B)}\set \; validc \'\ + \ \ \ NotR (x).M a \ \'" +| TNotL: "\a\\; \ \ M \ ((a,B)#\); set \' = {(x,NOT B)} \ set \; validn \'\ + \ \' \ NotL .M x \ \" +| TAndL1: "\x\(\,y); ((x,B1)#\) \ M \ \; set \' = {(y,B1 AND B2)} \ set \; validn \'\ + \ \' \ AndL1 (x).M y \ \" +| TAndL2: "\x\(\,y); ((x,B2)#\) \ M \ \; set \' = {(y,B1 AND B2)} \ set \; validn \'\ + \ \' \ AndL2 (x).M y \ \" +| TAndR: "\a\(\,N,c); b\(\,M,c); a\b; \ \ M \ ((a,B)#\); \ \ N \ ((b,C)#\); + set \' = {(c,B AND C)}\set \; validc \'\ + \ \ \ AndR .M .N c \ \'" +| TOrL: "\x\(\,N,z); y\(\,M,z); x\y; ((x,B)#\) \ M \ \; ((y,C)#\) \ N \ \; + set \' = {(z,B OR C)} \ set \; validn \'\ + \ \' \ OrL (x).M (y).N z \ \" +| TOrR1: "\a\(\,b); \ \ M \ ((a,B1)#\); set \' = {(b,B1 OR B2)}\set \; validc \'\ + \ \ \ OrR1 .M b \ \'" +| TOrR2: "\a\(\,b); \ \ M \ ((a,B2)#\); set \' = {(b,B1 OR B2)}\set \; validc \'\ + \ \ \ OrR2 .M b \ \'" +| TImpL: "\a\(\,N); x\(\,M,y); \ \ M \ ((a,B)#\); ((x,C)#\) \ N \ \; + set \' = {(y,B IMP C)} \ set \; validn \'\ + \ \' \ ImpL .M (x).N y \ \" +| TImpR: "\a\(\,b); x\\; ((x,B)#\) \ M \ ((a,C)#\); set \' = {(b,B IMP C)}\set \; validc \'\ + \ \ \ ImpR (x)..M b \ \'" +| TCut: "\a\(\,N); x\(\,M); \ \ M \ ((a,B)#\); ((x,B)#\) \ N \ \\ + \ \ \ Cut .M (x).N \ \" + +equivariance typing + +lemma fresh_set_member: + fixes x::"name" + and a::"coname" + shows "x\L \ e\set L \ x\e" + and "a\L \ e\set L \ a\e" +by (induct L) (auto simp add: fresh_list_cons) + +lemma fresh_subset: + fixes x::"name" + and a::"coname" + shows "x\L \ set L' \ set L \ x\L'" + and "a\L \ set L' \ set L \ a\L'" +apply(induct L' arbitrary: L) +apply(auto simp add: fresh_list_cons fresh_list_nil intro: fresh_set_member) +done + +lemma fresh_subset_ext: + fixes x::"name" + and a::"coname" + shows "x\L \ x\e \ set L' \ set (e#L) \ x\L'" + and "a\L \ a\e \ set L' \ set (e#L) \ a\L'" +apply(induct L' arbitrary: L) +apply(auto simp add: fresh_list_cons fresh_list_nil intro: fresh_set_member) +done + +lemma fresh_under_insert: + fixes x::"name" + and a::"coname" + and \::"ctxtn" + and \::"ctxtc" + shows "x\\ \ x\y \ set \' = insert (y,B) (set \) \ x\\'" + and "a\\ \ a\c \ set \' = insert (c,B) (set \) \ a\\'" +apply(rule fresh_subset_ext(1)) +apply(auto simp add: fresh_prod fresh_atm fresh_ty) +apply(rule fresh_subset_ext(2)) +apply(auto simp add: fresh_prod fresh_atm fresh_ty) +done + +nominal_inductive typing + apply (simp_all add: abs_fresh fresh_atm fresh_list_cons fresh_prod fresh_ty fresh_ctxt + fresh_list_append abs_supp fin_supp) + apply(auto intro: fresh_under_insert) + done + +lemma validn_elim: + assumes a: "validn ((x,B)#\)" + shows "validn \ \ x\\" +using a +apply(erule_tac validn.cases) +apply(auto) +done + +lemma validc_elim: + assumes a: "validc ((a,B)#\)" + shows "validc \ \ a\\" +using a +apply(erule_tac validc.cases) +apply(auto) +done + +lemma context_fresh: + fixes x::"name" + and a::"coname" + shows "x\\ \ \(\B. (x,B)\set \)" + and "a\\ \ \(\B. (a,B)\set \)" +apply - +apply(induct \) +apply(auto simp add: fresh_list_cons fresh_prod fresh_atm) +apply(induct \) +apply(auto simp add: fresh_list_cons fresh_prod fresh_atm) +done + +lemma typing_implies_valid: + assumes a: "\ \ M \ \" + shows "validn \ \ validc \" +using a +apply(nominal_induct rule: typing.strong_induct) +apply(auto dest: validn_elim validc_elim) +done + +lemma ty_perm: + fixes pi1::"name prm" + and pi2::"coname prm" + and B::"ty" + shows "pi1\B=B" and "pi2\B=B" +apply(nominal_induct B rule: ty.induct) +apply(auto simp add: perm_string) +done + +lemma ctxt_perm: + fixes pi1::"name prm" + and pi2::"coname prm" + and \::"ctxtn" + and \::"ctxtc" + shows "pi2\\=\" and "pi1\\=\" +apply - +apply(induct \) +apply(auto simp add: calc_atm ty_perm) +apply(induct \) +apply(auto simp add: calc_atm ty_perm) +done + +lemma typing_Ax_elim1: + assumes a: "\ \ Ax x a \ ((a,B)#\)" + shows "(x,B)\set \" +using a +apply(erule_tac typing.cases) +apply(simp_all add: trm.inject) +apply(auto) +apply(auto dest: validc_elim context_fresh) +done + +lemma typing_Ax_elim2: + assumes a: "((x,B)#\) \ Ax x a \ \" + shows "(a,B)\set \" +using a +apply(erule_tac typing.cases) +apply(simp_all add: trm.inject) +apply(auto dest!: validn_elim context_fresh) +done + +lemma psubst_Ax_aux: + assumes a: "\c cmaps a to Some (y,N)" + shows "lookupb x a \c c P = Cut .P (y).N" +using a +apply(induct \c) +apply(auto) +done + +lemma psubst_Ax: + assumes a: "\n nmaps x to Some (c,P)" + and b: "\c cmaps a to Some (y,N)" + shows "\n,\c = Cut .P (y).N" +using a b +apply(induct \n) +apply(auto simp add: psubst_Ax_aux) +done + +lemma psubst_Cut: + assumes a: "\x. M\Ax x c" + and b: "\a. N\Ax x a" + and c: "c\(\n,\c,N)" "x\(\n,\c,M)" + shows "\n,\c.M (x).N> = Cut .(\n,\c) (x).(\n,\c)" +using a b c +apply(simp) +done + +lemma all_CAND: + assumes a: "\ \ M \ \" + and b: "\n ncloses \" + and c: "\c ccloses \" + shows "SNa (\n,\c)" +using a b c +proof(nominal_induct avoiding: \n \c rule: typing.strong_induct) + case (TAx \ \ x B a \n \c) + then show ?case + apply - + apply(drule ncloses_elim) + apply(assumption) + apply(drule ccloses_elim) + apply(assumption) + apply(erule exE)+ + apply(erule conjE)+ + apply(rule_tac s="Cut .P (xa).Pa" and t="\n,\c" in subst) + apply(rule sym) + apply(simp only: psubst_Ax) + apply(simp add: CUT_SNa) + done +next + case (TNotR x \ B M \ \' a \n \c) + then show ?case + apply(simp) + apply(subgoal_tac "(a,NOT B) \ set \'") + apply(drule ccloses_elim) + apply(assumption) + apply(erule exE)+ + apply(simp) + apply(generate_fresh "coname") + apply(fresh_fun_simp) + apply(rule_tac B="NOT B" in CUT_SNa) + apply(simp) + apply(rule disjI2) + apply(rule disjI2) + apply(rule_tac x="c" in exI) + apply(rule_tac x="x" in exI) + apply(rule_tac x="\n,\c" in exI) + apply(simp) + apply(rule conjI) + apply(rule fic.intros) + apply(rule psubst_fresh_coname) + apply(simp) + apply(simp) + apply(simp) + apply(rule BINDING_implies_CAND) + apply(unfold BINDINGn_def) + apply(simp) + apply(rule_tac x="x" in exI) + apply(rule_tac x="\n,\c" in exI) + apply(simp) + apply(rule allI)+ + apply(rule impI) + apply(simp add: psubst_nsubst[symmetric]) + apply(drule_tac x="(x,aa,Pa)#\n" in meta_spec) + apply(drule_tac x="\c" in meta_spec) + apply(drule meta_mp) + apply(rule ncloses_extend) + apply(assumption) + apply(assumption) + apply(assumption) + apply(assumption) + apply(drule meta_mp) + apply(rule ccloses_subset) + apply(assumption) + apply(blast) + apply(assumption) + apply(simp) + apply(blast) + done +next + case (TNotL a \ \ M B \' x \n \c) + then show ?case + apply(simp) + apply(subgoal_tac "(x,NOT B) \ set \'") + apply(drule ncloses_elim) + apply(assumption) + apply(erule exE)+ + apply(simp del: NEGc.simps) + apply(generate_fresh "name") + apply(fresh_fun_simp) + apply(rule_tac B="NOT B" in CUT_SNa) + apply(simp) + apply(rule NEG_intro) + apply(simp (no_asm)) + apply(rule disjI2) + apply(rule disjI2) + apply(rule_tac x="a" in exI) + apply(rule_tac x="ca" in exI) + apply(rule_tac x="\n,\c" in exI) + apply(simp del: NEGc.simps) + apply(rule conjI) + apply(rule fin.intros) + apply(rule psubst_fresh_name) + apply(simp) + apply(simp) + apply(simp) + apply(rule BINDING_implies_CAND) + apply(unfold BINDINGc_def) + apply(simp (no_asm)) + apply(rule_tac x="a" in exI) + apply(rule_tac x="\n,\c" in exI) + apply(simp (no_asm)) + apply(rule allI)+ + apply(rule impI) + apply(simp del: NEGc.simps add: psubst_csubst[symmetric]) + apply(drule_tac x="\n" in meta_spec) + apply(drule_tac x="(a,xa,Pa)#\c" in meta_spec) + apply(drule meta_mp) + apply(rule ncloses_subset) + apply(assumption) + apply(blast) + apply(drule meta_mp) + apply(rule ccloses_extend) + apply(assumption) + apply(assumption) + apply(assumption) + apply(assumption) + apply(assumption) + apply(blast) + done +next + case (TAndL1 x \ y B1 M \ \' B2 \n \c) + then show ?case + apply(simp) + apply(subgoal_tac "(y,B1 AND B2) \ set \'") + apply(drule ncloses_elim) + apply(assumption) + apply(erule exE)+ + apply(simp del: NEGc.simps) + apply(generate_fresh "name") + apply(fresh_fun_simp) + apply(rule_tac B="B1 AND B2" in CUT_SNa) + apply(simp) + apply(rule NEG_intro) + apply(simp (no_asm)) + apply(rule disjI2) + apply(rule disjI2) + apply(rule disjI1) + apply(rule_tac x="x" in exI) + apply(rule_tac x="ca" in exI) + apply(rule_tac x="\n,\c" in exI) + apply(simp del: NEGc.simps) + apply(rule conjI) + apply(rule fin.intros) + apply(simp del: NEGc.simps add: abs_fresh fresh_prod fresh_atm) + apply(rule psubst_fresh_name) + apply(simp) + apply(simp) + apply(simp) + apply(rule BINDING_implies_CAND) + apply(unfold BINDINGn_def) + apply(simp (no_asm)) + apply(rule_tac x="x" in exI) + apply(rule_tac x="\n,\c" in exI) + apply(simp (no_asm)) + apply(rule allI)+ + apply(rule impI) + apply(simp del: NEGc.simps add: psubst_nsubst[symmetric]) + apply(drule_tac x="(x,a,Pa)#\n" in meta_spec) + apply(drule_tac x="\c" in meta_spec) + apply(drule meta_mp) + apply(rule ncloses_extend) + apply(rule ncloses_subset) + apply(assumption) + apply(blast) + apply(simp) + apply(simp) + apply(simp) + apply(drule meta_mp) + apply(assumption) + apply(assumption) + apply(blast) + done +next + case (TAndL2 x \ y B2 M \ \' B1 \n \c) + then show ?case + apply(simp) + apply(subgoal_tac "(y,B1 AND B2) \ set \'") + apply(drule ncloses_elim) + apply(assumption) + apply(erule exE)+ + apply(simp del: NEGc.simps) + apply(generate_fresh "name") + apply(fresh_fun_simp) + apply(rule_tac B="B1 AND B2" in CUT_SNa) + apply(simp) + apply(rule NEG_intro) + apply(simp (no_asm)) + apply(rule disjI2) + apply(rule disjI2) + apply(rule disjI2) + apply(rule_tac x="x" in exI) + apply(rule_tac x="ca" in exI) + apply(rule_tac x="\n,\c" in exI) + apply(simp del: NEGc.simps) + apply(rule conjI) + apply(rule fin.intros) + apply(simp del: NEGc.simps add: abs_fresh fresh_prod fresh_atm) + apply(rule psubst_fresh_name) + apply(simp) + apply(simp) + apply(simp) + apply(rule BINDING_implies_CAND) + apply(unfold BINDINGn_def) + apply(simp (no_asm)) + apply(rule_tac x="x" in exI) + apply(rule_tac x="\n,\c" in exI) + apply(simp (no_asm)) + apply(rule allI)+ + apply(rule impI) + apply(simp del: NEGc.simps add: psubst_nsubst[symmetric]) + apply(drule_tac x="(x,a,Pa)#\n" in meta_spec) + apply(drule_tac x="\c" in meta_spec) + apply(drule meta_mp) + apply(rule ncloses_extend) + apply(rule ncloses_subset) + apply(assumption) + apply(blast) + apply(simp) + apply(simp) + apply(simp) + apply(drule meta_mp) + apply(assumption) + apply(assumption) + apply(blast) + done +next + case (TAndR a \ N c b M \ B C \' \n \c) + then show ?case + apply(simp) + apply(subgoal_tac "(c,B AND C) \ set \'") + apply(drule ccloses_elim) + apply(assumption) + apply(erule exE)+ + apply(simp) + apply(generate_fresh "coname") + apply(fresh_fun_simp) + apply(rule_tac B="B AND C" in CUT_SNa) + apply(simp) + apply(rule disjI2) + apply(rule disjI2) + apply(rule_tac x="ca" in exI) + apply(rule_tac x="a" in exI) + apply(rule_tac x="b" in exI) + apply(rule_tac x="\n,\c" in exI) + apply(rule_tac x="\n,\c" in exI) + apply(simp) + apply(rule conjI) + apply(rule fic.intros) + apply(simp add: abs_fresh fresh_prod fresh_atm) + apply(rule psubst_fresh_coname) + apply(simp) + apply(simp) + apply(simp) + apply(simp add: abs_fresh fresh_prod fresh_atm) + apply(rule psubst_fresh_coname) + apply(simp) + apply(simp) + apply(simp) + apply(rule conjI) + apply(rule BINDING_implies_CAND) + apply(unfold BINDINGc_def) + apply(simp) + apply(rule_tac x="a" in exI) + apply(rule_tac x="\n,\c" in exI) + apply(simp) + apply(rule allI)+ + apply(rule impI) + apply(simp add: psubst_csubst[symmetric]) + apply(drule_tac x="\n" in meta_spec) + apply(drule_tac x="(a,xa,Pa)#\c" in meta_spec) + apply(drule meta_mp) + apply(assumption) + apply(drule meta_mp) + apply(rule ccloses_extend) + apply(rule ccloses_subset) + apply(assumption) + apply(blast) + apply(simp) + apply(simp) + apply(assumption) + apply(assumption) + apply(rule BINDING_implies_CAND) + apply(unfold BINDINGc_def) + apply(simp) + apply(rule_tac x="b" in exI) + apply(rule_tac x="\n,\c" in exI) + apply(simp) + apply(rule allI)+ + apply(rule impI) + apply(simp add: psubst_csubst[symmetric]) + apply(rotate_tac 14) + apply(drule_tac x="\n" in meta_spec) + apply(drule_tac x="(b,xa,Pa)#\c" in meta_spec) + apply(drule meta_mp) + apply(assumption) + apply(drule meta_mp) + apply(rule ccloses_extend) + apply(rule ccloses_subset) + apply(assumption) + apply(blast) + apply(simp) + apply(simp) + apply(assumption) + apply(assumption) + apply(simp) + apply(blast) + done +next + case (TOrL x \ N z y M B \ C \' \n \c) + then show ?case + apply(simp) + apply(subgoal_tac "(z,B OR C) \ set \'") + apply(drule ncloses_elim) + apply(assumption) + apply(erule exE)+ + apply(simp del: NEGc.simps) + apply(generate_fresh "name") + apply(fresh_fun_simp) + apply(rule_tac B="B OR C" in CUT_SNa) + apply(simp) + apply(rule NEG_intro) + apply(simp (no_asm)) + apply(rule disjI2) + apply(rule disjI2) + apply(rule_tac x="x" in exI) + apply(rule_tac x="y" in exI) + apply(rule_tac x="ca" in exI) + apply(rule_tac x="\n,\c" in exI) + apply(rule_tac x="\n,\c" in exI) + apply(simp del: NEGc.simps) + apply(rule conjI) + apply(rule fin.intros) + apply(simp del: NEGc.simps add: abs_fresh fresh_prod fresh_atm) + apply(rule psubst_fresh_name) + apply(simp) + apply(simp) + apply(simp) + apply(simp del: NEGc.simps add: abs_fresh fresh_prod fresh_atm) + apply(rule psubst_fresh_name) + apply(simp) + apply(simp) + apply(simp) + apply(rule conjI) + apply(rule BINDING_implies_CAND) + apply(unfold BINDINGn_def) + apply(simp del: NEGc.simps) + apply(rule_tac x="x" in exI) + apply(rule_tac x="\n,\c" in exI) + apply(simp del: NEGc.simps) + apply(rule allI)+ + apply(rule impI) + apply(simp del: NEGc.simps add: psubst_nsubst[symmetric]) + apply(drule_tac x="(x,a,Pa)#\n" in meta_spec) + apply(drule_tac x="\c" in meta_spec) + apply(drule meta_mp) + apply(rule ncloses_extend) + apply(rule ncloses_subset) + apply(assumption) + apply(blast) + apply(simp) + apply(simp) + apply(assumption) + apply(drule meta_mp) + apply(assumption) + apply(assumption) + apply(rule BINDING_implies_CAND) + apply(unfold BINDINGn_def) + apply(simp del: NEGc.simps) + apply(rule_tac x="y" in exI) + apply(rule_tac x="\n,\c" in exI) + apply(simp del: NEGc.simps) + apply(rule allI)+ + apply(rule impI) + apply(simp del: NEGc.simps add: psubst_nsubst[symmetric]) + apply(rotate_tac 14) + apply(drule_tac x="(y,a,Pa)#\n" in meta_spec) + apply(drule_tac x="\c" in meta_spec) + apply(drule meta_mp) + apply(rule ncloses_extend) + apply(rule ncloses_subset) + apply(assumption) + apply(blast) + apply(simp) + apply(simp) + apply(assumption) + apply(drule meta_mp) + apply(assumption) + apply(assumption) + apply(blast) + done +next + case (TOrR1 a \ b \ M B1 \' B2 \n \c) + then show ?case + apply(simp) + apply(subgoal_tac "(b,B1 OR B2) \ set \'") + apply(drule ccloses_elim) + apply(assumption) + apply(erule exE)+ + apply(simp del: NEGc.simps) + apply(generate_fresh "coname") + apply(fresh_fun_simp) + apply(rule_tac B="B1 OR B2" in CUT_SNa) + apply(simp) + apply(rule disjI2) + apply(rule disjI2) + apply(rule disjI1) + apply(rule_tac x="a" in exI) + apply(rule_tac x="c" in exI) + apply(rule_tac x="\n,\c" in exI) + apply(simp) + apply(rule conjI) + apply(rule fic.intros) + apply(simp del: NEGc.simps add: abs_fresh fresh_prod fresh_atm) + apply(rule psubst_fresh_coname) + apply(simp) + apply(simp) + apply(simp) + apply(rule BINDING_implies_CAND) + apply(unfold BINDINGc_def) + apply(simp (no_asm)) + apply(rule_tac x="a" in exI) + apply(rule_tac x="\n,\c" in exI) + apply(simp (no_asm)) + apply(rule allI)+ + apply(rule impI) + apply(simp del: NEGc.simps add: psubst_csubst[symmetric]) + apply(drule_tac x="\n" in meta_spec) + apply(drule_tac x="(a,xa,Pa)#\c" in meta_spec) + apply(drule meta_mp) + apply(assumption) + apply(drule meta_mp) + apply(rule ccloses_extend) + apply(rule ccloses_subset) + apply(assumption) + apply(blast) + apply(simp) + apply(simp) + apply(simp) + apply(assumption) + apply(simp) + apply(blast) + done +next + case (TOrR2 a \ b \ M B2 \' B1 \n \c) + then show ?case + apply(simp) + apply(subgoal_tac "(b,B1 OR B2) \ set \'") + apply(drule ccloses_elim) + apply(assumption) + apply(erule exE)+ + apply(simp del: NEGc.simps) + apply(generate_fresh "coname") + apply(fresh_fun_simp) + apply(rule_tac B="B1 OR B2" in CUT_SNa) + apply(simp) + apply(rule disjI2) + apply(rule disjI2) + apply(rule disjI2) + apply(rule_tac x="a" in exI) + apply(rule_tac x="c" in exI) + apply(rule_tac x="\n,\c" in exI) + apply(simp) + apply(rule conjI) + apply(rule fic.intros) + apply(simp del: NEGc.simps add: abs_fresh fresh_prod fresh_atm) + apply(rule psubst_fresh_coname) + apply(simp) + apply(simp) + apply(simp) + apply(rule BINDING_implies_CAND) + apply(unfold BINDINGc_def) + apply(simp (no_asm)) + apply(rule_tac x="a" in exI) + apply(rule_tac x="\n,\c" in exI) + apply(simp (no_asm)) + apply(rule allI)+ + apply(rule impI) + apply(simp del: NEGc.simps add: psubst_csubst[symmetric]) + apply(drule_tac x="\n" in meta_spec) + apply(drule_tac x="(a,xa,Pa)#\c" in meta_spec) + apply(drule meta_mp) + apply(assumption) + apply(drule meta_mp) + apply(rule ccloses_extend) + apply(rule ccloses_subset) + apply(assumption) + apply(blast) + apply(simp) + apply(simp) + apply(simp) + apply(assumption) + apply(simp) + apply(blast) + done +next + case (TImpL a \ N x \ M y B C \' \n \c) + then show ?case + apply(simp) + apply(subgoal_tac "(y,B IMP C) \ set \'") + apply(drule ncloses_elim) + apply(assumption) + apply(erule exE)+ + apply(simp del: NEGc.simps) + apply(generate_fresh "name") + apply(fresh_fun_simp) + apply(rule_tac B="B IMP C" in CUT_SNa) + apply(simp) + apply(rule NEG_intro) + apply(simp (no_asm)) + apply(rule disjI2) + apply(rule disjI2) + apply(rule_tac x="x" in exI) + apply(rule_tac x="a" in exI) + apply(rule_tac x="ca" in exI) + apply(rule_tac x="\n,\c" in exI) + apply(rule_tac x="\n,\c" in exI) + apply(simp del: NEGc.simps) + apply(rule conjI) + apply(rule fin.intros) + apply(rule psubst_fresh_name) + apply(simp) + apply(simp) + apply(simp) + apply(simp del: NEGc.simps add: abs_fresh fresh_prod fresh_atm) + apply(rule psubst_fresh_name) + apply(simp) + apply(simp) + apply(simp) + apply(rule conjI) + apply(rule BINDING_implies_CAND) + apply(unfold BINDINGc_def) + apply(simp del: NEGc.simps) + apply(rule_tac x="a" in exI) + apply(rule_tac x="\n,\c" in exI) + apply(simp del: NEGc.simps) + apply(rule allI)+ + apply(rule impI) + apply(simp del: NEGc.simps add: psubst_csubst[symmetric]) + apply(drule_tac x="\n" in meta_spec) + apply(drule_tac x="(a,xa,Pa)#\c" in meta_spec) + apply(drule meta_mp) + apply(rule ncloses_subset) + apply(assumption) + apply(blast) + apply(drule meta_mp) + apply(rule ccloses_extend) + apply(assumption) + apply(simp) + apply(simp) + apply(assumption) + apply(assumption) + apply(rule BINDING_implies_CAND) + apply(unfold BINDINGn_def) + apply(simp del: NEGc.simps) + apply(rule_tac x="x" in exI) + apply(rule_tac x="\n,\c" in exI) + apply(simp del: NEGc.simps) + apply(rule allI)+ + apply(rule impI) + apply(simp del: NEGc.simps add: psubst_nsubst[symmetric]) + apply(rotate_tac 12) + apply(drule_tac x="(x,aa,Pa)#\n" in meta_spec) + apply(drule_tac x="\c" in meta_spec) + apply(drule meta_mp) + apply(rule ncloses_extend) + apply(rule ncloses_subset) + apply(assumption) + apply(blast) + apply(simp) + apply(simp) + apply(assumption) + apply(drule meta_mp) + apply(assumption) + apply(assumption) + apply(blast) + done +next + case (TImpR a \ b x \ B M C \' \n \c) + then show ?case + apply(simp) + apply(subgoal_tac "(b,B IMP C) \ set \'") + apply(drule ccloses_elim) + apply(assumption) + apply(erule exE)+ + apply(simp) + apply(generate_fresh "coname") + apply(fresh_fun_simp) + apply(rule_tac B="B IMP C" in CUT_SNa) + apply(simp) + apply(rule disjI2) + apply(rule disjI2) + apply(rule_tac x="x" in exI) + apply(rule_tac x="a" in exI) + apply(rule_tac x="c" in exI) + apply(rule_tac x="\n,\c" in exI) + apply(simp) + apply(rule conjI) + apply(rule fic.intros) + apply(simp add: abs_fresh fresh_prod fresh_atm) + apply(rule psubst_fresh_coname) + apply(simp) + apply(simp) + apply(simp) + apply(rule conjI) + apply(rule allI)+ + apply(rule impI) + apply(simp add: psubst_csubst[symmetric]) + apply(rule BINDING_implies_CAND) + apply(unfold BINDINGn_def) + apply(simp) + apply(rule_tac x="x" in exI) + apply(rule_tac x="\n,((a,z,Pa)#\c)" in exI) + apply(simp) + apply(rule allI)+ + apply(rule impI) + apply(rule_tac t="\n,((a,z,Pa)#\c){x:=.Pb}" and + s="((x,aa,Pb)#\n),((a,z,Pa)#\c)" in subst) + apply(rule psubst_nsubst) + apply(simp add: fresh_prod fresh_atm fresh_list_cons) + apply(drule_tac x="(x,aa,Pb)#\n" in meta_spec) + apply(drule_tac x="(a,z,Pa)#\c" in meta_spec) + apply(drule meta_mp) + apply(rule ncloses_extend) + apply(assumption) + apply(simp) + apply(simp) + apply(simp) + apply(drule meta_mp) + apply(rule ccloses_extend) + apply(rule ccloses_subset) + apply(assumption) + apply(blast) + apply(auto intro: fresh_subset simp del: NEGc.simps)[1] + apply(simp) + apply(simp) + apply(assumption) + apply(rule allI)+ + apply(rule impI) + apply(simp add: psubst_nsubst[symmetric]) + apply(rule BINDING_implies_CAND) + apply(unfold BINDINGc_def) + apply(simp) + apply(rule_tac x="a" in exI) + apply(rule_tac x="((x,ca,Q)#\n),\c" in exI) + apply(simp) + apply(rule allI)+ + apply(rule impI) + apply(rule_tac t="((x,ca,Q)#\n),\c{a:=(xaa).Pa}" and + s="((x,ca,Q)#\n),((a,xaa,Pa)#\c)" in subst) + apply(rule psubst_csubst) + apply(simp add: fresh_prod fresh_atm fresh_list_cons) + apply(drule_tac x="(x,ca,Q)#\n" in meta_spec) + apply(drule_tac x="(a,xaa,Pa)#\c" in meta_spec) + apply(drule meta_mp) + apply(rule ncloses_extend) + apply(assumption) + apply(simp) + apply(simp) + apply(simp) + apply(drule meta_mp) + apply(rule ccloses_extend) + apply(rule ccloses_subset) + apply(assumption) + apply(blast) + apply(auto intro: fresh_subset simp del: NEGc.simps)[1] + apply(simp) + apply(simp) + apply(assumption) + apply(simp) + apply(blast) + done +next + case (TCut a \ N x \ M B \n \c) + then show ?case + apply - + apply(case_tac "\y. M\Ax y a") + apply(case_tac "\c. N\Ax x c") + apply(simp) + apply(rule_tac B="B" in CUT_SNa) + apply(rule BINDING_implies_CAND) + apply(unfold BINDINGc_def) + apply(simp) + apply(rule_tac x="a" in exI) + apply(rule_tac x="\n,\c" in exI) + apply(simp) + apply(rule allI) + apply(rule allI) + apply(rule impI) + apply(simp add: psubst_csubst[symmetric]) (*?*) + apply(drule_tac x="\n" in meta_spec) + apply(drule_tac x="(a,xa,P)#\c" in meta_spec) + apply(drule meta_mp) + apply(assumption) + apply(drule meta_mp) + apply(rule ccloses_extend) + apply(assumption) + apply(assumption) + apply(assumption) + apply(assumption) + apply(assumption) + apply(rule BINDING_implies_CAND) + apply(unfold BINDINGn_def) + apply(simp) + apply(rule_tac x="x" in exI) + apply(rule_tac x="\n,\c" in exI) + apply(simp) + apply(rule allI) + apply(rule allI) + apply(rule impI) + apply(simp add: psubst_nsubst[symmetric]) (*?*) + apply(rotate_tac 11) + apply(drule_tac x="(x,aa,P)#\n" in meta_spec) + apply(drule_tac x="\c" in meta_spec) + apply(drule meta_mp) + apply(rule ncloses_extend) + apply(assumption) + apply(assumption) + apply(assumption) + apply(assumption) + apply(drule_tac meta_mp) + apply(assumption) + apply(assumption) + (* cases at least one axiom *) + apply(simp (no_asm_use)) + apply(erule exE) + apply(simp del: psubst.simps) + apply(drule typing_Ax_elim2) + apply(auto simp add: trm.inject)[1] + apply(rule_tac B="B" in CUT_SNa) + (* left term *) + apply(rule BINDING_implies_CAND) + apply(unfold BINDINGc_def) + apply(simp) + apply(rule_tac x="a" in exI) + apply(rule_tac x="\n,\c" in exI) + apply(simp) + apply(rule allI)+ + apply(rule impI) + apply(drule_tac x="\n" in meta_spec) + apply(drule_tac x="(a,xa,P)#\c" in meta_spec) + apply(drule meta_mp) + apply(assumption) + apply(drule meta_mp) + apply(rule ccloses_extend) + apply(assumption) + apply(assumption) + apply(assumption) + apply(assumption) + apply(simp add: psubst_csubst[symmetric]) (*?*) + (* right term -axiom *) + apply(drule ccloses_elim) + apply(assumption) + apply(erule exE)+ + apply(erule conjE) + apply(frule_tac y="x" in lookupd_cmaps) + apply(drule cmaps_fresh) + apply(assumption) + apply(simp) + apply(subgoal_tac "(x):P[xa\n>x] = (xa):P") + apply(simp) + apply(simp add: ntrm.inject) + apply(simp add: alpha fresh_prod fresh_atm) + apply(rule sym) + apply(rule nrename_swap) + apply(simp) + (* M is axiom *) + apply(simp) + apply(auto)[1] + (* both are axioms *) + apply(rule_tac B="B" in CUT_SNa) + apply(drule typing_Ax_elim1) + apply(drule ncloses_elim) + apply(assumption) + apply(erule exE)+ + apply(erule conjE) + apply(frule_tac a="a" in lookupc_nmaps) + apply(drule_tac a="a" in nmaps_fresh) + apply(assumption) + apply(simp) + apply(subgoal_tac ":P[c\c>a] = :P") + apply(simp) + apply(simp add: ctrm.inject) + apply(simp add: alpha fresh_prod fresh_atm) + apply(rule sym) + apply(rule crename_swap) + apply(simp) + apply(drule typing_Ax_elim2) + apply(drule ccloses_elim) + apply(assumption) + apply(erule exE)+ + apply(erule conjE) + apply(frule_tac y="x" in lookupd_cmaps) + apply(drule cmaps_fresh) + apply(assumption) + apply(simp) + apply(subgoal_tac "(x):P[xa\n>x] = (xa):P") + apply(simp) + apply(simp add: ntrm.inject) + apply(simp add: alpha fresh_prod fresh_atm) + apply(rule sym) + apply(rule nrename_swap) + apply(simp) + (* N is not axioms *) + apply(rule_tac B="B" in CUT_SNa) + (* left term *) + apply(drule typing_Ax_elim1) + apply(drule ncloses_elim) + apply(assumption) + apply(erule exE)+ + apply(erule conjE) + apply(frule_tac a="a" in lookupc_nmaps) + apply(drule_tac a="a" in nmaps_fresh) + apply(assumption) + apply(simp) + apply(subgoal_tac ":P[c\c>a] = :P") + apply(simp) + apply(simp add: ctrm.inject) + apply(simp add: alpha fresh_prod fresh_atm) + apply(rule sym) + apply(rule crename_swap) + apply(simp) + apply(rule BINDING_implies_CAND) + apply(unfold BINDINGn_def) + apply(simp) + apply(rule_tac x="x" in exI) + apply(rule_tac x="\n,\c" in exI) + apply(simp) + apply(rule allI) + apply(rule allI) + apply(rule impI) + apply(simp add: psubst_nsubst[symmetric]) (*?*) + apply(rotate_tac 10) + apply(drule_tac x="(x,aa,P)#\n" in meta_spec) + apply(drule_tac x="\c" in meta_spec) + apply(drule meta_mp) + apply(rule ncloses_extend) + apply(assumption) + apply(assumption) + apply(assumption) + apply(assumption) + apply(drule_tac meta_mp) + apply(assumption) + apply(assumption) + done +qed + +consts + "idn" :: "(name\ty) list\coname\(name\coname\trm) list" +primrec + "idn [] a = []" + "idn (p#\) a = ((fst p),a,Ax (fst p) a)#(idn \ a)" + +consts + "idc" :: "(coname\ty) list\name\(coname\name\trm) list" +primrec + "idc [] x = []" + "idc (p#\) x = ((fst p),x,Ax x (fst p))#(idc \ x)" + +lemma idn_eqvt[eqvt]: + fixes pi1::"name prm" + and pi2::"coname prm" + shows "(pi1\(idn \ a)) = idn (pi1\\) (pi1\a)" + and "(pi2\(idn \ a)) = idn (pi2\\) (pi2\a)" +apply(induct \) +apply(auto) +done + +lemma idc_eqvt[eqvt]: + fixes pi1::"name prm" + and pi2::"coname prm" + shows "(pi1\(idc \ x)) = idc (pi1\\) (pi1\x)" + and "(pi2\(idc \ x)) = idc (pi2\\) (pi2\x)" +apply(induct \) +apply(auto) +done + +lemma ccloses_id: + shows "(idc \ x) ccloses \" +apply(induct \) +apply(auto simp add: ccloses_def) +apply(rule Ax_in_CANDs) +apply(rule Ax_in_CANDs) +done + +lemma ncloses_id: + shows "(idn \ a) ncloses \" +apply(induct \) +apply(auto simp add: ncloses_def) +apply(rule Ax_in_CANDs) +apply(rule Ax_in_CANDs) +done + +lemma fresh_idn: + fixes x::"name" + and a::"coname" + shows "x\\ \ x\idn \ a" + and "a\(\,b) \ a\idn \ b" +apply(induct \) +apply(auto simp add: fresh_list_cons fresh_list_nil fresh_atm fresh_prod) +done + +lemma fresh_idc: + fixes x::"name" + and a::"coname" + shows "x\(\,y) \ x\idc \ y" + and "a\\ \ a\idc \ y" +apply(induct \) +apply(auto simp add: fresh_list_cons fresh_list_nil fresh_atm fresh_prod) +done + +lemma idc_cmaps: + assumes a: "idc \ y cmaps b to Some (x,M)" + shows "M=Ax x b" +using a +apply(induct \) +apply(auto) +apply(case_tac "b=a") +apply(auto) +done + +lemma idn_nmaps: + assumes a: "idn \ a nmaps x to Some (b,M)" + shows "M=Ax x b" +using a +apply(induct \) +apply(auto) +apply(case_tac "aa=x") +apply(auto) +done + +lemma lookup1: + assumes a: "x\(idn \ b)" + shows "lookup x a (idn \ b) \c = lookupa x a \c" +using a +apply(induct \) +apply(auto simp add: fresh_list_cons fresh_prod fresh_atm) +done + +lemma lookup2: + assumes a: "\(x\(idn \ b))" + shows "lookup x a (idn \ b) \c = lookupb x a \c b (Ax x b)" +using a +apply(induct \) +apply(auto simp add: fresh_list_cons fresh_prod fresh_atm fresh_list_nil) +done + +lemma lookup3: + assumes a: "a\(idc \ y)" + shows "lookupa x a (idc \ y) = Ax x a" +using a +apply(induct \) +apply(auto simp add: fresh_list_cons fresh_prod fresh_atm) +done + +lemma lookup4: + assumes a: "\(a\(idc \ y))" + shows "lookupa x a (idc \ y) = Cut .(Ax x a) (y).Ax y a" +using a +apply(induct \) +apply(auto simp add: fresh_list_cons fresh_prod fresh_atm fresh_list_nil) +done + +lemma lookup5: + assumes a: "a\(idc \ y)" + shows "lookupb x a (idc \ y) c P = Cut .P (x).Ax x a" +using a +apply(induct \) +apply(auto simp add: fresh_list_cons fresh_prod fresh_atm fresh_list_nil) +done + +lemma lookup6: + assumes a: "\(a\(idc \ y))" + shows "lookupb x a (idc \ y) c P = Cut .P (y).Ax y a" +using a +apply(induct \) +apply(auto simp add: fresh_list_cons fresh_prod fresh_atm fresh_list_nil) +done + +lemma lookup7: + shows "lookupc x a (idn \ b) = Ax x a" +apply(induct \) +apply(auto simp add: fresh_list_cons fresh_prod fresh_atm fresh_list_nil) +done + +lemma lookup8: + shows "lookupd x a (idc \ y) = Ax x a" +apply(induct \) +apply(auto simp add: fresh_list_cons fresh_prod fresh_atm fresh_list_nil) +done + +lemma id_redu: + shows "(idn \ x),(idc \ a) \\<^isub>a* M" +apply(nominal_induct M avoiding: \ \ x a rule: trm.induct) +apply(auto) +(* Ax *) +apply(case_tac "name\(idn \ x)") +apply(simp add: lookup1) +apply(case_tac "coname\(idc \ a)") +apply(simp add: lookup3) +apply(simp add: lookup4) +apply(rule a_star_trans) +apply(rule a_starI) +apply(rule al_redu) +apply(rule better_LAxR_intro) +apply(rule fic.intros) +apply(simp) +apply(simp add: lookup2) +apply(case_tac "coname\(idc \ a)") +apply(simp add: lookup5) +apply(rule a_star_trans) +apply(rule a_starI) +apply(rule al_redu) +apply(rule better_LAxR_intro) +apply(rule fic.intros) +apply(simp) +apply(simp add: lookup6) +apply(rule a_star_trans) +apply(rule a_starI) +apply(rule al_redu) +apply(rule better_LAxR_intro) +apply(rule fic.intros) +apply(simp) +(* Cut *) +apply(auto simp add: fresh_idn fresh_idc psubst_fresh_name psubst_fresh_coname fresh_atm fresh_prod )[1] +apply(simp add: lookup7 lookup8) +apply(simp add: lookup7 lookup8) +apply(simp add: a_star_Cut) +apply(simp add: lookup7 lookup8) +apply(simp add: a_star_Cut) +apply(simp add: a_star_Cut) +(* NotR *) +apply(simp add: fresh_idn fresh_idc) +apply(case_tac "findc (idc \ a) coname") +apply(simp) +apply(simp add: a_star_NotR) +apply(auto)[1] +apply(generate_fresh "coname") +apply(fresh_fun_simp) +apply(drule idc_cmaps) +apply(simp) +apply(subgoal_tac "c\idn \ x,idc \ a") +apply(rule a_star_trans) +apply(rule a_starI) +apply(rule al_redu) +apply(rule better_LAxR_intro) +apply(rule fic.intros) +apply(assumption) +apply(simp add: crename_fresh) +apply(simp add: a_star_NotR) +apply(rule psubst_fresh_coname) +apply(rule fresh_idn) +apply(simp) +apply(rule fresh_idc) +apply(simp) +apply(simp) +(* NotL *) +apply(simp add: fresh_idn fresh_idc) +apply(case_tac "findn (idn \ x) name") +apply(simp) +apply(simp add: a_star_NotL) +apply(auto)[1] +apply(generate_fresh "name") +apply(fresh_fun_simp) +apply(drule idn_nmaps) +apply(simp) +apply(subgoal_tac "c\idn \ x,idc \ a") +apply(rule a_star_trans) +apply(rule a_starI) +apply(rule al_redu) +apply(rule better_LAxL_intro) +apply(rule fin.intros) +apply(assumption) +apply(simp add: nrename_fresh) +apply(simp add: a_star_NotL) +apply(rule psubst_fresh_name) +apply(rule fresh_idn) +apply(simp) +apply(rule fresh_idc) +apply(simp) +apply(simp) +(* AndR *) +apply(simp add: fresh_idn fresh_idc) +apply(case_tac "findc (idc \ a) coname3") +apply(simp) +apply(simp add: a_star_AndR) +apply(auto)[1] +apply(generate_fresh "coname") +apply(fresh_fun_simp) +apply(drule idc_cmaps) +apply(simp) +apply(subgoal_tac "c\idn \ x,idc \ a") +apply(subgoal_tac "c\idn \ x,idc \ a") +apply(rule a_star_trans) +apply(rule a_starI) +apply(rule al_redu) +apply(rule better_LAxR_intro) +apply(rule fic.intros) +apply(simp add: abs_fresh) +apply(simp add: abs_fresh) +apply(auto simp add: fresh_idn fresh_idc psubst_fresh_name crename_fresh fresh_atm fresh_prod )[1] +apply(rule aux3) +apply(rule crename.simps) +apply(auto simp add: fresh_prod fresh_atm)[1] +apply(rule psubst_fresh_coname) +apply(rule fresh_idn) +apply(simp add: fresh_prod fresh_atm) +apply(rule fresh_idc) +apply(simp) +apply(simp) +apply(auto simp add: fresh_prod fresh_atm)[1] +apply(rule psubst_fresh_coname) +apply(rule fresh_idn) +apply(simp add: fresh_prod fresh_atm) +apply(rule fresh_idc) +apply(simp) +apply(simp) +apply(simp) +apply(simp) +apply(simp add: crename_fresh) +apply(simp add: a_star_AndR) +apply(rule psubst_fresh_coname) +apply(rule fresh_idn) +apply(simp) +apply(rule fresh_idc) +apply(simp) +apply(simp) +apply(rule psubst_fresh_coname) +apply(rule fresh_idn) +apply(simp) +apply(rule fresh_idc) +apply(simp) +apply(simp) +(* AndL1 *) +apply(simp add: fresh_idn fresh_idc) +apply(case_tac "findn (idn \ x) name2") +apply(simp) +apply(simp add: a_star_AndL1) +apply(auto)[1] +apply(generate_fresh "name") +apply(fresh_fun_simp) +apply(drule idn_nmaps) +apply(simp) +apply(subgoal_tac "c\idn \ x,idc \ a") +apply(rule a_star_trans) +apply(rule a_starI) +apply(rule al_redu) +apply(rule better_LAxL_intro) +apply(rule fin.intros) +apply(simp add: abs_fresh) +apply(rule aux3) +apply(rule nrename.simps) +apply(auto simp add: fresh_prod fresh_atm)[1] +apply(simp) +apply(simp add: nrename_fresh) +apply(simp add: a_star_AndL1) +apply(rule psubst_fresh_name) +apply(rule fresh_idn) +apply(simp) +apply(rule fresh_idc) +apply(simp) +apply(simp) +(* AndL2 *) +apply(simp add: fresh_idn fresh_idc) +apply(case_tac "findn (idn \ x) name2") +apply(simp) +apply(simp add: a_star_AndL2) +apply(auto)[1] +apply(generate_fresh "name") +apply(fresh_fun_simp) +apply(drule idn_nmaps) +apply(simp) +apply(subgoal_tac "c\idn \ x,idc \ a") +apply(rule a_star_trans) +apply(rule a_starI) +apply(rule al_redu) +apply(rule better_LAxL_intro) +apply(rule fin.intros) +apply(simp add: abs_fresh) +apply(rule aux3) +apply(rule nrename.simps) +apply(auto simp add: fresh_prod fresh_atm)[1] +apply(simp) +apply(simp add: nrename_fresh) +apply(simp add: a_star_AndL2) +apply(rule psubst_fresh_name) +apply(rule fresh_idn) +apply(simp) +apply(rule fresh_idc) +apply(simp) +apply(simp) +(* OrR1 *) +apply(simp add: fresh_idn fresh_idc) +apply(case_tac "findc (idc \ a) coname2") +apply(simp) +apply(simp add: a_star_OrR1) +apply(auto)[1] +apply(generate_fresh "coname") +apply(fresh_fun_simp) +apply(drule idc_cmaps) +apply(simp) +apply(subgoal_tac "c\idn \ x,idc \ a") +apply(rule a_star_trans) +apply(rule a_starI) +apply(rule al_redu) +apply(rule better_LAxR_intro) +apply(rule fic.intros) +apply(simp add: abs_fresh) +apply(rule aux3) +apply(rule crename.simps) +apply(auto simp add: fresh_prod fresh_atm)[1] +apply(simp) +apply(simp add: crename_fresh) +apply(simp add: a_star_OrR1) +apply(rule psubst_fresh_coname) +apply(rule fresh_idn) +apply(simp) +apply(rule fresh_idc) +apply(simp) +apply(simp) +(* OrR2 *) +apply(simp add: fresh_idn fresh_idc) +apply(case_tac "findc (idc \ a) coname2") +apply(simp) +apply(simp add: a_star_OrR2) +apply(auto)[1] +apply(generate_fresh "coname") +apply(fresh_fun_simp) +apply(drule idc_cmaps) +apply(simp) +apply(subgoal_tac "c\idn \ x,idc \ a") +apply(rule a_star_trans) +apply(rule a_starI) +apply(rule al_redu) +apply(rule better_LAxR_intro) +apply(rule fic.intros) +apply(simp add: abs_fresh) +apply(rule aux3) +apply(rule crename.simps) +apply(auto simp add: fresh_prod fresh_atm)[1] +apply(simp) +apply(simp add: crename_fresh) +apply(simp add: a_star_OrR2) +apply(rule psubst_fresh_coname) +apply(rule fresh_idn) +apply(simp) +apply(rule fresh_idc) +apply(simp) +apply(simp) +(* OrL *) +apply(simp add: fresh_idn fresh_idc) +apply(case_tac "findn (idn \ x) name3") +apply(simp) +apply(simp add: a_star_OrL) +apply(auto)[1] +apply(generate_fresh "name") +apply(fresh_fun_simp) +apply(drule idn_nmaps) +apply(simp) +apply(subgoal_tac "c\idn \ x,idc \ a") +apply(subgoal_tac "c\idn \ x,idc \ a") +apply(rule a_star_trans) +apply(rule a_starI) +apply(rule al_redu) +apply(rule better_LAxL_intro) +apply(rule fin.intros) +apply(simp add: abs_fresh) +apply(simp add: abs_fresh) +apply(rule aux3) +apply(rule nrename.simps) +apply(auto simp add: fresh_prod fresh_atm)[1] +apply(rule psubst_fresh_name) +apply(rule fresh_idn) +apply(simp) +apply(rule fresh_idc) +apply(simp add: fresh_prod fresh_atm) +apply(simp) +apply(auto simp add: fresh_prod fresh_atm)[1] +apply(rule psubst_fresh_name) +apply(rule fresh_idn) +apply(simp) +apply(rule fresh_idc) +apply(simp add: fresh_prod fresh_atm) +apply(simp) +apply(simp) +apply(simp) +apply(simp add: nrename_fresh) +apply(simp add: a_star_OrL) +apply(rule psubst_fresh_name) +apply(rule fresh_idn) +apply(simp) +apply(rule fresh_idc) +apply(simp) +apply(simp) +apply(rule psubst_fresh_name) +apply(rule fresh_idn) +apply(simp) +apply(rule fresh_idc) +apply(simp) +apply(simp) +(* ImpR *) +apply(simp add: fresh_idn fresh_idc) +apply(case_tac "findc (idc \ a) coname2") +apply(simp) +apply(simp add: a_star_ImpR) +apply(auto)[1] +apply(generate_fresh "coname") +apply(fresh_fun_simp) +apply(drule idc_cmaps) +apply(simp) +apply(subgoal_tac "c\idn \ x,idc \ a") +apply(rule a_star_trans) +apply(rule a_starI) +apply(rule al_redu) +apply(rule better_LAxR_intro) +apply(rule fic.intros) +apply(simp add: abs_fresh) +apply(rule aux3) +apply(rule crename.simps) +apply(auto simp add: fresh_prod fresh_atm)[1] +apply(simp) +apply(simp add: crename_fresh) +apply(simp add: a_star_ImpR) +apply(rule psubst_fresh_coname) +apply(rule fresh_idn) +apply(simp) +apply(rule fresh_idc) +apply(simp) +apply(simp) +(* ImpL *) +apply(simp add: fresh_idn fresh_idc) +apply(case_tac "findn (idn \ x) name2") +apply(simp) +apply(simp add: a_star_ImpL) +apply(auto)[1] +apply(generate_fresh "name") +apply(fresh_fun_simp) +apply(drule idn_nmaps) +apply(simp) +apply(subgoal_tac "c\idn \ x,idc \ a") +apply(subgoal_tac "c\idn \ x,idc \ a") +apply(rule a_star_trans) +apply(rule a_starI) +apply(rule al_redu) +apply(rule better_LAxL_intro) +apply(rule fin.intros) +apply(simp add: abs_fresh) +apply(simp add: abs_fresh) +apply(rule aux3) +apply(rule nrename.simps) +apply(auto simp add: fresh_prod fresh_atm)[1] +apply(rule psubst_fresh_coname) +apply(rule fresh_idn) +apply(simp add: fresh_atm) +apply(rule fresh_idc) +apply(simp add: fresh_prod fresh_atm) +apply(simp) +apply(auto simp add: fresh_prod fresh_atm)[1] +apply(rule psubst_fresh_name) +apply(rule fresh_idn) +apply(simp) +apply(rule fresh_idc) +apply(simp add: fresh_prod fresh_atm) +apply(simp) +apply(simp) +apply(simp add: nrename_fresh) +apply(simp add: a_star_ImpL) +apply(rule psubst_fresh_name) +apply(rule fresh_idn) +apply(simp) +apply(rule fresh_idc) +apply(simp) +apply(simp) +apply(rule psubst_fresh_name) +apply(rule fresh_idn) +apply(simp) +apply(rule fresh_idc) +apply(simp) +apply(simp) +done + +theorem ALL_SNa: + assumes a: "\ \ M \ \" + shows "SNa M" +proof - + have "(idc \ x) ccloses \" by (simp add: ccloses_id) + moreover + have "(idn \ a) ncloses \" by (simp add: ncloses_id) + ultimately have "SNa ((idn \ a),(idc \ x))" using a by (simp add: all_CAND) + moreover + have "((idn \ a),(idc \ x)) \\<^isub>a* M" by (simp add: id_redu) + ultimately show "SNa M" by (simp add: a_star_preserves_SNa) +qed end