# HG changeset patch # User wenzelm # Date 1714687671 -7200 # Node ID d9b8831a6a998150aa5570efca31a6eadafa50ab # Parent 5e64a54f6790beb22ec061af7b88c3576adac542# Parent 96f60533ec1da892b07f957eeafed0597acbaa43 merged, resoving conflicts in src/HOL/Nominal/Nominal.thy; diff -r 5e64a54f6790 -r d9b8831a6a99 Admin/Release/CHECKLIST --- a/Admin/Release/CHECKLIST Thu May 02 15:40:05 2024 +0200 +++ b/Admin/Release/CHECKLIST Fri May 03 00:07:51 2024 +0200 @@ -27,8 +27,7 @@ - test "isabelle build HOL-Codegenerator_Test" with -new-syntax (see also etc/settings:ISABELLE_SCALAC_OPTIONS); -- test Windows 10 subsystem for Linux: - https://docs.microsoft.com/en-us/windows/wsl/install-win10 +- test Windows Subsystem for Linux (e.g. on Windows Server 2022); - check (non-)executable files: $ find . "(" -name "*.thy" -o -name "*.ML" -o -name "*.scala" -o -name ROOT ")" -executable diff -r 5e64a54f6790 -r d9b8831a6a99 Admin/Windows/Cygwin/README --- a/Admin/Windows/Cygwin/README Thu May 02 15:40:05 2024 +0200 +++ b/Admin/Windows/Cygwin/README Fri May 03 00:07:51 2024 +0200 @@ -71,11 +71,6 @@ - run ssh-host-config ``` StrictMode: yes - privilege separation: yes - new local account 'sshd': yes - sshd as a service: yes - different name (than cyg_server): no - new privileged account cyg_server: yes ``` - user setup @@ -85,3 +80,8 @@ passwd -R USER ``` + + - firewall +``` + netsh advfirewall firewall add rule name="Open SSH Port 22" dir=in action=allow protocol=TCP localport=22 remoteip=any +``` diff -r 5e64a54f6790 -r d9b8831a6a99 Admin/components/README.md --- a/Admin/components/README.md Thu May 02 15:40:05 2024 +0200 +++ b/Admin/components/README.md Fri May 03 00:07:51 2024 +0200 @@ -45,7 +45,7 @@ - macOS 14 Sonoma (`studio1` Mac13,2 M1 Ultra, 16+4 cores) * `x86_64-windows` - - **Windows Server 2012 Rev 2** (`vmnipkow9`) + - **Windows Server 2022** (`se0.proof.cit.tum.de`) - **Windows 10** - Windows 11 * `x86_64-cygwin` diff -r 5e64a54f6790 -r d9b8831a6a99 Admin/polyml/INSTALL-MinGW --- a/Admin/polyml/INSTALL-MinGW Thu May 02 15:40:05 2024 +0200 +++ b/Admin/polyml/INSTALL-MinGW Fri May 03 00:07:51 2024 +0200 @@ -14,8 +14,7 @@ pacman -S --needed --noconfirm base-devel gmp-devel mingw-w64-x86_64-toolchain mingw-w64-x86_64-gmp mingw-w64-x86_64-lapack mingw-w64-x86_64-openblas -- build (as regular user) e.g. on vmnipkow9 with Cygwin-Terminal from Isabelle2018 +- build (as regular user) e.g. with Cygwin-Terminal from Isabelle2018 (to avoid address relocation problems): - isabelle component_polyml -M /cygdrive/c/msys64 -m32 -s sha1 src - isabelle component_polyml -M /cygdrive/c/msys64 -m64 -s sha1 src + isabelle component_polyml -M /cygdrive/c/msys64 diff -r 5e64a54f6790 -r d9b8831a6a99 CONTRIBUTORS --- a/CONTRIBUTORS Thu May 02 15:40:05 2024 +0200 +++ b/CONTRIBUTORS Fri May 03 00:07:51 2024 +0200 @@ -3,6 +3,10 @@ listed as an author in one of the source files of this Isabelle distribution. +Contributions to this Isabelle version +-------------------------------------- + + Contributions to Isabelle2024 ----------------------------- diff -r 5e64a54f6790 -r d9b8831a6a99 NEWS --- a/NEWS Thu May 02 15:40:05 2024 +0200 +++ b/NEWS Fri May 03 00:07:51 2024 +0200 @@ -4,6 +4,11 @@ (Note: Isabelle/jEdit shows a tree-view of the NEWS file in Sidekick.) +New in this Isabelle version +---------------------------- + + + New in Isabelle2024 (May 2024) ------------------------------ diff -r 5e64a54f6790 -r d9b8831a6a99 src/HOL/Nominal/Examples/Class1.thy --- a/src/HOL/Nominal/Examples/Class1.thy Thu May 02 15:40:05 2024 +0200 +++ b/src/HOL/Nominal/Examples/Class1.thy Fri May 03 00:07:51 2024 +0200 @@ -1,5 +1,5 @@ theory Class1 -imports "HOL-Nominal.Nominal" +imports "HOL-Nominal.Nominal" begin section \Term-Calculus from Urban's PhD\ @@ -44,9 +44,9 @@ and T::"ty" shows "a\T" and "x\T" by (nominal_induct T rule: ty.strong_induct) - (auto simp add: fresh_string) - -text \terms\ + (auto simp: fresh_string) + +text \terms\ nominal_datatype trm = Ax "name" "coname" @@ -89,11 +89,7 @@ | "a\(d,e,b) \ (ImpR (x)..M b)[d\c>e] = (if b=d then ImpR (x)..(M[d\c>e]) e else ImpR (x)..(M[d\c>e]) b)" | "\a\(d,e,N);x\(M,y)\ \ (ImpL .M (x).N y)[d\c>e] = ImpL .(M[d\c>e]) (x).(N[d\c>e]) y" -apply(finite_guess)+ -apply(rule TrueI)+ -apply(simp add: abs_fresh abs_supp fin_supp)+ -apply(fresh_guess)+ -done + by(finite_guess | simp add: abs_fresh abs_supp fin_supp | fresh_guess)+ nominal_primrec (freshness_context: "(u::name,v::name)") nrename :: "trm \ name \ name \ trm" ("_[_\n>_]" [100,100,100] 100) @@ -112,41 +108,29 @@ | "\a\b; x\(u,v)\ \ (ImpR (x)..M b)[u\n>v] = ImpR (x)..(M[u\n>v]) b" | "\a\N;x\(u,v,M,y)\ \ (ImpL .M (x).N y)[u\n>v] = (if y=u then ImpL .(M[u\n>v]) (x).(N[u\n>v]) v else ImpL .(M[u\n>v]) (x).(N[u\n>v]) y)" -apply(finite_guess)+ -apply(rule TrueI)+ -apply(simp add: abs_fresh abs_supp fs_name1 fs_coname1)+ -apply(fresh_guess)+ -done + by(finite_guess | simp add: abs_fresh abs_supp fin_supp | fresh_guess)+ 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)]" -apply(nominal_induct M avoiding: d e rule: trm.strong_induct) -apply(auto simp add: fresh_bij eq_bij) -done + by (nominal_induct M avoiding: d e rule: trm.strong_induct) (auto simp: fresh_bij eq_bij) lemma crename_coname_eqvt[eqvt]: fixes pi::"coname prm" shows "pi\(M[d\c>e]) = (pi\M)[(pi\d)\c>(pi\e)]" -apply(nominal_induct M avoiding: d e rule: trm.strong_induct) -apply(auto simp add: fresh_bij eq_bij) -done + by (nominal_induct M avoiding: d e rule: trm.strong_induct)(auto simp: fresh_bij eq_bij) lemma nrename_name_eqvt[eqvt]: fixes pi::"name prm" shows "pi\(M[x\n>y]) = (pi\M)[(pi\x)\n>(pi\y)]" -apply(nominal_induct M avoiding: x y rule: trm.strong_induct) -apply(auto simp add: fresh_bij eq_bij) -done + by(nominal_induct M avoiding: x y rule: trm.strong_induct) (auto simp: fresh_bij eq_bij) lemma nrename_coname_eqvt[eqvt]: fixes pi::"coname prm" shows "pi\(M[x\n>y]) = (pi\M)[(pi\x)\n>(pi\y)]" -apply(nominal_induct M avoiding: x y rule: trm.strong_induct) -apply(auto simp add: fresh_bij eq_bij) -done + by(nominal_induct M avoiding: x y rule: trm.strong_induct)(auto simp: fresh_bij eq_bij) lemmas rename_eqvts = crename_name_eqvt crename_coname_eqvt nrename_name_eqvt nrename_coname_eqvt @@ -155,58 +139,57 @@ shows "M[x\n>y] = M" using a by (nominal_induct M avoiding: x y rule: trm.strong_induct) - (auto simp add: trm.inject fresh_atm abs_fresh fin_supp abs_supp) + (auto simp: 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.strong_induct) - (auto simp add: trm.inject fresh_atm abs_fresh) + (auto simp: 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.strong_induct) - (auto simp add: fresh_atm abs_fresh abs_supp fin_supp) + (auto simp: 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.strong_induct) - (auto simp add: fresh_atm abs_fresh abs_supp fin_supp) + (auto simp: 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.strong_induct) - (auto simp add: fresh_atm abs_fresh abs_supp fin_supp) + (auto simp: 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.strong_induct) - (auto simp add: fresh_atm abs_fresh abs_supp fin_supp) + (auto simp: 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.strong_induct) - (auto simp add: fresh_prod fresh_atm abs_fresh abs_supp fin_supp) + (auto simp: 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.strong_induct) - (auto simp add: fresh_prod fresh_atm abs_fresh abs_supp fin_supp) + (auto simp: 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.strong_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) + apply(auto simp: abs_fresh abs_supp fin_supp fresh_left calc_atm fresh_atm) done lemma crename_rename: @@ -214,8 +197,7 @@ shows "([(a',a)]\M)[a'\c>b]= M[a\c>b]" using a apply(nominal_induct M avoiding: a a' b rule: trm.strong_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) + apply(auto simp: abs_fresh abs_supp fin_supp fresh_left calc_atm fresh_atm) done lemmas rename_fresh = nrename_fresh crename_fresh @@ -230,21 +212,13 @@ 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) + using fs1 fs2 by (rule_tac sym, auto simp: 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 + by (intro nrename.simps; simp add: fresh_left calc_atm) 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 + by(simp add: trm.inject alpha fresh_prod rename_eqvts calc_atm rename_fresh fresh_atm) finally show "(Cut .M (x).N)[u\n>v] = Cut .(M[u\n>v]) (x).(N[u\n>v])" using eq1 by simp qed @@ -256,21 +230,13 @@ 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) + using fs1 fs2 by (rule_tac sym, auto simp: 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 + by (intro crename.simps; simp add: fresh_left calc_atm) 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 + by(simp add: trm.inject alpha calc_atm rename_fresh fresh_atm fresh_prod rename_eqvts) finally show "(Cut .M (x).N)[b\c>c] = Cut .(M[b\c>c]) (x).(N[b\c>c])" using eq1 by simp qed @@ -286,12 +252,14 @@ lemma nrename_swap: shows "x\M \ [(x,y)]\M = M[y\n>x]" by (nominal_induct M avoiding: x y rule: trm.strong_induct) - (simp_all add: calc_atm fresh_atm trm.inject alpha abs_fresh abs_supp fin_supp) + (auto simp: abs_fresh abs_supp fin_supp fresh_left calc_atm fresh_atm) + lemma crename_swap: shows "a\M \ [(a,b)]\M = M[b\c>a]" by (nominal_induct M avoiding: a b rule: trm.strong_induct) - (simp_all add: calc_atm fresh_atm trm.inject alpha abs_fresh abs_supp fin_supp) + (auto simp: abs_fresh abs_supp fin_supp fresh_left calc_atm fresh_atm) + lemma crename_ax: assumes a: "M[a\c>b] = Ax x c" "c\a" "c\b" @@ -317,11 +285,7 @@ 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 + by (simp add: assms fresh_perm_app) lemma fresh_perm_name: fixes x::"name" @@ -329,30 +293,29 @@ 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 + by (simp add: assms fresh_perm_app) 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 +proof (rule fresh_fun_app) + show "pt (TYPE(trm)::trm itself) (TYPE(name)::name itself)" + by(rule pt_name_inst) + show "at (TYPE(name)::name itself)" + by(rule at_name_inst) + show "finite (supp (\x'. Cut .P (x').NotL .M x')::name set)" + using a by(finite_guess) + obtain n::name where n: "n\(c,P,a,M)" + by (metis assms fresh_atm(3) fresh_prod) + with assms + show "\b. b \ (\x'. Cut .P (x').NotL .M x', Cut .P (b).NotL .M b)" + apply(intro exI [where x="n"]) + apply(simp add: fresh_prod abs_fresh) + apply(fresh_guess) + done + show "x' \ (\x'. Cut .P (x').NotL .M x')" + using assms by(fresh_guess) +qed lemma fresh_fun_NotL[eqvt_force]: fixes pi1::"name prm" @@ -361,48 +324,36 @@ 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] swap_simps) -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 + apply(perm_simp) + apply(generate_fresh "name") + apply(auto simp: fresh_prod fresh_fun_simp_NotL) + apply(rule sym) + apply(rule trans) + apply(rule fresh_fun_simp_NotL) + apply(blast intro: fresh_perm_name) + apply(metis fresh_perm_name) + apply(simp add: at_prm_fresh[OF at_name_inst] swap_simps) + + apply(perm_simp) + apply(subgoal_tac "\n::name. n\(P,M,pi2\P,pi2\M,pi2)") + apply (metis fresh_fun_simp_NotL fresh_prodD swap_simps(8) trm.perm(14) trm.perm(16)) + by (meson exists_fresh(1) fs_name1) 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 + shows "fresh_fun (\z'. Cut .P(z').AndL1 (x).M z') = Cut .P (z').AndL1 (x).M z'" +proof (rule fresh_fun_app [OF pt_name_inst at_name_inst]) + obtain n::name where "n\(c,P,x,M)" + by (meson exists_fresh(1) fs_name1) + then show "\a. a \ (\z'. Cut .P(z').AndL1 x. M z', Cut .P(a).AndL1 x. M a)" + apply(intro exI [where x="n"]) + apply(simp add: fresh_prod abs_fresh) + apply(fresh_guess) + done +next + show "z' \ (\z'. Cut .P(z').AndL1 x. M z')" + using a by(fresh_guess) +qed finite_guess lemma fresh_fun_AndL1[eqvt_force]: fixes pi1::"name prm" @@ -411,41 +362,30 @@ 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] swap_simps) -apply(rule exists_fresh') -apply(simp add: fin_supp) +apply(force simp add: fresh_prod fresh_fun_simp_AndL1 at_prm_fresh[OF at_name_inst] swap_simps) + apply (meson exists_fresh(1) fs_name1) 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 +apply(force simp add: fresh_prod fresh_fun_simp_AndL1 calc_atm) + by (meson exists_fresh'(1) fs_name1) 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 +proof (rule fresh_fun_app [OF pt_name_inst at_name_inst]) + obtain n::name where "n\(c,P,x,M)" + by (meson exists_fresh(1) fs_name1) + then show "\a. a \ (\z'. Cut .P(z').AndL2 x. M z', Cut .P(a).AndL2 x. M a)" + apply(intro exI [where x="n"]) + apply(simp add: fresh_prod abs_fresh) + apply(fresh_guess) + done +next + show "z' \ (\z'. Cut .P(z').AndL2 x. M z')" + using a by(fresh_guess) +qed finite_guess lemma fresh_fun_AndL2[eqvt_force]: fixes pi1::"name prm" @@ -454,84 +394,62 @@ 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] swap_simps) -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 + apply(perm_simp) + apply(subgoal_tac "\n::name. n\(P,M,x,pi1\P,pi1\M,pi1\x,pi1)") + apply(force simp add: fresh_prod fresh_fun_simp_AndL2 at_prm_fresh[OF at_name_inst] swap_simps) + apply (meson exists_fresh(1) fs_name1) + apply(perm_simp) + apply(subgoal_tac "\n::name. n\(P,M,x,pi2\P,pi2\M,pi2\x,pi2)") + apply(force simp add: fresh_prod fresh_fun_simp_AndL2 calc_atm) + by (meson exists_fresh(1) fs_name1) 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 + shows "fresh_fun (\z'. Cut .P(z').OrL(x).M(u).N z') = Cut .P (z').OrL (x).M (u).N z'" +proof (rule fresh_fun_app [OF pt_name_inst at_name_inst]) + obtain n::name where "n\(c,P,x,M,u,N)" + by (meson exists_fresh(1) fs_name1) + then show "\a. a \ (\z'. Cut .P(z').OrL(x).M(u).N(z'), Cut .P(a).OrL(x).M(u).N(a))" + apply(intro exI [where x="n"]) + apply(simp add: fresh_prod abs_fresh) + apply(fresh_guess) + done +next + show "z' \ (\z'. Cut .P(z').OrL(x).M(u).N(z'))" + using a by(fresh_guess) +qed finite_guess 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] swap_simps) -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 + 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(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(force simp add: fresh_prod fresh_fun_simp_OrL at_prm_fresh[OF at_name_inst] swap_simps) + apply (meson exists_fresh(1) fs_name1) + 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(force simp add: fresh_prod fresh_fun_simp_OrL calc_atm) + by (meson exists_fresh(1) fs_name1) 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 +proof (rule fresh_fun_app [OF pt_name_inst at_name_inst]) + obtain n::name where "n\(c,P,x,M,N)" + by (meson exists_fresh(1) fs_name1) + then show "\aa. aa \ (\z'. Cut .P(z').ImpL .M(x).N z', Cut .P(aa).ImpL .M(x).N aa)" + apply(intro exI [where x="n"]) + apply(simp add: fresh_prod abs_fresh) + apply(fresh_guess) + done +next + show "z' \ (\z'. Cut .P(z').ImpL .M(x).N z')" + using a by(fresh_guess) +qed finite_guess lemma fresh_fun_ImpL[eqvt_force]: fixes pi1::"name prm" @@ -540,41 +458,27 @@ 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] swap_simps) -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 + apply(perm_simp) + apply(subgoal_tac "\n::name. n\(P,M,N,x,pi1\P,pi1\M,pi1\N,pi1\x,pi1)") + apply(force simp add: fresh_prod fresh_fun_simp_ImpL at_prm_fresh[OF at_name_inst] swap_simps) + apply (meson exists_fresh(1) fs_name1) + apply(perm_simp) + apply(subgoal_tac "\n::name. n\(P,M,N,x,pi2\P,pi2\M,pi2\N,pi2\x,pi2)") + apply(force simp add: fresh_prod fresh_fun_simp_ImpL calc_atm) + by (meson exists_fresh(1) fs_name1) 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 +proof (rule fresh_fun_app [OF pt_coname_inst at_coname_inst]) + obtain n::coname where "n\(x,P,y,M)" + by (metis assms(1) assms(2) fresh_atm(4) fresh_prod) + then show "\a. a \ (\a'. Cut .(NotR (y).M a') (x).P, Cut .NotR(y).M(a) (x).P)" + apply(intro exI [where x="n"]) + apply(simp add: fresh_prod abs_fresh) + apply(fresh_guess) + done +qed (use a in \fresh_guess|finite_guess\)+ lemma fresh_fun_NotR[eqvt_force]: fixes pi1::"name prm" @@ -583,41 +487,27 @@ 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] swap_simps) -apply(rule exists_fresh') -apply(simp add: fin_supp) -done + apply(perm_simp) + apply(subgoal_tac "\n::coname. n\(P,M,pi1\P,pi1\M,pi1)") + apply(force simp add: fresh_prod fresh_fun_simp_NotR calc_atm) + apply (meson exists_fresh(2) fs_coname1) + apply(perm_simp) + apply(subgoal_tac "\n::coname. n\(P,M,pi2\P,pi2\M,pi2)") + apply(force simp: fresh_prod fresh_fun_simp_NotR at_prm_fresh[OF at_coname_inst] swap_simps) + by (meson exists_fresh(2) fs_coname1) 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 +proof (rule fresh_fun_app [OF pt_coname_inst at_coname_inst]) + obtain n::coname where "n\(x,P,b,M,c,N)" + by (meson exists_fresh(2) fs_coname1) + then show "\a. a \ (\a'. Cut .AndR .M .N(a') (x).P, Cut .AndR .M .N(a) (x).P)" + apply(intro exI [where x="n"]) + apply(simp add: fresh_prod abs_fresh) + apply(fresh_guess) + done +qed (use a in \fresh_guess|finite_guess\)+ lemma fresh_fun_AndR[eqvt_force]: fixes pi1::"name prm" @@ -626,41 +516,27 @@ 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] swap_simps) -apply(rule exists_fresh') -apply(simp add: fin_supp) -done + 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(force simp add: fresh_prod fresh_fun_simp_AndR calc_atm) + apply (meson exists_fresh(2) fs_coname1) + 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(force simp add: fresh_prod fresh_fun_simp_AndR at_prm_fresh[OF at_coname_inst] swap_simps) + by (meson exists_fresh(2) fs_coname1) 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 +proof (rule fresh_fun_app [OF pt_coname_inst at_coname_inst]) + obtain n::coname where "n\(x,P,b,M)" + by (meson exists_fresh(2) fs_coname1) + then show "\a. a \ (\a'. Cut .OrR1 .M(a') (x).P, Cut .OrR1 .M(a) (x).P)" + apply(intro exI [where x="n"]) + apply(simp add: fresh_prod abs_fresh) + apply(fresh_guess) + done +qed (use a in \fresh_guess|finite_guess\)+ lemma fresh_fun_OrR1[eqvt_force]: fixes pi1::"name prm" @@ -669,41 +545,27 @@ 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] swap_simps) -apply(rule exists_fresh') -apply(simp add: fin_supp) -done + apply(perm_simp) + apply(subgoal_tac "\n::coname. n\(P,M,b,pi1\P,pi1\M,pi1\b,pi1)") + apply(force simp add: fresh_prod fresh_fun_simp_OrR1 calc_atm) + apply (meson exists_fresh(2) fs_coname1) + apply(perm_simp) + apply(subgoal_tac "\n::coname. n\(P,M,b,pi2\P,pi2\M,pi2\b,pi2)") + apply(force simp add: fresh_prod fresh_fun_simp_OrR1 at_prm_fresh[OF at_coname_inst] swap_simps) + by (meson exists_fresh(2) fs_coname1) 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 +proof (rule fresh_fun_app [OF pt_coname_inst at_coname_inst]) + obtain n::coname where "n\(x,P,b,M)" + by (meson exists_fresh(2) fs_coname1) + then show "\a. a \ (\a'. Cut .OrR2 .M(a') (x).P, Cut .OrR2 .M(a) (x).P)" + apply(intro exI [where x="n"]) + apply(simp add: fresh_prod abs_fresh) + apply(fresh_guess) + done +qed (use a in \fresh_guess|finite_guess\)+ lemma fresh_fun_OrR2[eqvt_force]: fixes pi1::"name prm" @@ -712,41 +574,27 @@ 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] swap_simps) -apply(rule exists_fresh') -apply(simp add: fin_supp) -done + apply(perm_simp) + apply(subgoal_tac "\n::coname. n\(P,M,b,pi1\P,pi1\M,pi1\b,pi1)") + apply(force simp add: fresh_prod fresh_fun_simp_OrR2 calc_atm) + apply (meson exists_fresh(2) fs_coname1) + apply(perm_simp) + apply(subgoal_tac "\n::coname. n\(P,M,b,pi2\P,pi2\M,pi2\b,pi2)") + apply(force simp add: fresh_prod fresh_fun_simp_OrR2 at_prm_fresh[OF at_coname_inst] swap_simps) + by (meson exists_fresh(2) fs_coname1) 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 +proof (rule fresh_fun_app [OF pt_coname_inst at_coname_inst]) + obtain n::coname where "n\(x,P,y,b,M)" + by (meson exists_fresh(2) fs_coname1) + then show "\a. a \ (\a'. Cut .(ImpR (y)..M a') (x).P, Cut .(ImpR (y)..M a) (x).P)" + apply(intro exI [where x="n"]) + apply(simp add: fresh_prod abs_fresh) + apply(fresh_guess) + done +qed (use a in \fresh_guess|finite_guess\)+ lemma fresh_fun_ImpR[eqvt_force]: fixes pi1::"name prm" @@ -755,22 +603,14 @@ 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] swap_simps) -apply(rule exists_fresh') -apply(simp add: fin_supp) -done + apply(perm_simp) + apply(subgoal_tac "\n::coname. n\(P,M,b,pi1\P,pi1\M,pi1\b,pi1)") + apply(force simp add: fresh_prod fresh_fun_simp_ImpR calc_atm) + apply (meson exists_fresh(2) fs_coname1) + apply(perm_simp) + apply(subgoal_tac "\n::coname. n\(P,M,b,pi2\P,pi2\M,pi2\b,pi2)") + apply(force simp add: fresh_prod fresh_fun_simp_ImpR at_prm_fresh[OF at_coname_inst] swap_simps) + by (meson exists_fresh(2) fs_coname1) nominal_primrec (freshness_context: "(y::name,c::coname,P::trm)") substn :: "trm \ name \ coname \ trm \ trm" ("_{_:=<_>._}" [100,100,100,100] 100) @@ -798,50 +638,40 @@ | "\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(finite_guess | simp add: abs_fresh abs_supp fin_supp | fresh_guess | rule strip)+ +apply(subgoal_tac "\x::name. x\(x1,P,y1)") +apply(force simp add: fresh_prod fresh_fun_simp_NotL abs_fresh fresh_atm) +apply (meson exists_fresh(1) fs_name1) + +apply(simp add: abs_fresh abs_supp fin_supp | rule strip)+ +apply(subgoal_tac "\x::name. x\(x1,P,y1)") +apply(force simp add: fresh_prod fresh_fun_simp_AndL1 abs_fresh fresh_atm) +apply (meson exists_fresh(1) fs_name1) + +apply(simp add: abs_fresh abs_supp fin_supp | rule strip)+ +apply(subgoal_tac "\x::name. x\(x1,P,y1)") +apply(force simp add: fresh_prod fresh_fun_simp_AndL2 abs_fresh fresh_atm) +apply (meson exists_fresh(1) fs_name1) + +apply(simp add: abs_fresh abs_supp fin_supp | rule strip)+ +apply(subgoal_tac "\x::name. x\(x1,P,y1,x3,y2)") +apply(force simp add: fresh_prod fresh_fun_simp_OrL abs_fresh fresh_atm) +apply (meson exists_fresh(1) fs_name1) + +apply(simp add: abs_fresh abs_supp fin_supp | rule strip)+ +apply(subgoal_tac "\x::name. x\(x1,P,y1,x3,y2)") +apply(force simp add: fresh_prod fresh_fun_simp_OrL abs_fresh fresh_atm) +apply (meson exists_fresh(1) fs_name1) + +apply(simp add: abs_fresh abs_supp fin_supp | rule strip)+ +apply(subgoal_tac "\x::name. x\(x3,P,y1,y2)") +apply(force simp add: fresh_prod fresh_fun_simp_ImpL abs_fresh fresh_atm) +apply (meson exists_fresh(1) fs_name1) + +apply(simp add: abs_fresh abs_supp fin_supp | rule strip)+ +apply(subgoal_tac "\x::name. x\(x3,P,y1,y2)") +apply(force simp add: fresh_prod fresh_fun_simp_ImpL abs_fresh fresh_atm) +apply (meson exists_fresh(1) fs_name1) apply(fresh_guess)+ done @@ -870,801 +700,377 @@ 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(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)+ +apply(finite_guess | simp add: abs_fresh abs_supp fs_name1 fs_coname1 | rule strip)+ +apply(subgoal_tac "\x::coname. x\(x1,P,y1)") +apply(force simp add: fresh_prod fresh_fun_simp_NotR abs_fresh fresh_atm) +apply(meson exists_fresh' fin_supp) + +apply(simp add: abs_fresh abs_supp fs_name1 fs_coname1 | rule strip)+ +apply(subgoal_tac "\x::coname. x\(x1,P,y1,x3,y2)") +apply(force simp add: fresh_prod fresh_fun_simp_AndR abs_fresh fresh_atm) +apply(meson exists_fresh' fin_supp) + +apply(simp add: abs_fresh abs_supp fs_name1 fs_coname1 | rule strip)+ +apply(subgoal_tac "\x::coname. x\(x1,P,y1,x3,y2)") +apply(force simp add: fresh_prod fresh_fun_simp_AndR abs_fresh fresh_atm) +apply(meson exists_fresh' fin_supp) + +apply(simp add: abs_fresh abs_supp fs_name1 fs_coname1 | rule strip)+ +apply(subgoal_tac "\x::coname. x\(x1,P,y1)") +apply(force simp add: fresh_prod fresh_fun_simp_OrR1 abs_fresh fresh_atm) +apply(meson exists_fresh' fin_supp) + +apply(simp add: abs_fresh abs_supp fs_name1 fs_coname1 | rule strip)+ +apply(subgoal_tac "\x::coname. x\(x1,P,y1)") +apply(force simp add: fresh_prod fresh_fun_simp_OrR2 abs_fresh fresh_atm) +apply(meson exists_fresh' fin_supp) + +apply(simp add: abs_fresh abs_supp | rule strip)+ +apply(subgoal_tac "\x::coname. x\(x1,P,x2,y1)") +apply(force simp add: fresh_prod fresh_fun_simp_ImpR abs_fresh fresh_atm abs_supp) +apply(meson exists_fresh' fin_supp) + +apply(simp add: abs_fresh abs_supp | rule strip)+ +apply(subgoal_tac "\x::coname. x\(x1,P,x2,y1)") +apply(force simp add: fresh_prod fresh_fun_simp_ImpR abs_fresh fresh_atm) +apply(meson exists_fresh' fin_supp) + +apply(simp add: abs_fresh | fresh_guess add: abs_fresh)+ 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.strong_induct) -apply(auto simp add: eq_bij fresh_bij eqvts) -apply(perm_simp)+ -done +by (nominal_induct M avoiding: c x N rule: trm.strong_induct) + (auto simp: eq_bij fresh_bij eqvts; perm_simp)+ 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.strong_induct) -apply(auto simp add: eq_bij fresh_bij eqvts) -apply(perm_simp)+ -done +by (nominal_induct M avoiding: c x N rule: trm.strong_induct) + (auto simp: eq_bij fresh_bij eqvts; perm_simp)+ lemma supp_subst1: shows "supp (M{y:=.P}) \ ((supp M) - {y}) \ (supp P)" -apply(nominal_induct M avoiding: y P c rule: trm.strong_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 - +proof (nominal_induct M avoiding: y P c rule: trm.strong_induct) + case (NotL coname trm name) + obtain x'::name where "x'\(trm{y:=.P},P)" + by (meson exists_fresh(1) fs_name1) + with NotL + show ?case + by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_NotL; blast) +next + case (AndL1 name1 trm name2) + obtain x'::name where "x'\(trm{y:=.P},P,name1)" + by (meson exists_fresh(1) fs_name1) + with AndL1 show ?case + by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_AndL1 fresh_atm; blast) +next + case (AndL2 name1 trm name2) + obtain x'::name where "x'\(trm{y:=.P},P,name1)" + by (meson exists_fresh(1) fs_name1) + with AndL2 show ?case + by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_AndL2 fresh_atm; blast) +next + case (OrL name1 trm1 name2 trm2 name3) + obtain x'::name where "x'\(trm1{y:=.P},P,name1,trm2{y:=.P},name2)" + by (meson exists_fresh(1) fs_name1) + with OrL show ?case + by (auto simp: fs_name1 fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_OrL fresh_atm; blast) +next + case (ImpL coname trm1 name1 trm2 name2) + obtain x'::name where "x'\(trm1{name2:=.P},P,name1,trm2{name2:=.P})" + by (meson exists_fresh(1) fs_name1) + with ImpL show ?case + by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_ImpL fresh_atm; blast) +qed (simp add: abs_supp supp_atm trm.supp fs_name1; blast)+ + +text \Identical to the previous proof\ lemma supp_subst2: shows "supp (M{y:=.P}) \ supp (M) \ ((supp P) - {c})" -apply(nominal_induct M avoiding: y P c rule: trm.strong_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 +proof (nominal_induct M avoiding: y P c rule: trm.strong_induct) + case (NotL coname trm name) + obtain x'::name where "x'\(trm{y:=.P},P)" + by (meson exists_fresh(1) fs_name1) + with NotL + show ?case + by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_NotL; blast) +next + case (AndL1 name1 trm name2) + obtain x'::name where "x'\(trm{y:=.P},P,name1)" + by (meson exists_fresh(1) fs_name1) + with AndL1 show ?case + by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_AndL1 fresh_atm; blast) +next + case (AndL2 name1 trm name2) + obtain x'::name where "x'\(trm{y:=.P},P,name1)" + by (meson exists_fresh(1) fs_name1) + with AndL2 show ?case + by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_AndL2 fresh_atm; blast) +next + case (OrL name1 trm1 name2 trm2 name3) + obtain x'::name where "x'\(trm1{y:=.P},P,name1,trm2{y:=.P},name2)" + by (meson exists_fresh(1) fs_name1) + with OrL show ?case + by (auto simp: fs_name1 fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_OrL fresh_atm; blast) +next + case (ImpL coname trm1 name1 trm2 name2) + obtain x'::name where "x'\(trm1{name2:=.P},P,name1,trm2{name2:=.P})" + by (meson exists_fresh(1) fs_name1) + with ImpL show ?case + by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_ImpL fresh_atm; blast) +qed (simp add: abs_supp supp_atm trm.supp fs_name1; blast)+ lemma supp_subst3: shows "supp (M{c:=(x).P}) \ ((supp M) - {c}) \ (supp P)" -apply(nominal_induct M avoiding: x P c rule: trm.strong_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 +proof (nominal_induct M avoiding: x P c rule: trm.strong_induct) + case (NotR name trm coname) + obtain x'::coname where "x'\(trm{coname:=(x).P},P)" + by (meson exists_fresh'(2) fs_coname1) + with NotR show ?case + by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_NotR; blast) +next + case (AndR coname1 trm1 coname2 trm2 coname3) + obtain x'::coname where x': "x'\(trm1{coname3:=(x).P},P,trm2{coname3:=(x).P},coname1,coname2)" + by (meson exists_fresh'(2) fs_coname1) + with AndR show ?case + by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_AndR; blast) +next + case (OrR1 coname1 trm coname2) + obtain x'::coname where x': "x'\(trm{coname2:=(x).P},P,coname1)" + by (meson exists_fresh'(2) fs_coname1) + with OrR1 show ?case + by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_OrR1; blast) +next + case (OrR2 coname1 trm coname2) + obtain x'::coname where x': "x'\(trm{coname2:=(x).P},P,coname1)" + by (meson exists_fresh'(2) fs_coname1) + with OrR2 show ?case + by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_OrR2; blast) +next + case (ImpR name coname1 trm coname2) + obtain x'::coname where x': "x'\(trm{coname2:=(x).P},P,coname1)" + by (meson exists_fresh'(2) fs_coname1) + with ImpR show ?case + by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_ImpR; blast) +qed (simp add: abs_supp supp_atm trm.supp fs_name1; blast)+ + lemma supp_subst4: shows "supp (M{c:=(x).P}) \ (supp M) \ ((supp P) - {x})" -apply(nominal_induct M avoiding: x P c rule: trm.strong_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 +proof (nominal_induct M avoiding: x P c rule: trm.strong_induct) + case (NotR name trm coname) + obtain x'::coname where "x'\(trm{coname:=(x).P},P)" + by (meson exists_fresh'(2) fs_coname1) + with NotR show ?case + by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_NotR; blast) +next + case (AndR coname1 trm1 coname2 trm2 coname3) + obtain x'::coname where x': "x'\(trm1{coname3:=(x).P},P,trm2{coname3:=(x).P},coname1,coname2)" + by (meson exists_fresh'(2) fs_coname1) + with AndR show ?case + by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_AndR; blast) +next + case (OrR1 coname1 trm coname2) + obtain x'::coname where x': "x'\(trm{coname2:=(x).P},P,coname1)" + by (meson exists_fresh'(2) fs_coname1) + with OrR1 show ?case + by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_OrR1; blast) +next + case (OrR2 coname1 trm coname2) + obtain x'::coname where x': "x'\(trm{coname2:=(x).P},P,coname1)" + by (meson exists_fresh'(2) fs_coname1) + with OrR2 show ?case + by (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_OrR2; blast) +next + case (ImpR name coname1 trm coname2) + obtain x'::coname where x': "x'\(trm{coname2:=(x).P},P,coname1)" + by (meson exists_fresh'(2) fs_coname1) + with ImpR show ?case + by (auto simp: fresh_prod abs_supp supp_atm trm.supp fs_name1 fresh_fun_simp_ImpR; blast) +qed (simp add: abs_supp supp_atm trm.supp fs_name1; blast)+ lemma supp_subst5: shows "(supp M - {y}) \ supp (M{y:=.P})" -apply(nominal_induct M avoiding: y P c rule: trm.strong_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 +proof (nominal_induct M avoiding: y P c rule: trm.strong_induct) + case (NotL coname trm name) + obtain x'::name where "x'\(trm{y:=.P},P)" + by (meson exists_fresh(1) fs_name1) + with NotL + show ?case + apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_NotL) + apply (auto simp: fresh_def) + done +next + case (AndL1 name1 trm name2) + obtain x'::name where "x'\(trm{y:=.P},P,name1)" + by (meson exists_fresh(1) fs_name1) + with AndL1 show ?case + apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_AndL1) + apply (auto simp: fresh_def) + done +next + case (AndL2 name1 trm name2) + obtain x'::name where "x'\(trm{y:=.P},P,name1)" + by (meson exists_fresh(1) fs_name1) + with AndL2 show ?case + apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_AndL2) + apply (auto simp: fresh_def) + done +next + case (OrL name1 trm1 name2 trm2 name3) + obtain x'::name where "x'\(trm1{y:=.P},P,name1,trm2{y:=.P},name2)" + by (meson exists_fresh(1) fs_name1) + with OrL show ?case + apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_OrL) + apply (fastforce simp: fresh_def)+ + done +next + case (ImpL coname trm1 name1 trm2 name2) + obtain x'::name where "x'\(trm1{name2:=.P},P,name1,trm2{name2:=.P})" + by (meson exists_fresh(1) fs_name1) + with ImpL show ?case + apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_ImpL) + apply (fastforce simp: fresh_def)+ + done +qed (simp add: abs_supp supp_atm trm.supp fs_name1; blast)+ lemma supp_subst6: shows "(supp M) \ ((supp (M{y:=.P}))::coname set)" -apply(nominal_induct M avoiding: y P c rule: trm.strong_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 +proof (nominal_induct M avoiding: y P c rule: trm.strong_induct) + case (NotL coname trm name) + obtain x'::name where "x'\(trm{y:=.P},P)" + by (meson exists_fresh(1) fs_name1) + with NotL + show ?case + apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_NotL) + apply (auto simp: fresh_def) + done +next + case (AndL1 name1 trm name2) + obtain x'::name where "x'\(trm{y:=.P},P,name1)" + by (meson exists_fresh(1) fs_name1) + with AndL1 show ?case + apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_AndL1) + apply (auto simp: fresh_def) + done +next + case (AndL2 name1 trm name2) + obtain x'::name where "x'\(trm{y:=.P},P,name1)" + by (meson exists_fresh(1) fs_name1) + with AndL2 show ?case + apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_AndL2) + apply (auto simp: fresh_def) + done +next + case (OrL name1 trm1 name2 trm2 name3) + obtain x'::name where "x'\(trm1{y:=.P},P,name1,trm2{y:=.P},name2)" + by (meson exists_fresh(1) fs_name1) + with OrL show ?case + apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_OrL) + apply (fastforce simp: fresh_def)+ + done +next + case (ImpL coname trm1 name1 trm2 name2) + obtain x'::name where "x'\(trm1{name2:=.P},P,name1,trm2{name2:=.P})" + by (meson exists_fresh(1) fs_name1) + with ImpL show ?case + apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_ImpL) + apply (fastforce simp: fresh_def)+ + done +qed (simp add: abs_supp supp_atm trm.supp fs_name1; blast)+ lemma supp_subst7: shows "(supp M - {c}) \ supp (M{c:=(x).P})" -apply(nominal_induct M avoiding: x P c rule: trm.strong_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 - +proof (nominal_induct M avoiding: x P c rule: trm.strong_induct) + case (NotR name trm coname) + obtain x'::coname where "x'\(trm{coname:=(x).P},P)" + by (meson exists_fresh'(2) fs_coname1) + with NotR show ?case + apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_NotR) + apply (auto simp: fresh_def) + done +next + case (AndR coname1 trm1 coname2 trm2 coname3) + obtain x'::coname where "x'\(trm1{coname3:=(x).P},P,trm2{coname3:=(x).P},coname1,coname2)" + by (meson exists_fresh'(2) fs_coname1) + with AndR show ?case + apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_AndR) + apply (fastforce simp: fresh_def)+ + done +next + case (OrR1 coname1 trm coname2) + obtain x'::coname where "x'\(trm{coname2:=(x).P},P,coname1)" + by (meson exists_fresh'(2) fs_coname1) + with OrR1 show ?case + apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_OrR1) + apply (auto simp: fresh_def) + done +next + case (OrR2 coname1 trm coname2) + obtain x'::coname where "x'\(trm{coname2:=(x).P},P,coname1)" + by (meson exists_fresh'(2) fs_coname1) + with OrR2 show ?case + apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_OrR2) + apply (auto simp: fresh_def) + done +next + case (ImpR name coname1 trm coname2) + obtain x'::coname where "x'\(trm{coname2:=(x).P},P,coname1)" + by (meson exists_fresh'(2) fs_coname1) + with ImpR show ?case + apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_ImpR) + apply (auto simp: fresh_def) + done +qed (simp add: abs_supp supp_atm trm.supp fs_name1; blast)+ + lemma supp_subst8: shows "(supp M) \ ((supp (M{c:=(x).P}))::name set)" -apply(nominal_induct M avoiding: x P c rule: trm.strong_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 - +proof (nominal_induct M avoiding: x P c rule: trm.strong_induct) + case (NotR name trm coname) + obtain x'::coname where "x'\(trm{coname:=(x).P},P)" + by (meson exists_fresh'(2) fs_coname1) + with NotR show ?case + apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_NotR) + apply (auto simp: fresh_def) + done +next + case (AndR coname1 trm1 coname2 trm2 coname3) + obtain x'::coname where "x'\(trm1{coname3:=(x).P},P,trm2{coname3:=(x).P},coname1,coname2)" + by (meson exists_fresh'(2) fs_coname1) + with AndR show ?case + apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_AndR) + apply (fastforce simp: fresh_def)+ + done +next + case (OrR1 coname1 trm coname2) + obtain x'::coname where "x'\(trm{coname2:=(x).P},P,coname1)" + by (meson exists_fresh'(2) fs_coname1) + with OrR1 show ?case + apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_OrR1) + apply (auto simp: fresh_def) + done +next + case (OrR2 coname1 trm coname2) + obtain x'::coname where "x'\(trm{coname2:=(x).P},P,coname1)" + by (meson exists_fresh'(2) fs_coname1) + with OrR2 show ?case + apply (auto simp: fresh_prod abs_supp supp_atm trm.supp fresh_fun_simp_OrR2) + apply (auto simp: fresh_def) + done +next + case (ImpR name coname1 trm coname2) + obtain x'::coname where "x'\(trm{coname2:=(x).P},P,coname1)" + by (meson exists_fresh'(2) fs_coname1) + with ImpR show ?case + by (force simp: fresh_prod abs_supp fs_name1 supp_atm trm.supp fresh_fun_simp_ImpR) +qed (simp add: abs_supp supp_atm trm.supp fs_name1; blast)+ + lemmas subst_supp = supp_subst1 supp_subst2 supp_subst3 supp_subst4 supp_subst5 supp_subst6 supp_subst7 supp_subst8 lemma subst_fresh: @@ -1678,17 +1084,14 @@ 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 + using subst_supp + by(fastforce simp add: fresh_def supp_prod)+ 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.strong_induct) -apply(auto simp add: fresh_atm abs_fresh abs_supp fin_supp) +apply(auto simp: fresh_atm abs_fresh abs_supp fin_supp) done lemma substc_rename1: @@ -1696,198 +1099,81 @@ shows "M{a:=(x).N} = ([(c,a)]\M){c:=(x).N}" using a proof(nominal_induct M avoiding: c a x N rule: trm.strong_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 - by(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod) -next case (AndR c1 M c2 M' c3) then show ?case - apply(auto simp add: fresh_prod calc_atm fresh_atm abs_fresh fresh_left) - apply (metis (erased, opaque_lifting)) - by metis -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 - by(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod) -next - case (OrR2 d M e) - then show ?case - by(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod) -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) + apply(auto simp: fresh_prod calc_atm fresh_atm abs_fresh fresh_left) + apply (metis (no_types, lifting))+ + done next case ImpL then show ?case - by (auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) + by (auto simp: calc_atm alpha fresh_atm abs_fresh fresh_prod fresh_left) metis next - case (ImpR y d M e) - then show ?case - by(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) -next case (Cut d M y M') then show ?case by(simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) (metis crename.simps(1) crename_id crename_rename) -qed +qed (auto simp: calc_atm alpha fresh_atm abs_fresh fresh_prod fresh_left trm.inject) 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.strong_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 + obtain a::coname where "a\(N,M{d:=(y).([(y,x)]\N)},[(y,x)]\N)" + by (meson exists_fresh(2) fs_coname1) + with NotR show ?case + apply(auto simp: calc_atm alpha fresh_atm fresh_prod fresh_left) + by (metis (no_types, opaque_lifting) alpha(1) trm.inject(2)) 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) + obtain a'::coname where "a'\(N,M{c3:=(y).([(y,x)]\N)},M'{c3:=(y).([(y,x)]\N)},[(y,x)]\N,c1,c2,c3)" + by (meson exists_fresh(2) fs_coname1) + with AndR show ?case + apply(auto simp: calc_atm alpha fresh_atm fresh_prod fresh_left) + by (metis (no_types, opaque_lifting) alpha(1) trm.inject(2)) 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 + obtain a'::coname where "a'\(N,M{e:=(y).([(y,x)]\N)},[(y,x)]\N,d,e)" + by (meson exists_fresh(2) fs_coname1) + with OrR1 show ?case + by (auto simp: perm_swap calc_atm trm.inject alpha fresh_atm fresh_prod fresh_left fresh_fun_simp_OrR1) 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) + obtain a'::coname where "a'\(N,M{e:=(y).([(y,x)]\N)},[(y,x)]\N,d,e)" + by (meson exists_fresh(2) fs_coname1) + with OrR2 show ?case + by (auto simp: perm_swap calc_atm trm.inject alpha fresh_atm fresh_prod fresh_left fresh_fun_simp_OrR2) 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 + obtain a'::coname where "a'\(N,M{e:=(y).([(y,x)]\N)},[(y,x)]\N,d,e)" + by (meson exists_fresh(2) fs_coname1) + with ImpR show ?case + by (auto simp: perm_swap calc_atm trm.inject alpha fresh_atm fresh_prod fresh_left fresh_fun_simp_ImpR) +qed (auto simp: calc_atm trm.inject alpha fresh_atm fresh_prod fresh_left perm_swap) 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.strong_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 - by(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod) -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 - by(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod) -next - case (AndL2 u M v) - then show ?case - by(auto simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod) -next case (OrL x1 M x2 M' x3) then show ?case - by(simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) - (metis (poly_guards_query)) -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) + apply(auto simp add: calc_atm fresh_atm abs_fresh fresh_prod fresh_left) + by (metis (mono_tags))+ next case (ImpL d M v M' u) then show ?case - by(simp add: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) - (metis (poly_guards_query)) + apply(auto simp add: calc_atm fresh_atm abs_fresh fresh_prod fresh_left) + by (metis (mono_tags))+ 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) - apply metis - done -qed + apply(auto simp: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left) + by (metis nrename.simps(1) nrename_id nrename_rename)+ +qed (auto simp: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_left abs_supp fin_supp fresh_prod) lemma substn_rename4: assumes a: "c\(N,a)" @@ -1896,15 +1182,15 @@ proof(nominal_induct M avoiding: x c a N rule: trm.strong_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) + by (auto simp: 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) + by (auto simp: 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(auto simp: 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)+ @@ -1915,25 +1201,25 @@ 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(auto simp: 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 (auto simp: 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) + then show ?case by (auto simp: 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) + then show ?case by (auto simp: 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(auto simp: 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)+ @@ -1944,7 +1230,7 @@ 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(auto simp: 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)+ @@ -1955,15 +1241,15 @@ 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) + by (auto simp: 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) + by (auto simp: 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(auto simp: 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)+ @@ -1974,7 +1260,7 @@ 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) + by (auto simp: calc_atm trm.inject alpha fresh_atm abs_fresh fresh_prod fresh_left perm_swap) qed lemma subst_rename5: @@ -2005,9 +1291,9 @@ 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) + using fs1 fs2 by (rule_tac sym, auto simp: 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(auto simp: calc_atm) apply(drule pt_bij1[OF pt_coname_inst, OF at_coname_inst]) apply(simp add: calc_atm) done @@ -2015,12 +1301,12 @@ 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) + using fs1 fs2 by (auto simp: 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(auto simp: 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) @@ -2040,9 +1326,9 @@ 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) + using fs1 fs2 by (rule_tac sym, auto simp: 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(auto simp: calc_atm) apply(drule pt_bij1[OF pt_name_inst, OF at_name_inst]) apply(simp add: calc_atm) done @@ -2055,7 +1341,7 @@ using fs1 fs2 a apply - apply(simp only: eq2[symmetric]) - apply(auto simp add: trm.inject) + apply(auto simp: 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) @@ -2078,7 +1364,7 @@ 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(auto simp: 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) @@ -2091,7 +1377,7 @@ 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(auto simp: fresh_left calc_atm forget) apply(generate_fresh "coname") apply(rule_tac f="fresh_fun" in arg_cong) apply(simp add: fun_eq_iff) @@ -2107,7 +1393,7 @@ 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(auto simp: fresh_left calc_atm forget) apply(generate_fresh "name") apply(rule_tac f="fresh_fun" in arg_cong) apply(simp add: fun_eq_iff) @@ -2123,7 +1409,7 @@ 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(auto simp: fresh_left calc_atm forget abs_fresh)[1] apply(generate_fresh "name") apply(rule_tac f="fresh_fun" in arg_cong) apply(simp add: fun_eq_iff) @@ -2148,7 +1434,7 @@ 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(auto simp: fresh_left calc_atm forget abs_fresh)[1] apply(generate_fresh "name") apply(rule_tac f="fresh_fun" in arg_cong) apply(simp add: fun_eq_iff) @@ -2174,12 +1460,12 @@ 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(auto simp: 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(auto simp: fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(auto simp: fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(auto simp: fresh_prod fresh_atm)[1] apply(simp) apply(rule_tac f="fresh_fun" in arg_cong) apply(simp add: fun_eq_iff) @@ -2187,9 +1473,9 @@ 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(auto simp: fresh_left calc_atm abs_fresh)[1] apply(rule forget) -apply(auto simp add: fresh_left calc_atm abs_fresh)[1] +apply(auto simp: 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 @@ -2202,12 +1488,12 @@ 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(auto simp: 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(auto simp: fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(auto simp: fresh_left calc_atm fresh_prod fresh_atm)[1] +apply(auto simp: fresh_prod fresh_atm)[1] apply(simp) apply(rule_tac f="fresh_fun" in arg_cong) apply(simp add: fun_eq_iff) @@ -2215,9 +1501,9 @@ 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(auto simp: fresh_left calc_atm abs_fresh)[1] apply(rule forget) -apply(auto simp add: fresh_left calc_atm abs_fresh)[1] +apply(auto simp: 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 @@ -2229,7 +1515,7 @@ 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(auto simp: fresh_left calc_atm forget abs_fresh)[1] apply(rule_tac f="fresh_fun" in arg_cong) apply(simp add: fun_eq_iff) apply(rule allI) @@ -2253,7 +1539,7 @@ 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(auto simp: fresh_left calc_atm forget abs_fresh)[1] apply(rule_tac f="fresh_fun" in arg_cong) apply(simp add: fun_eq_iff) apply(rule allI) @@ -2278,7 +1564,7 @@ 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(auto simp: fresh_left calc_atm forget abs_fresh)[1] apply(rule_tac f="fresh_fun" in arg_cong) apply(simp add: fun_eq_iff) apply(rule allI) @@ -2303,7 +1589,7 @@ 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(auto simp: fresh_left calc_atm forget abs_fresh)[1] apply(rule_tac f="fresh_fun" in arg_cong) apply(simp add: fun_eq_iff) apply(rule allI) @@ -2357,7 +1643,7 @@ 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.strong_induct) -apply(auto simp add: subst_fresh rename_fresh trm.inject) +apply(auto simp: 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") @@ -2381,11 +1667,11 @@ 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(auto simp: fresh_atm)[1] apply(rule trans) apply(rule better_crename_Cut) apply(simp add: fresh_atm) -apply(auto simp add: fresh_atm)[1] +apply(auto simp: fresh_atm)[1] apply(drule crename_ax) apply(simp add: fresh_atm) apply(simp add: fresh_atm) @@ -2449,7 +1735,7 @@ 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.strong_induct) -apply(auto simp add: subst_fresh rename_fresh trm.inject) +apply(auto simp: subst_fresh rename_fresh trm.inject) apply(rule trans) apply(rule better_crename_Cut) apply(simp add: fresh_atm fresh_prod) @@ -2523,7 +1809,7 @@ 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.strong_induct) -apply(auto simp add: subst_fresh rename_fresh trm.inject) +apply(auto simp: subst_fresh rename_fresh trm.inject) apply(rule trans) apply(rule better_nrename_Cut) apply(simp add: fresh_prod fresh_atm) @@ -2595,7 +1881,7 @@ 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.strong_induct) -apply(auto simp add: subst_fresh rename_fresh trm.inject) +apply(auto simp: subst_fresh rename_fresh trm.inject) apply(rule trans) apply(rule better_nrename_Cut) apply(simp add: fresh_atm fresh_prod) @@ -2865,9 +2151,9 @@ 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) + show "a\\" by (induct \) (auto simp: fresh_list_nil fresh_list_cons fresh_prod fresh_atm fresh_ty) +next + show "x\\" by (induct \) (auto simp: fresh_list_nil fresh_list_cons fresh_prod fresh_atm fresh_ty) qed text \cut-reductions\ @@ -2891,7 +2177,7 @@ shows "x=y" using a apply(erule_tac fin.cases) -apply(auto simp add: trm.inject) +apply(auto simp: trm.inject) done lemma fin_NotL_elim: @@ -2899,7 +2185,7 @@ shows "x=y \ x\M" using a apply(erule_tac fin.cases) -apply(auto simp add: trm.inject) +apply(auto simp: trm.inject) apply(subgoal_tac "y\[aa].Ma") apply(drule sym) apply(simp_all add: abs_fresh) @@ -2910,7 +2196,7 @@ shows "z=y \ z\[x].M" using a apply(erule_tac fin.cases) -apply(auto simp add: trm.inject) +apply(auto simp: trm.inject) done lemma fin_AndL2_elim: @@ -2918,7 +2204,7 @@ shows "z=y \ z\[x].M" using a apply(erule_tac fin.cases) -apply(auto simp add: trm.inject) +apply(auto simp: trm.inject) done lemma fin_OrL_elim: @@ -2926,7 +2212,7 @@ shows "z=u \ z\[x].M \ z\[y].N" using a apply(erule_tac fin.cases) -apply(auto simp add: trm.inject) +apply(auto simp: trm.inject) done lemma fin_ImpL_elim: @@ -2934,7 +2220,7 @@ shows "z=y \ z\M \ z\[x].N" using a apply(erule_tac fin.cases) -apply(auto simp add: trm.inject) +apply(auto simp: trm.inject) apply(subgoal_tac "y\[aa].Ma") apply(drule sym) apply(simp_all add: abs_fresh) @@ -2957,7 +2243,7 @@ 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) + (auto simp: calc_atm simp add: fresh_left abs_fresh) lemma not_fin_subst1: assumes a: "\(fin M x)" @@ -2995,13 +2281,13 @@ 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(auto simp: 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(auto simp: abs_fresh)[1] apply(drule freshn_after_substc) apply(subgoal_tac "name2\[name1].trm") apply(simp add: fin.intros) @@ -3023,7 +2309,7 @@ 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(auto simp: abs_fresh)[1] apply(drule freshn_after_substc)+ apply(subgoal_tac "name3\[name1].trm1 \ name3\[name2].trm2") apply(simp add: fin.intros) @@ -3037,7 +2323,7 @@ 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(auto simp: abs_fresh)[1] apply(drule freshn_after_substc)+ apply(subgoal_tac "x\[name1].trm2") apply(simp add: fin.intros) @@ -3075,7 +2361,7 @@ 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(auto simp: abs_fresh)[1] apply(drule freshn_after_substn) apply(simp) apply(subgoal_tac "name2\[name1]. trm") @@ -3089,7 +2375,7 @@ 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(auto simp: abs_fresh)[1] apply(drule freshn_after_substn) apply(simp) apply(subgoal_tac "name2\[name1].trm") @@ -3105,7 +2391,7 @@ 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(auto simp: abs_fresh)[1] apply(drule freshn_after_substn) apply(simp) apply(drule freshn_after_substn) @@ -3122,7 +2408,7 @@ 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(auto simp: abs_fresh)[1] apply(drule freshn_after_substn) apply(simp) apply(drule freshn_after_substn) @@ -3315,7 +2601,7 @@ shows "a=b" using a apply(erule_tac fic.cases) -apply(auto simp add: trm.inject) +apply(auto simp: trm.inject) done lemma fic_NotR_elim: @@ -3323,7 +2609,7 @@ shows "a=b \ b\M" using a apply(erule_tac fic.cases) -apply(auto simp add: trm.inject) +apply(auto simp: trm.inject) apply(subgoal_tac "b\[xa].Ma") apply(drule sym) apply(simp_all add: abs_fresh) @@ -3334,7 +2620,7 @@ shows "b=c \ c\[a].M" using a apply(erule_tac fic.cases) -apply(auto simp add: trm.inject) +apply(auto simp: trm.inject) done lemma fic_OrR2_elim: @@ -3342,7 +2628,7 @@ shows "b=c \ c\[a].M" using a apply(erule_tac fic.cases) -apply(auto simp add: trm.inject) +apply(auto simp: trm.inject) done lemma fic_AndR_elim: @@ -3350,7 +2636,7 @@ shows "c=d \ d\[a].M \ d\[b].N" using a apply(erule_tac fic.cases) -apply(auto simp add: trm.inject) +apply(auto simp: trm.inject) done lemma fic_ImpR_elim: @@ -3358,7 +2644,7 @@ shows "b=c \ b\[a].M" using a apply(erule_tac fic.cases) -apply(auto simp add: trm.inject) +apply(auto simp: trm.inject) apply(subgoal_tac "c\[xa].[aa].Ma") apply(drule sym) apply(simp_all add: abs_fresh) @@ -3379,7 +2665,7 @@ 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) + (auto simp: calc_atm simp add: fresh_left abs_fresh) lemma not_fic_subst1: assumes a: "\(fic M a)" @@ -3734,13 +3020,13 @@ and "M \\<^sub>l M' \ a\M \ a\M'" apply - apply(induct rule: l_redu.induct) -apply(auto simp add: abs_fresh rename_fresh) +apply(auto simp: 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(auto simp: abs_fresh rename_fresh) apply(case_tac "aa=a") apply(simp add: rename_fresh) apply(simp add: rename_fresh fresh_atm) @@ -3754,7 +3040,7 @@ 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) + using fs1 fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm) also have "\ \\<^sub>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) @@ -3768,7 +3054,7 @@ 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) + using fs1 fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm) also have "\ \\<^sub>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) @@ -3786,17 +3072,17 @@ 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) + by (rule_tac sym, auto simp: 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) + by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm) also have "\ \\<^sub>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) + using f1 f2 f3 f4 by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm) finally show ?thesis by simp qed @@ -3815,7 +3101,7 @@ 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) + apply(auto simp: perm_fresh_fresh) done also have "\ = Cut .(AndR .([(b1',b1)]\M1) .([(b2',b2)]\M2) a') (y').(AndL1 (x').([(x',x)]\N) y')" @@ -3827,10 +3113,10 @@ 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) + apply(auto simp: 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) + using f1 f2 f3 f4 f5 fs by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm) finally show ?thesis by simp qed @@ -3849,7 +3135,7 @@ 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) + apply(auto simp: perm_fresh_fresh) done also have "\ = Cut .(AndR .([(b1',b1)]\M1) .([(b2',b2)]\M2) a') (y').(AndL2 (x').([(x',x)]\N) y')" @@ -3861,10 +3147,10 @@ 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) + apply(auto simp: 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) + using f1 f2 f3 f4 f5 fs by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm) finally show ?thesis by simp qed @@ -3883,7 +3169,7 @@ 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) + apply(auto simp: perm_fresh_fresh) done also have "\ = Cut .(OrR1 .([(a',a)]\M) b') (y').(OrL (x1').([(x1',x1)]\N1) (x2').([(x2',x2)]\N2) y')" @@ -3895,10 +3181,10 @@ 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) + apply(auto simp: 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) + using f1 f2 f3 f4 f5 fs by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm) finally show ?thesis by simp qed @@ -3917,7 +3203,7 @@ 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) + apply(auto simp: perm_fresh_fresh) done also have "\ = Cut .(OrR2 .([(a',a)]\M) b') (y').(OrL (x1').([(x1',x1)]\N1) (x2').([(x2',x2)]\N2) y')" @@ -3929,10 +3215,10 @@ 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) + apply(auto simp: 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) + using f1 f2 f3 f4 f5 fs by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm) finally show ?thesis by simp qed @@ -3952,7 +3238,7 @@ 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) + apply(auto simp: 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')" @@ -3973,7 +3259,7 @@ 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) + apply(auto simp: 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 @@ -3986,7 +3272,7 @@ 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(auto simp: fresh_prod fresh_atm)[1] apply(perm_simp add: alpha) apply(perm_simp add: alpha) apply(perm_simp add: alpha) @@ -4007,7 +3293,7 @@ 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 fresh_prod fresh_atm) +apply(auto simp: alpha_fresh fresh_prod fresh_atm) apply(drule sym) apply(perm_simp) done @@ -4018,7 +3304,7 @@ 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 fresh_prod fresh_atm) +apply(auto simp: alpha_fresh fresh_prod fresh_atm) apply(drule sym) apply(perm_simp) done @@ -4030,7 +3316,7 @@ 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 fresh_prod fresh_atm +apply(auto simp: alpha_fresh fresh_prod fresh_atm abs_supp fin_supp abs_fresh abs_perm fresh_left calc_atm) apply(drule sym) apply(simp) @@ -4099,7 +3385,7 @@ apply(rule exI)+ apply(rule conjI) apply(rule refl) -apply(auto simp add: calc_atm abs_fresh fresh_left)[1] +apply(auto simp: calc_atm abs_fresh fresh_left)[1] apply(case_tac "y=x") apply(perm_simp) apply(perm_simp) @@ -4136,7 +3422,7 @@ 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(auto simp: 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] @@ -4148,7 +3434,7 @@ 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(auto simp: 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) @@ -4161,7 +3447,7 @@ 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(auto simp: 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) @@ -4174,7 +3460,7 @@ 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(auto simp: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] apply(perm_simp)+ (* and2 case *) apply(rule disjI2) @@ -4206,7 +3492,7 @@ 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(auto simp: 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] @@ -4218,7 +3504,7 @@ 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(auto simp: 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) @@ -4231,7 +3517,7 @@ 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(auto simp: 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) @@ -4244,7 +3530,7 @@ 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(auto simp: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] apply(perm_simp)+ (* or1 case *) apply(rule disjI2) @@ -4278,7 +3564,7 @@ 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(auto simp: 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) @@ -4292,7 +3578,7 @@ 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(auto simp: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] apply(perm_simp)+ (* or2 case *) apply(rule disjI2) @@ -4326,7 +3612,7 @@ 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(auto simp: 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) @@ -4339,7 +3625,7 @@ 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(auto simp: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] apply(perm_simp)+ (* imp-case *) apply(rule disjI2) @@ -4373,7 +3659,7 @@ 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(auto simp: 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) @@ -4386,7 +3672,7 @@ 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(auto simp: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] apply(perm_simp)+ done @@ -4408,7 +3694,7 @@ 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) + using fs1 fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm) also have "\ \\<^sub>c ([(a',a)]\M){a':=(x').([(x',x)]\N)}" using fs1 fs2 not_fic apply - apply(rule left) @@ -4429,7 +3715,7 @@ 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) + using fs1 fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm) also have "\ \\<^sub>c ([(x',x)]\N){x':=.([(a',a)]\M)}" using fs1 fs2 not_fin apply - apply(rule right) @@ -4450,9 +3736,9 @@ and "M \\<^sub>c M' \ c\M \ c\M'" apply - apply(induct rule: c_redu.induct) -apply(auto simp add: abs_fresh rename_fresh subst_fresh) +apply(auto simp: abs_fresh rename_fresh subst_fresh) apply(induct rule: c_redu.induct) -apply(auto simp add: abs_fresh rename_fresh subst_fresh) +apply(auto simp: abs_fresh rename_fresh subst_fresh) done inductive @@ -4485,11 +3771,11 @@ 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(auto simp: 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) +apply(auto simp: abs_fresh abs_supp fin_supp) done equivariance a_redu @@ -4504,11 +3790,11 @@ 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) + using fs1 fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm) also have "\ \\<^sub>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) + using fs1 fs2 red by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu) finally show ?thesis by simp qed @@ -4519,11 +3805,11 @@ 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) + using fs1 fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm) also have "\ \\<^sub>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) + using fs1 fs2 red by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu) finally show ?thesis by simp qed @@ -4534,11 +3820,11 @@ 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) + using fs1 fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm) also have "\ \\<^sub>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) + using fs1 fs2 red by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu) finally show ?thesis by simp qed @@ -4549,11 +3835,11 @@ 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) + using fs1 fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm) also have "\ \\<^sub>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) + using fs1 fs2 red by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu) finally show ?thesis by simp qed @@ -4563,11 +3849,11 @@ assume red: "M\\<^sub>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) + using fs1 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm) also have "\ \\<^sub>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) + using fs1 red by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu) finally show ?thesis by simp qed @@ -4577,11 +3863,11 @@ assume red: "M\\<^sub>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) + using fs1 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm) also have "\ \\<^sub>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) + using fs1 red by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu) finally show ?thesis by simp qed @@ -4592,11 +3878,11 @@ 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) + using fs1 fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm) also have "\ \\<^sub>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) + using fs1 fs2 red by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu) finally show ?thesis by simp qed @@ -4607,11 +3893,11 @@ 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) + using fs1 fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm) also have "\ \\<^sub>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) + using fs1 fs2 red by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu) finally show ?thesis by simp qed @@ -4621,11 +3907,11 @@ assume red: "M\\<^sub>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) + using fs1 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm) also have "\ \\<^sub>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) + using fs1 red by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu) finally show ?thesis by simp qed @@ -4635,11 +3921,11 @@ assume red: "M\\<^sub>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) + using fs1 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm) also have "\ \\<^sub>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) + using fs1 red by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu) finally show ?thesis by simp qed @@ -4650,11 +3936,11 @@ 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) + using fs1 fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm) also have "\ \\<^sub>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) + using fs1 fs2 red by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu) finally show ?thesis by simp qed @@ -4665,11 +3951,11 @@ 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) + using fs1 fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm) also have "\ \\<^sub>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) + using fs1 fs2 red by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu) finally show ?thesis by simp qed @@ -4679,11 +3965,11 @@ assume red: "M\\<^sub>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) + using fs2 by (rule_tac sym, auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm) also have "\ \\<^sub>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) + using fs2 red by (auto simp: trm.inject alpha fresh_atm fresh_prod calc_atm fresh_a_redu) finally show ?thesis by simp qed @@ -4700,7 +3986,7 @@ lemma ax_do_not_a_reduce: shows "Ax x a \\<^sub>a M \ False" apply(erule_tac a_redu.cases) -apply(auto simp add: trm.inject) +apply(auto simp: trm.inject) apply(drule ax_do_not_l_reduce) apply(simp) apply(drule ax_do_not_c_reduce) @@ -4719,12 +4005,12 @@ 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(auto simp: 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(auto simp: 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(auto simp: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu) apply(simp add: perm_swap) done @@ -4740,12 +4026,12 @@ 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(auto simp: 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(auto simp: 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(auto simp: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu) apply(simp add: perm_swap) done @@ -4761,81 +4047,81 @@ 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] done lemma a_redu_AndL1_elim: @@ -4850,12 +4136,12 @@ 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(auto simp: 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(auto simp: 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(auto simp: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu) apply(simp add: perm_swap) done @@ -4871,12 +4157,12 @@ 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(auto simp: 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(auto simp: 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(auto simp: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu) apply(simp add: perm_swap) done @@ -4892,81 +4178,81 @@ 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] done lemma a_redu_OrR1_elim: @@ -4981,12 +4267,12 @@ 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(auto simp: 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(auto simp: 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(auto simp: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu) apply(simp add: perm_swap) done @@ -5002,12 +4288,12 @@ 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(auto simp: 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(auto simp: 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(auto simp: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu) apply(simp add: perm_swap) done @@ -5023,81 +4309,81 @@ 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] done lemma a_redu_ImpR_elim: @@ -5112,43 +4398,43 @@ 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: 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(auto simp: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap) apply(rule sym) apply(rule trans) apply(rule perm_compose) @@ -5301,7 +4587,7 @@ apply(induct set: rtranclp) apply(auto) apply(drule a_redu_AndR_elim) -apply(auto simp add: alpha trm.inject) +apply(auto simp: alpha trm.inject) done lemma a_star_redu_AndL1_elim: @@ -5311,7 +4597,7 @@ apply(induct set: rtranclp) apply(auto) apply(drule a_redu_AndL1_elim) -apply(auto simp add: alpha trm.inject) +apply(auto simp: alpha trm.inject) done lemma a_star_redu_AndL2_elim: @@ -5321,7 +4607,7 @@ apply(induct set: rtranclp) apply(auto) apply(drule a_redu_AndL2_elim) -apply(auto simp add: alpha trm.inject) +apply(auto simp: alpha trm.inject) done lemma a_star_redu_OrL_elim: @@ -5331,7 +4617,7 @@ apply(induct set: rtranclp) apply(auto) apply(drule a_redu_OrL_elim) -apply(auto simp add: alpha trm.inject) +apply(auto simp: alpha trm.inject) done lemma a_star_redu_OrR1_elim: @@ -5341,7 +4627,7 @@ apply(induct set: rtranclp) apply(auto) apply(drule a_redu_OrR1_elim) -apply(auto simp add: alpha trm.inject) +apply(auto simp: alpha trm.inject) done lemma a_star_redu_OrR2_elim: @@ -5351,7 +4637,7 @@ apply(induct set: rtranclp) apply(auto) apply(drule a_redu_OrR2_elim) -apply(auto simp add: alpha trm.inject) +apply(auto simp: alpha trm.inject) done lemma a_star_redu_ImpR_elim: @@ -5361,7 +4647,7 @@ apply(induct set: rtranclp) apply(auto) apply(drule a_redu_ImpR_elim) -apply(auto simp add: alpha trm.inject) +apply(auto simp: alpha trm.inject) done lemma a_star_redu_ImpL_elim: @@ -5371,7 +4657,7 @@ apply(induct set: rtranclp) apply(auto) apply(drule a_redu_ImpL_elim) -apply(auto simp add: alpha trm.inject) +apply(auto simp: alpha trm.inject) done text \Substitution\ @@ -5701,7 +4987,7 @@ shows "fin M' x" using b a apply(induct set: rtranclp) -apply(auto simp add: fin_a_reduce) +apply(auto simp: fin_a_reduce) done lemma fic_l_reduce: @@ -5767,7 +5053,7 @@ shows "fic M' x" using b a apply(induct set: rtranclp) -apply(auto simp add: fic_a_reduce) +apply(auto simp: fic_a_reduce) done text \substitution properties\ @@ -5882,7 +5168,7 @@ 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) + by (auto simp: fresh_prod fresh_atm nrename_fresh) also have "\ \\<^sub>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} \\<^sub>a* (AndL1 (u).M v)[x\n>y]" by simp @@ -5916,7 +5202,7 @@ 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) + by (auto simp: fresh_prod fresh_atm nrename_fresh) also have "\ \\<^sub>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} \\<^sub>a* (AndL2 (u).M v)[x\n>y]" by simp @@ -5966,7 +5252,7 @@ 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) + by (auto simp: fresh_prod fresh_atm nrename_fresh subst_fresh) also have "\ \\<^sub>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 @@ -6012,7 +5298,7 @@ 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) + by (auto simp: fresh_prod subst_fresh fresh_atm trm.inject alpha rename_fresh) also have "\ \\<^sub>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 @@ -6132,7 +5418,7 @@ 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) + by (auto simp: fresh_prod fresh_atm subst_fresh crename_fresh) also have "\ \\<^sub>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} \\<^sub>a* (AndR .M .N e)[b\c>a]" by simp @@ -6182,7 +5468,7 @@ 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) + by (auto simp: fresh_prod fresh_atm crename_fresh subst_fresh) also have "\ \\<^sub>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} \\<^sub>a* (OrR1 .M d)[b\c>a]" by simp @@ -6216,7 +5502,7 @@ 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) + by (auto simp: fresh_prod fresh_atm crename_fresh subst_fresh) also have "\ \\<^sub>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} \\<^sub>a* (OrR2 .M d)[b\c>a]" by simp @@ -6259,7 +5545,7 @@ 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) + by (auto simp: fresh_prod crename_fresh subst_fresh fresh_atm) also have "\ \\<^sub>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} \\<^sub>a* (ImpR (z)..M d)[b\c>a]" using eq by simp @@ -6288,7 +5574,7 @@ 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.strong_induct) -apply(auto simp add: fresh_atm abs_fresh abs_supp fin_supp) +apply(auto simp: 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) @@ -6360,7 +5646,7 @@ lemma not_Ax2: shows "\(x\M) \ M{x:=.Q} \ Ax y a" apply(nominal_induct M avoiding: b y Q x a rule: trm.strong_induct) -apply(auto simp add: fresh_atm abs_fresh abs_supp fin_supp) +apply(auto simp: 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) @@ -6442,7 +5728,7 @@ proof(nominal_induct N avoiding: x y c P rule: trm.strong_induct) case Ax then show ?case - by (auto simp add: abs_fresh fresh_atm forget trm.inject) + by (auto simp: abs_fresh fresh_atm forget trm.inject) next case (Cut d M u M' x' y' c P) with assms show ?case @@ -6529,11 +5815,11 @@ next case NotR then show ?case - by (auto simp add: abs_fresh fresh_atm forget) + by (auto simp: abs_fresh fresh_atm forget) next case (NotL d M u) then show ?case - apply (auto simp add: abs_fresh fresh_atm forget) + apply (auto simp: 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)+ @@ -6543,7 +5829,7 @@ apply(simp add: abs_fresh) apply(simp) apply(simp) - apply(auto simp add: fresh_atm) + apply(auto simp: 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)") @@ -6566,11 +5852,11 @@ next case (AndR d1 M d2 M' d3) then show ?case - by (auto simp add: abs_fresh fresh_atm forget trm.inject subst_fresh) + by (auto simp: 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(auto simp: 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)+ @@ -6580,7 +5866,7 @@ apply(simp add: abs_fresh) apply(simp) apply(simp) - apply(auto simp add: fresh_atm) + apply(auto simp: 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)") @@ -6593,13 +5879,13 @@ apply(simp add: abs_fresh) apply(simp) apply(simp) - apply(auto simp add: fresh_atm) + apply(auto simp: 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(auto simp: 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)+ @@ -6609,7 +5895,7 @@ apply(simp add: abs_fresh) apply(simp) apply(simp) - apply(auto simp add: fresh_atm) + apply(auto simp: 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)") @@ -6622,21 +5908,21 @@ apply(simp add: abs_fresh) apply(simp) apply(simp) - apply(auto simp add: fresh_atm) + apply(auto simp: 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) + by (auto simp: 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) + by (auto simp: 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(auto simp: 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) @@ -6647,7 +5933,7 @@ apply(simp add: abs_fresh) apply(simp) apply(simp) - apply(auto simp add: fresh_atm) + apply(auto simp: fresh_atm) apply(simp add: trm.inject alpha forget) apply(rule trans) apply(rule substn.simps) @@ -6674,17 +5960,17 @@ apply(simp add: abs_fresh subst_fresh fresh_atm) apply(force) apply(simp) - apply(auto simp add: fresh_atm) + apply(auto simp: 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) + by (auto simp: 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(auto simp: 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) @@ -6695,7 +5981,7 @@ apply(simp add: abs_fresh) apply(simp) apply(simp) - apply(auto simp add: fresh_atm) + apply(auto simp: fresh_atm) apply(simp add: trm.inject alpha forget) apply(rule trans) apply(rule substn.simps) @@ -6720,7 +6006,7 @@ 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(auto simp: fresh_atm) apply(rule exists_fresh'(1)[OF fs_name1]) done qed @@ -6750,12 +6036,12 @@ proof(nominal_induct N avoiding: a b y P rule: trm.strong_induct) case Ax then show ?case - by (auto simp add: abs_fresh fresh_atm forget trm.inject) + by (auto simp: abs_fresh fresh_atm forget trm.inject) next case (Cut d M u M' x' y' c P) with assms show ?case apply(simp) - apply(auto simp add: trm.inject) + apply(auto simp: trm.inject) apply(rule trans) apply(rule better_Cut_substc) apply(simp) @@ -6815,11 +6101,11 @@ next case NotL then show ?case - by (auto simp add: abs_fresh fresh_atm forget) + by (auto simp: abs_fresh fresh_atm forget) next case (NotR u M d) then show ?case - apply (auto simp add: abs_fresh fresh_atm forget) + apply (auto simp: 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)+ @@ -6829,7 +6115,7 @@ apply(simp add: abs_fresh) apply(simp add: abs_fresh) apply(simp) - apply(auto simp add: fresh_atm) + apply(auto simp: 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)") @@ -6852,7 +6138,7 @@ 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(auto simp: 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) @@ -6863,7 +6149,7 @@ apply(simp add: abs_fresh fresh_atm) apply(simp add: abs_fresh fresh_atm) apply(simp) - apply(auto simp add: fresh_atm) + apply(auto simp: fresh_atm) apply(simp add: trm.inject alpha forget) apply(rule trans) apply(rule substc.simps) @@ -6890,21 +6176,21 @@ apply(simp add: abs_fresh subst_fresh fresh_atm) apply(force) apply(simp) - apply(auto simp add: fresh_atm) + apply(auto simp: 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) + by (auto simp: 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) + by (auto simp: 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 (auto simp: 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)+ @@ -6914,7 +6200,7 @@ apply(simp add: abs_fresh) apply(simp add: abs_fresh) apply(simp) - apply(auto simp add: fresh_atm) + apply(auto simp: 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)") @@ -6937,7 +6223,7 @@ next case (OrR2 d M e) then show ?case - apply (auto simp add: abs_fresh fresh_atm forget) + apply (auto simp: 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)+ @@ -6947,7 +6233,7 @@ apply(simp add: abs_fresh) apply(simp add: abs_fresh) apply(simp) - apply(auto simp add: fresh_atm) + apply(auto simp: 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)") @@ -6970,15 +6256,15 @@ next case (OrL x1 M x2 M' x3) then show ?case - by(auto simp add: abs_fresh fresh_atm forget trm.inject subst_fresh) + by(auto simp: 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) + by (auto simp: 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(auto simp: 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)+ @@ -6988,7 +6274,7 @@ apply(simp add: abs_fresh) apply(simp add: abs_fresh) apply(simp) - apply(auto simp add: fresh_atm) + apply(auto simp: 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})") @@ -7007,7 +6293,7 @@ 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(auto simp: fresh_atm) apply(rule exists_fresh'(2)[OF fs_coname1]) done qed @@ -7043,7 +6329,7 @@ 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) + using fs asm by (auto simp: 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 @@ -7055,14 +6341,14 @@ 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) + by (auto simp: 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 "\ = Cut .(P{b:=(y).Q}) (x).Ax z c" using fs asm by (auto simp: 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 @@ -7078,12 +6364,12 @@ have "(Cut .M (z).N){x:=.P}{b:=(y).Q} = (Cut .P (z).(N{x:=.P})){b:=(y).Q}" using Cut asm by simp also have "\ = (Cut .P (z).N){b:=(y).Q}" using Cut asm by (simp add: fresh_atm) - also have "\ = (Cut .(P{b:=(y).Q}) (y).Q)" using Cut asm by (auto simp add: fresh_prod fresh_atm) + also have "\ = (Cut .(P{b:=(y).Q}) (y).Q)" using Cut asm by (auto simp: 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 Cut asm by (simp add: fresh_atm) also have "\ = Cut .(P{b:=(y).Q}) (y).(Q{x:=.(P{b:=(y).Q})})" using Cut asm - by (auto simp add: fresh_prod fresh_atm subst_fresh) + by (auto simp: fresh_prod fresh_atm subst_fresh) also have "\ = Cut .(P{b:=(y).Q}) (y).Q" using Cut asm 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 @@ -7109,7 +6395,7 @@ 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 Cut asm by simp also have "\ = Cut .(M{b:=(y).Q}{x:=.(P{b:=(y).Q})}) (y).(Q{x:=.(P{b:=(y).Q})})" - using Cut asm neq by (auto simp add: fresh_prod fresh_atm subst_fresh abs_fresh) + using Cut asm neq by (auto simp: fresh_prod fresh_atm subst_fresh abs_fresh) also have "\ = Cut .(M{b:=(y).Q}{x:=.(P{b:=(y).Q})}) (y).Q" using Cut asm 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 @@ -7135,7 +6421,7 @@ = (Cut .(M{b:=(y).Q}) (z).(N{b:=(y).Q})){x:=.(P{b:=(y).Q})}" using Cut asm by auto also have "\ = (Cut .M (z).(N{b:=(y).Q})){x:=.(P{b:=(y).Q})}" - using Cut asm by (auto simp add: fresh_atm) + using Cut asm by (auto simp: fresh_atm) also have "\ = Cut .(P{b:=(y).Q}) (z).(N{b:=(y).Q}{x:=.(P{b:=(y).Q})})" using Cut asm by (simp add: fresh_prod fresh_atm subst_fresh) finally @@ -7179,7 +6465,7 @@ next case (NotR z M c) then show ?case - apply(auto simp add: fresh_prod fresh_atm subst_fresh) + apply(auto simp: 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) @@ -7197,7 +6483,7 @@ next case (NotL c M z) then show ?case - apply(auto simp add: fresh_prod fresh_atm subst_fresh) + apply(auto simp: 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) @@ -7208,7 +6494,7 @@ next case (AndR c1 M c2 N c3) then show ?case - apply(auto simp add: fresh_prod fresh_atm subst_fresh) + apply(auto simp: 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) @@ -7225,7 +6511,7 @@ next case (AndL1 z1 M z2) then show ?case - apply(auto simp add: fresh_prod fresh_atm subst_fresh) + apply(auto simp: 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) @@ -7236,7 +6522,7 @@ next case (AndL2 z1 M z2) then show ?case - apply(auto simp add: fresh_prod fresh_atm subst_fresh) + apply(auto simp: 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) @@ -7247,7 +6533,7 @@ next case (OrL z1 M z2 N z3) then show ?case - apply(auto simp add: fresh_prod fresh_atm subst_fresh) + apply(auto simp: 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) @@ -7263,7 +6549,7 @@ next case (OrR1 c1 M c2) then show ?case - apply(auto simp add: fresh_prod fresh_atm subst_fresh) + apply(auto simp: 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) @@ -7277,7 +6563,7 @@ next case (OrR2 c1 M c2) then show ?case - apply(auto simp add: fresh_prod fresh_atm subst_fresh) + apply(auto simp: 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) @@ -7291,7 +6577,7 @@ next case (ImpR z c M d) then show ?case - apply(auto simp add: fresh_prod fresh_atm subst_fresh) + apply(auto simp: 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) @@ -7304,7 +6590,7 @@ next case (ImpL c M z N u) then show ?case - apply(auto simp add: fresh_prod fresh_atm subst_fresh) + apply(auto simp: 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) @@ -7326,7 +6612,7 @@ proof(nominal_induct M avoiding: a x N y b P rule: trm.strong_induct) case (Ax z c) then show ?case - by (auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget trm.inject) + by (auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget trm.inject) next case (Cut d M' u M'') then show ?case @@ -7385,7 +6671,7 @@ next case (NotR z M' d) then show ?case - apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget) + apply(auto simp: 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)+ @@ -7405,7 +6691,7 @@ 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(auto simp: 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)+ @@ -7427,7 +6713,7 @@ 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(auto simp: 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) @@ -7450,7 +6736,7 @@ 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(auto simp: 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)+ @@ -7471,7 +6757,7 @@ 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(auto simp: 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)+ @@ -7492,7 +6778,7 @@ 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(auto simp: 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) @@ -7516,7 +6802,7 @@ 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(auto simp: 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) @@ -7537,7 +6823,7 @@ 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(auto simp: 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) @@ -7558,7 +6844,7 @@ 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(auto simp: 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) @@ -7581,7 +6867,7 @@ 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(auto simp: 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) @@ -7610,7 +6896,7 @@ proof(nominal_induct N avoiding: x y a c M P rule: trm.strong_induct) case (Ax z c) then show ?case - by(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget) + by(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget) next case (Cut d M' u M'') then show ?case @@ -7650,11 +6936,11 @@ next case NotR then show ?case - by(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget) + by(auto simp: 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(auto simp: 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)+ @@ -7693,11 +6979,11 @@ next case AndR then show ?case - by(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget) + by(auto simp: 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(auto simp: 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)+ @@ -7736,7 +7022,7 @@ next case (AndL2 u M' v) then show ?case - apply(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget) + apply(auto simp: 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)+ @@ -7775,15 +7061,15 @@ next case OrR1 then show ?case - by(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget) + by(auto simp: 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) + by(auto simp: 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(auto simp: 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) @@ -7827,11 +7113,11 @@ next case ImpR then show ?case - by(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget) + by(auto simp: 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(auto simp: 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) @@ -7879,7 +7165,7 @@ proof(nominal_induct N avoiding: x y a c M P rule: trm.strong_induct) case (Ax z c) then show ?case - by (auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget) + by (auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget) next case (Cut d M' u M'') then show ?case @@ -7919,11 +7205,11 @@ next case NotL then show ?case - by(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget) + by(auto simp: 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(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget) apply(generate_fresh "coname") apply(fresh_fun_simp) apply(fresh_fun_simp) @@ -7937,7 +7223,7 @@ apply(rule trans) apply(rule substc.simps) apply(simp add: fresh_prod fresh_atm) - apply(auto simp add: fresh_atm fresh_prod)[1] + apply(auto simp: fresh_atm fresh_prod)[1] apply(generate_fresh "coname") apply(fresh_fun_simp) apply(fresh_fun_simp) @@ -7946,25 +7232,25 @@ 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(auto simp: 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) + apply(auto simp: fresh_prod fresh_atm) done next case AndL1 then show ?case - by(auto simp add: subst_fresh abs_fresh fresh_atm fresh_prod forget) + by(auto simp: 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) + by(auto simp: 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(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget) apply(generate_fresh "coname") apply(fresh_fun_simp) apply(fresh_fun_simp) @@ -7977,10 +7263,10 @@ 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: fresh_prod fresh_atm subst_fresh)[1] + apply(auto simp: fresh_prod fresh_atm subst_fresh)[1] apply(simp) - apply(auto simp add: fresh_atm fresh_prod)[1] + apply(auto simp: fresh_atm fresh_prod)[1] apply(generate_fresh "coname") apply(fresh_fun_simp) apply(fresh_fun_simp) @@ -7989,23 +7275,23 @@ 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(auto simp: 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: fresh_prod fresh_atm subst_fresh)[1] + apply(auto simp: fresh_prod fresh_atm subst_fresh)[1] apply(simp) - apply(auto simp add: fresh_atm fresh_prod)[1] + apply(auto simp: 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) + by(auto simp: 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(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget) apply(generate_fresh "coname") apply(fresh_fun_simp) apply(fresh_fun_simp) @@ -8018,8 +7304,8 @@ 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: fresh_prod fresh_atm subst_fresh)[1] + apply(auto simp: fresh_prod fresh_atm subst_fresh)[1] apply(generate_fresh "coname") apply(fresh_fun_simp) apply(fresh_fun_simp) @@ -8028,17 +7314,17 @@ 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(auto simp: 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: fresh_prod fresh_atm subst_fresh)[1] + apply(auto simp: 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(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget) apply(generate_fresh "coname") apply(fresh_fun_simp) apply(fresh_fun_simp) @@ -8051,8 +7337,8 @@ 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: fresh_prod fresh_atm subst_fresh)[1] + apply(auto simp: fresh_prod fresh_atm subst_fresh)[1] apply(generate_fresh "coname") apply(fresh_fun_simp) apply(fresh_fun_simp) @@ -8061,21 +7347,21 @@ 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(auto simp: 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: fresh_prod fresh_atm subst_fresh)[1] + apply(auto simp: 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) + by(auto simp: 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(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget) apply(generate_fresh "coname") apply(fresh_fun_simp) apply(fresh_fun_simp) @@ -8088,9 +7374,9 @@ 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(auto simp: fresh_prod fresh_atm subst_fresh)[1] + apply(auto simp: fresh_prod fresh_atm subst_fresh)[1] + apply(auto simp: 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) @@ -8099,14 +7385,14 @@ 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(auto simp: 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] + apply(auto simp: fresh_prod fresh_atm subst_fresh)[1] + apply(auto simp: fresh_prod fresh_atm subst_fresh)[1] + apply(auto simp: fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp)[1] + apply(auto simp: fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp)[1] done qed diff -r 5e64a54f6790 -r d9b8831a6a99 src/HOL/Nominal/Examples/Compile.thy --- a/src/HOL/Nominal/Examples/Compile.thy Thu May 02 15:40:05 2024 +0200 +++ b/src/HOL/Nominal/Examples/Compile.thy Fri May 03 00:07:51 2024 +0200 @@ -110,11 +110,7 @@ | "\z\x; x\y; x\e; x\e2; z\y; z\e; z\e1; x\t'; z\t'\ \ (Case e of inl x \ e1 | inr z \ e2)[y::=t'] = (Case (e[y::=t']) of inl x \ (e1[y::=t']) | inr z \ (e2[y::=t']))" - apply(finite_guess)+ - apply(rule TrueI)+ - apply(simp add: abs_fresh)+ - apply(fresh_guess)+ - done + by(finite_guess | simp add: abs_fresh | fresh_guess)+ instance .. @@ -135,11 +131,7 @@ | "(IRef e)[y::=t'] = IRef (e[y::=t'])" | "(ISeq e1 e2)[y::=t'] = ISeq (e1[y::=t']) (e2[y::=t'])" | "(Iif e e1 e2)[y::=t'] = Iif (e[y::=t']) (e1[y::=t']) (e2[y::=t'])" - apply(finite_guess)+ - apply(rule TrueI)+ - apply(simp add: abs_fresh)+ - apply(fresh_guess)+ - done + by(finite_guess | simp add: abs_fresh | fresh_guess)+ instance .. @@ -151,33 +143,28 @@ and t2::"trmI" and x::"name" shows "pi\(t1[x::=t2]) = ((pi\t1)[(pi\x)::=(pi\t2)])" - apply (nominal_induct t1 avoiding: x t2 rule: trmI.strong_induct) - apply (simp_all add: subst_trmI.simps eqvts fresh_bij) - done + by (nominal_induct t1 avoiding: x t2 rule: trmI.strong_induct) + (simp_all add: subst_trmI.simps eqvts fresh_bij) lemma Isubst_supp: fixes t1::"trmI" and t2::"trmI" and x::"name" shows "((supp (t1[x::=t2]))::name set) \ (supp t2)\((supp t1)-{x})" - apply (nominal_induct t1 avoiding: x t2 rule: trmI.strong_induct) - apply (auto simp add: subst_trmI.simps trmI.supp supp_atm abs_supp supp_nat) - apply blast+ - done +proof (nominal_induct t1 avoiding: x t2 rule: trmI.strong_induct) + case (IVar name) + then show ?case + by (simp add: supp_atm trmI.supp(1)) +qed (fastforce simp add: subst_trmI.simps trmI.supp supp_atm abs_supp supp_nat)+ lemma Isubst_fresh: fixes x::"name" and y::"name" and t1::"trmI" and t2::"trmI" - assumes a: "x\[y].t1" "x\t2" + assumes "x\[y].t1" "x\t2" shows "x\(t1[y::=t2])" -using a -apply(auto simp add: fresh_def Isubst_supp) -apply(drule rev_subsetD) -apply(rule Isubst_supp) -apply(simp add: abs_supp) -done + using assms Isubst_supp abs_supp(2) unfolding fresh_def Isubst_supp by fastforce text \big-step evaluation for trms\ @@ -240,11 +227,8 @@ let v1 = (trans e1) in let v2 = (trans e2) in Iif (IRef (ISucc v)) (v2[x2::=IRef (ISucc (ISucc v))]) (v1[x1::=IRef (ISucc (ISucc v))]))" - apply(finite_guess add: Let_def)+ - apply(rule TrueI)+ - apply(simp add: abs_fresh Isubst_fresh)+ - apply(fresh_guess add: Let_def)+ - done + unfolding Let_def + by(finite_guess | simp add: abs_fresh Isubst_fresh | fresh_guess)+ nominal_primrec trans_type :: "ty \ tyI" diff -r 5e64a54f6790 -r d9b8831a6a99 src/HOL/Nominal/Examples/Crary.thy --- a/src/HOL/Nominal/Examples/Crary.thy Thu May 02 15:40:05 2024 +0200 +++ b/src/HOL/Nominal/Examples/Crary.thy Fri May 03 00:07:51 2024 +0200 @@ -76,7 +76,7 @@ lemma lookup_eqvt[eqvt]: fixes pi::"name prm" shows "pi\(lookup \ x) = lookup (pi\\) (pi\x)" -by (induct \) (auto simp add: perm_bij) +by (induct \) (auto simp: perm_bij) lemma lookup_fresh: fixes z::"name" @@ -84,14 +84,14 @@ shows "z\ lookup \ x" using a by (induct rule: lookup.induct) - (auto simp add: fresh_list_cons) + (auto simp: fresh_list_cons) lemma lookup_fresh': assumes a: "z\\" shows "lookup \ z = Var z" using a by (induct rule: lookup.induct) - (auto simp add: fresh_list_cons fresh_prod fresh_atm) + (auto simp: fresh_list_cons fresh_prod fresh_atm) nominal_primrec psubst :: "Subst \ trm \ trm" ("_<_>" [100,100] 130) @@ -101,11 +101,7 @@ | "x\\ \ \<(Lam [x].t)> = Lam [x].(\)" | "\<(Const n)> = Const n" | "\<(Unit)> = Unit" -apply(finite_guess)+ -apply(rule TrueI)+ -apply(simp add: abs_fresh)+ -apply(fresh_guess)+ -done + by(finite_guess | simp add: abs_fresh | fresh_guess)+ abbreviation subst :: "trm \ name \ trm \ trm" ("_[_::=_]" [100,100,100] 100) @@ -131,9 +127,8 @@ assumes a: "c\t\<^sub>1" shows "t\<^sub>1[a::=t\<^sub>2] = ([(c,a)]\t\<^sub>1)[c::=t\<^sub>2]" using a -apply(nominal_induct t\<^sub>1 avoiding: a c t\<^sub>2 rule: trm.strong_induct) -apply(simp add: trm.inject calc_atm fresh_atm abs_fresh perm_nat_def)+ -done + by (nominal_induct t\<^sub>1 avoiding: a c t\<^sub>2 rule: trm.strong_induct) + (auto simp: trm.inject calc_atm fresh_atm abs_fresh perm_nat_def) lemma fresh_psubst: fixes z::"name" @@ -141,7 +136,7 @@ shows "z\(\)" using a by (nominal_induct t avoiding: z \ t rule: trm.strong_induct) - (auto simp add: abs_fresh lookup_fresh) + (auto simp: abs_fresh lookup_fresh) lemma fresh_subst'': fixes z::"name" @@ -149,7 +144,7 @@ shows "z\t\<^sub>1[z::=t\<^sub>2]" using assms by (nominal_induct t\<^sub>1 avoiding: t\<^sub>2 z rule: trm.strong_induct) - (auto simp add: abs_fresh fresh_nat fresh_atm) + (auto simp: abs_fresh fresh_nat fresh_atm) lemma fresh_subst': fixes z::"name" @@ -157,14 +152,14 @@ shows "z\t\<^sub>1[y::=t\<^sub>2]" using assms by (nominal_induct t\<^sub>1 avoiding: y t\<^sub>2 z rule: trm.strong_induct) - (auto simp add: abs_fresh fresh_nat fresh_atm) + (auto simp: abs_fresh fresh_nat fresh_atm) lemma fresh_subst: fixes z::"name" assumes a: "z\t\<^sub>1" "z\t\<^sub>2" shows "z\t\<^sub>1[y::=t\<^sub>2]" using a -by (auto simp add: fresh_subst' abs_fresh) +by (auto simp: fresh_subst' abs_fresh) lemma fresh_psubst_simp: assumes "x\t" @@ -181,7 +176,7 @@ by (simp add: fresh_list_cons fresh_prod) moreover have " \ = Lam [y]. (\)" using fs by simp ultimately show "((x,u)#\) = \" by auto -qed (auto simp add: fresh_atm abs_fresh) +qed (auto simp: fresh_atm abs_fresh) lemma forget: fixes x::"name" @@ -189,7 +184,7 @@ shows "t[x::=t'] = t" using a by (nominal_induct t avoiding: x t' rule: trm.strong_induct) - (auto simp add: fresh_atm abs_fresh) + (auto simp: fresh_atm abs_fresh) lemma subst_fun_eq: fixes u::trm @@ -212,20 +207,20 @@ lemma psubst_empty[simp]: shows "[] = t" by (nominal_induct t rule: trm.strong_induct) - (auto simp add: fresh_list_nil) + (auto simp: fresh_list_nil) lemma psubst_subst_psubst: assumes h:"c\\" shows "\[c::=s] = ((c,s)#\)" using h by (nominal_induct t avoiding: \ c s rule: trm.strong_induct) - (auto simp add: fresh_list_cons fresh_atm forget lookup_fresh lookup_fresh' fresh_psubst) + (auto simp: fresh_list_cons fresh_atm forget lookup_fresh lookup_fresh' fresh_psubst) lemma subst_fresh_simp: assumes a: "x\\" shows "\ = Var x" using a -by (induct \ arbitrary: x) (auto simp add:fresh_list_cons fresh_prod fresh_atm) +by (induct \ arbitrary: x) (auto simp:fresh_list_cons fresh_prod fresh_atm) lemma psubst_subst_propagate: assumes "x\\" @@ -239,7 +234,7 @@ } moreover { assume h:"x\n" - then have "x\Var n" by (auto simp add: fresh_atm) + then have "x\Var n" by (auto simp: fresh_atm) moreover have "x\\" by fact ultimately have "x\\" using fresh_psubst by blast then have " \[x::=\] = \" using forget by auto @@ -291,7 +286,7 @@ shows "\(\\::ty. (a,\)\set \)" using assms by (induct \) - (auto simp add: fresh_prod fresh_list_cons fresh_atm) + (auto simp: fresh_prod fresh_list_cons fresh_atm) lemma type_unicity_in_context: assumes a: "valid \" @@ -415,15 +410,7 @@ and b: "x \ b" shows "a=b" using a b -apply (induct arbitrary: b) -apply (erule whr_inv_auto(3)) -apply (clarify) -apply (rule subst_fun_eq) -apply (simp) -apply (force) -apply (erule whr_inv_auto(6)) -apply (blast)+ -done +by (induct arbitrary: b) (use subst_fun_eq in blast)+ lemma nf_unicity : assumes "x \ a" and "x \ b" @@ -433,8 +420,7 @@ case (QAN_Reduce x t a b) have h:"x \ t" "t \ a" by fact+ have ih:"\b. t \ b \ a = b" by fact - have "x \ b" by fact - then obtain t' where "x \ t'" and hl:"t' \ b" using h by auto + obtain t' where "x \ t'" and hl:"t' \ b" using h \x \ b\ by auto then have "t=t'" using h red_unicity by auto then show "a=b" using ih hl by auto qed (auto) @@ -535,7 +521,7 @@ shows "\ \ s \ t : T \ \ \ t \ s : T" and "\ \ s \ t : T \ \ \ t \ s : T" by (induct rule: alg_equiv_alg_path_equiv.inducts) - (auto simp add: fresh_prod) + (auto simp: fresh_prod) lemma algorithmic_transitivity: shows "\ \ s \ t : T \ \ \ t \ u : T \ \ \ s \ u : T" @@ -560,7 +546,7 @@ then have "(x,T\<^sub>1)#\ \ App t (Var x) \ App u (Var x) : T\<^sub>2" using fs by (simp add: Q_Arrow_strong_inversion) with ih have "(x,T\<^sub>1)#\ \ App s (Var x) \ App u (Var x) : T\<^sub>2" by simp - then show "\ \ s \ u : T\<^sub>1\T\<^sub>2" using fs by (auto simp add: fresh_prod) + then show "\ \ s \ u : T\<^sub>1\T\<^sub>2" using fs by (auto simp: fresh_prod) next case (QAP_App \ p q T\<^sub>1 T\<^sub>2 s t u) have "\ \ App q t \ u : T\<^sub>2" by fact @@ -578,10 +564,9 @@ lemma algorithmic_weak_head_closure: shows "\ \ s \ t : T \ s' \ s \ t' \ t \ \ \ s' \ t' : T" -apply (nominal_induct \ s t T avoiding: s' t' - rule: alg_equiv_alg_path_equiv.strong_inducts(1) [of _ _ _ _ "%a b c d e. True"]) -apply(auto intro!: QAT_Arrow) -done +proof (nominal_induct \ s t T avoiding: s' t' + rule: alg_equiv_alg_path_equiv.strong_inducts(1) [of _ _ _ _ "\a b c d e. True"]) +qed(auto intro!: QAT_Arrow) lemma algorithmic_monotonicity: shows "\ \ s \ t : T \ \ \ \' \ valid \' \ \' \ s \ t : T" @@ -596,7 +581,7 @@ moreover have sub: "(x,T\<^sub>1)#\ \ (x,T\<^sub>1)#\'" using h2 by auto ultimately have "(x,T\<^sub>1)#\' \ App s (Var x) \ App t (Var x) : T\<^sub>2" using ih by simp - then show "\' \ s \ t : T\<^sub>1\T\<^sub>2" using fs by (auto simp add: fresh_prod) + then show "\' \ s \ t : T\<^sub>1\T\<^sub>2" using fs by (auto simp: fresh_prod) qed (auto) lemma path_equiv_implies_nf: @@ -613,11 +598,7 @@ | "\ \ s is t : TBase = \ \ s \ t : TBase" | "\ \ s is t : (T\<^sub>1 \ T\<^sub>2) = (\\' s' t'. \\\' \ valid \' \ \' \ s' is t' : T\<^sub>1 \ (\' \ (App s s') is (App t t') : T\<^sub>2))" -apply (auto simp add: ty.inject) -apply (subgoal_tac "(\T\<^sub>1 T\<^sub>2. b=T\<^sub>1 \ T\<^sub>2) \ b=TUnit \ b=TBase" ) -apply (force) -apply (rule ty_cases) -done +using ty_cases by (force simp: ty.inject)+ termination by lexicographic_order @@ -656,7 +637,7 @@ then have "(x,T\<^sub>1)#\ \ Var x is Var x : T\<^sub>1" using ih2 by auto then have "(x,T\<^sub>1)#\ \ App s (Var x) is App t (Var x) : T\<^sub>2" using h v by auto then have "(x,T\<^sub>1)#\ \ App s (Var x) \ App t (Var x) : T\<^sub>2" using ih1 v by auto - then show "\ \ s \ t : T\<^sub>1\T\<^sub>2" using fs by (auto simp add: fresh_prod) + then show "\ \ s \ t : T\<^sub>1\T\<^sub>2" using fs by (auto simp: fresh_prod) next case (2 \ p q) have h: "\ \ p \ q : T\<^sub>1\T\<^sub>2" by fact @@ -693,7 +674,7 @@ shows "\ \ t is s : T" using a by (nominal_induct arbitrary: \ s t rule: ty.strong_induct) - (auto simp add: algorithmic_symmetry) + (auto simp: algorithmic_symmetry) lemma logical_transitivity: assumes "\ \ s is t : T" "\ \ t is u : T" @@ -726,8 +707,13 @@ and c: "t' \ t" shows "\ \ s' is t' : T" using a b c algorithmic_weak_head_closure -by (nominal_induct arbitrary: \ s t s' t' rule: ty.strong_induct) - (auto, blast) +proof (nominal_induct arbitrary: \ s t s' t' rule: ty.strong_induct) +next + case (Arrow ty1 ty2) + then show ?case + by (smt (verit, ccfv_threshold) QAR_App log_equiv.simps(3)) +qed auto + lemma logical_weak_head_closure': assumes "\ \ s is t : T" and "s' \ s" @@ -764,11 +750,8 @@ lemma logical_pseudo_reflexivity: assumes "\' \ t is s over \" - shows "\' \ s is s over \" -proof - - from assms have "\' \ s is t over \" using logical_symmetry by blast - with assms show "\' \ s is s over \" using logical_transitivity by blast -qed + shows "\' \ s is s over \" + by (meson assms logical_symmetry logical_transitivity) lemma logical_subst_monotonicity : fixes \ \' \'' :: Ctxt @@ -796,8 +779,8 @@ moreover { assume hl:"(y,U) \ set \" - then have "\ y\\" by (induct \) (auto simp add: fresh_list_cons fresh_atm fresh_prod) - then have hf:"x\ Var y" using fs by (auto simp add: fresh_atm) + then have "\ y\\" by (induct \) (auto simp: fresh_list_cons fresh_atm fresh_prod) + then have hf:"x\ Var y" using fs by (auto simp: fresh_atm) then have "((x,s)#\) = \" "((x,t)#\') = \'" using fresh_psubst_simp by blast+ moreover have "\' \ \ is \' : U" using h1 hl by auto @@ -847,18 +830,11 @@ using h1 h2 h3 proof (nominal_induct \ s t T avoiding: \' \ \' rule: def_equiv.strong_induct) case (Q_Refl \ t T \' \ \') - have "\ \ t : T" - and "valid \'" by fact+ - moreover - have "\' \ \ is \' over \" by fact - ultimately show "\' \ \ is \' : T" using fundamental_theorem_1 by blast + then show "\' \ \ is \' : T" using fundamental_theorem_1 + by blast next case (Q_Symm \ t s T \' \ \') - have "\' \ \ is \' over \" - and "valid \'" by fact+ - moreover - have ih: "\ \' \ \'. \\' \ \ is \' over \; valid \'\ \ \' \ \ is \' : T" by fact - ultimately show "\' \ \ is \' : T" using logical_symmetry by blast + then show "\' \ \ is \' : T" using logical_symmetry by blast next case (Q_Trans \ s t T u \' \ \') have ih1: "\ \' \ \'. \\' \ \ is \' over \; valid \'\ \ \' \ \ is \' : T" by fact @@ -943,13 +919,8 @@ shows "\ \ s \ t : T" proof - have val: "valid \" using def_equiv_implies_valid asm by simp - moreover - { - fix x T - assume "(x,T) \ set \" "valid \" - then have "\ \ Var x is Var x : T" using main_lemma(2) by blast - } - ultimately have "\ \ [] is [] over \" by auto + then have "\ \ [] is [] over \" + by (simp add: QAP_Var main_lemma(2)) then have "\ \ [] is [] : T" using fundamental_theorem_2 val asm by blast then have "\ \ s is t : T" by simp then show "\ \ s \ t : T" using main_lemma(1) val by simp diff -r 5e64a54f6790 -r d9b8831a6a99 src/HOL/Nominal/Examples/Fsub.thy --- a/src/HOL/Nominal/Examples/Fsub.thy Thu May 02 15:40:05 2024 +0200 +++ b/src/HOL/Nominal/Examples/Fsub.thy Fri May 03 00:07:51 2024 +0200 @@ -145,7 +145,7 @@ lemma finite_doms: shows "finite (ty_dom \)" and "finite (trm_dom \)" -by (induct \) (auto simp add: finite_vrs) +by (induct \) (auto simp: finite_vrs) lemma ty_dom_supp: shows "(supp (ty_dom \)) = (ty_dom \)" @@ -167,12 +167,12 @@ assumes a: "X\(ty_dom \)" shows "\T.(TVarB X T)\set \" using a - apply (induct \, auto) - apply (rename_tac a \') - apply (subgoal_tac "\T.(TVarB X T=a)") - apply (auto) - apply (auto simp add: ty_binding_existence) -done +proof (induction \) + case Nil then show ?case by simp +next + case (Cons a \) then show ?case + using ty_binding_existence by fastforce +qed lemma doms_append: shows "ty_dom (\@\) = ((ty_dom \) \ (ty_dom \))" @@ -183,18 +183,19 @@ fixes pi::"vrs prm" and S::"ty" shows "pi\S = S" -by (induct S rule: ty.induct) (auto simp add: calc_atm) +by (induct S rule: ty.induct) (auto simp: calc_atm) lemma fresh_ty_dom_cons: fixes X::"tyvrs" shows "X\(ty_dom (Y#\)) = (X\(tyvrs_of Y) \ X\(ty_dom \))" - apply (nominal_induct rule:binding.strong_induct) - apply (auto) - apply (simp add: fresh_def supp_def eqvts) - apply (simp add: fresh_fin_insert [OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst] finite_doms) - apply (simp add: fresh_def supp_def eqvts) - apply (simp add: fresh_fin_insert [OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst] finite_doms)+ - done +proof (nominal_induct rule:binding.strong_induct) + case (VarB vrs ty) + then show ?case by auto +next + case (TVarB tyvrs ty) + then show ?case + by (simp add: at_fin_set_supp at_tyvrs_inst finite_doms(1) fresh_def supp_atm(1)) +qed lemma tyvrs_fresh: fixes X::"tyvrs" @@ -202,21 +203,19 @@ shows "X \ tyvrs_of a" and "X \ vrs_of a" using assms - apply (nominal_induct a rule:binding.strong_induct) - apply (auto) - apply (fresh_guess)+ -done + by (nominal_induct a rule:binding.strong_induct) (force simp: fresh_singleton)+ lemma fresh_dom: fixes X::"tyvrs" assumes a: "X\\" shows "X\(ty_dom \)" using a -apply(induct \) -apply(simp add: fresh_set_empty) -apply(simp only: fresh_ty_dom_cons) -apply(auto simp add: fresh_prod fresh_list_cons tyvrs_fresh) -done +proof (induct \) + case Nil then show ?case by auto +next + case (Cons a \) then show ?case + by (meson fresh_list_cons fresh_ty_dom_cons tyvrs_fresh(1)) +qed text \Not all lists of type \<^typ>\env\ are well-formed. One condition requires that in \<^term>\TVarB X S#\\ all free variables of \<^term>\S\ must be @@ -240,10 +239,7 @@ lemma tyvrs_vrs_prm_simp: fixes pi::"vrs prm" shows "tyvrs_of (pi\a) = tyvrs_of a" - apply (nominal_induct rule:binding.strong_induct) - apply (simp_all add: eqvts) - apply (simp add: dj_perm_forget[OF dj_tyvrs_vrs]) - done + by (nominal_induct rule:binding.strong_induct) (auto simp: vrs_prm_tyvrs_def) lemma ty_vrs_fresh: fixes x::"vrs" @@ -255,17 +251,13 @@ fixes pi::"vrs prm" and \::"env" shows "(ty_dom (pi\\)) = (ty_dom \)" - apply(induct \) - apply (simp add: eqvts) - apply(simp add: tyvrs_vrs_prm_simp) -done +by (induct \) (auto simp: tyvrs_vrs_prm_simp) lemma closed_in_eqvt'[eqvt]: fixes pi::"vrs prm" assumes a: "S closed_in \" shows "(pi\S) closed_in (pi\\)" -using a -by (simp add: closed_in_def ty_dom_vrs_prm_simp ty_vrs_prm_simp) + using assms closed_in_def ty_dom_vrs_prm_simp ty_vrs_prm_simp by force lemma fresh_vrs_of: fixes x::"vrs" @@ -277,12 +269,12 @@ fixes x::"vrs" shows "x\ trm_dom \ = x\\" by (induct \) - (simp_all add: fresh_set_empty fresh_list_cons + (simp_all add: fresh_list_cons fresh_fin_union [OF pt_vrs_inst at_vrs_inst fs_vrs_inst] - finite_doms finite_vrs fresh_vrs_of fresh_list_nil) + finite_doms finite_vrs fresh_vrs_of) lemma closed_in_fresh: "(X::tyvrs) \ ty_dom \ \ T closed_in \ \ X \ T" - by (auto simp add: closed_in_def fresh_def ty_dom_supp) + by (auto simp: closed_in_def fresh_def ty_dom_supp) text \Now validity of a context is a straightforward inductive definition.\ @@ -313,8 +305,7 @@ proof (induct \) case (Cons a \') then show ?case - by (nominal_induct a rule:binding.strong_induct) - (auto elim: validE) + by (nominal_induct a rule:binding.strong_induct) auto qed (auto) lemma replace_type: @@ -324,12 +315,12 @@ using a b proof(induct \) case Nil - then show ?case by (auto elim: validE intro: valid_cons simp add: doms_append closed_in_def) + then show ?case by (auto intro: valid_cons simp add: doms_append closed_in_def) next case (Cons a \') then show ?case by (nominal_induct a rule:binding.strong_induct) - (auto elim: validE intro!: valid_cons simp add: doms_append closed_in_def) + (auto intro!: valid_cons simp add: doms_append closed_in_def) qed text \Well-formed contexts have a unique type-binding for a type-variable.\ @@ -353,10 +344,10 @@ by (simp add: fresh_ty_dom_cons fresh_fin_union[OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst] finite_vrs finite_doms, - auto simp add: fresh_atm fresh_singleton) + auto simp: fresh_atm fresh_singleton) qed (simp) } - ultimately show "T=S" by (auto simp add: binding.inject) + ultimately show "T=S" by (auto simp: binding.inject) qed (auto) lemma uniqueness_of_ctxt': @@ -377,10 +368,10 @@ thus "\ (\T.(VarB x' T)\set(y#\'))" by (simp add: fresh_fin_union[OF pt_vrs_inst at_vrs_inst fs_vrs_inst] finite_vrs finite_doms, - auto simp add: fresh_atm fresh_singleton) + auto simp: fresh_atm fresh_singleton) qed (simp) } - ultimately show "T=S" by (auto simp add: binding.inject) + ultimately show "T=S" by (auto simp: binding.inject) qed (auto) section \Size and Capture-Avoiding Substitution for Types\ @@ -392,11 +383,7 @@ | "size_ty (Top) = 1" | "size_ty (T1 \ T2) = (size_ty T1) + (size_ty T2) + 1" | "X \ T1 \ size_ty (\X<:T1. T2) = (size_ty T1) + (size_ty T2) + 1" - apply (finite_guess)+ - apply (rule TrueI)+ - apply (simp add: fresh_nat) - apply (fresh_guess)+ - done + by (finite_guess | fresh_guess | simp)+ nominal_primrec subst_ty :: "ty \ tyvrs \ ty \ ty" ("_[_ \ _]\<^sub>\" [300, 0, 0] 300) @@ -405,11 +392,7 @@ | "(Top)[Y \ T]\<^sub>\ = Top" | "(T\<^sub>1 \ T\<^sub>2)[Y \ T]\<^sub>\ = T\<^sub>1[Y \ T]\<^sub>\ \ T\<^sub>2[Y \ T]\<^sub>\" | "X\(Y,T,T\<^sub>1) \ (\X<:T\<^sub>1. T\<^sub>2)[Y \ T]\<^sub>\ = (\X<:T\<^sub>1[Y \ T]\<^sub>\. T\<^sub>2[Y \ T]\<^sub>\)" - apply (finite_guess)+ - apply (rule TrueI)+ - apply (simp add: abs_fresh) - apply (fresh_guess)+ - done + by (finite_guess | fresh_guess | simp add: abs_fresh)+ lemma subst_eqvt[eqvt]: fixes pi::"tyvrs prm" @@ -429,16 +412,16 @@ fixes X::"tyvrs" assumes "X\T" and "X\P" shows "X\T[Y \ P]\<^sub>\" -using assms -by (nominal_induct T avoiding: X Y P rule:ty.strong_induct) - (auto simp add: abs_fresh) + using assms + by (nominal_induct T avoiding: X Y P rule:ty.strong_induct) + (auto simp: abs_fresh) lemma fresh_type_subst_fresh: - assumes "X\T'" - shows "X\T[X \ T']\<^sub>\" -using assms -by (nominal_induct T avoiding: X T' rule: ty.strong_induct) - (auto simp add: fresh_atm abs_fresh fresh_nat) + assumes "X\T'" + shows "X\T[X \ T']\<^sub>\" + using assms + by (nominal_induct T avoiding: X T' rule: ty.strong_induct) + (auto simp: fresh_atm abs_fresh) lemma type_subst_identity: "X\T \ T[X \ U]\<^sub>\ = T" @@ -448,7 +431,7 @@ lemma type_substitution_lemma: "X \ Y \ X\L \ M[X \ N]\<^sub>\[Y \ L]\<^sub>\ = M[Y \ L]\<^sub>\[X \ N[Y \ L]\<^sub>\]\<^sub>\" by (nominal_induct M avoiding: X Y N L rule: ty.strong_induct) - (auto simp add: type_subst_fresh type_subst_identity) + (auto simp: type_subst_fresh type_subst_identity) lemma type_subst_rename: "Y\T \ ([(Y,X)]\T)[Y \ U]\<^sub>\ = T[X \ U]\<^sub>\" @@ -469,7 +452,7 @@ shows "X\a[Y \ P]\<^sub>b" using assms by (nominal_induct a rule: binding.strong_induct) - (auto simp add: type_subst_fresh) + (auto simp: type_subst_fresh) lemma binding_subst_identity: shows "X\B \ B[X \ U]\<^sub>b = B" @@ -487,7 +470,7 @@ shows "X\\[Y \ P]\<^sub>e" using assms by (induct \) - (auto simp add: fresh_list_cons binding_subst_fresh) + (auto simp: fresh_list_cons binding_subst_fresh) lemma ctxt_subst_mem_TVarB: "TVarB X T \ set \ \ TVarB X (T[Y \ U]\<^sub>\) \ set (\[Y \ U]\<^sub>e)" by (induct \) auto @@ -509,17 +492,13 @@ | "(t \\<^sub>\ T)[y \ t'] = t[y \ t'] \\<^sub>\ T" | "X\(T,t') \ (\X<:T. t)[y \ t'] = (\X<:T. t[y \ t'])" | "x\(y,t') \ (\x:T. t)[y \ t'] = (\x:T. t[y \ t'])" -apply(finite_guess)+ -apply(rule TrueI)+ -apply(simp add: abs_fresh)+ -apply(fresh_guess add: ty_vrs_fresh abs_fresh)+ -done + by (finite_guess | simp add: abs_fresh | fresh_guess add: ty_vrs_fresh abs_fresh)+ lemma subst_trm_fresh_tyvar: fixes X::"tyvrs" shows "X\t \ X\u \ X\t[x \ u]" by (nominal_induct t avoiding: x u rule: trm.strong_induct) - (auto simp add: trm.fresh abs_fresh) + (auto simp: abs_fresh) lemma subst_trm_fresh_var: "x\u \ x\t[x \ u]" @@ -553,10 +532,7 @@ | "(t1 \\<^sub>\ T)[Y \\<^sub>\ T2] = t1[Y \\<^sub>\ T2] \\<^sub>\ T[Y \ T2]\<^sub>\" | "X\(Y,T,T2) \ (\X<:T. t)[Y \\<^sub>\ T2] = (\X<:T[Y \ T2]\<^sub>\. t[Y \\<^sub>\ T2])" | "(\x:T. t)[Y \\<^sub>\ T2] = (\x:T[Y \ T2]\<^sub>\. t[Y \\<^sub>\ T2])" -apply(finite_guess)+ -apply(rule TrueI)+ -apply(simp add: abs_fresh ty_vrs_fresh)+ -apply(simp add: type_subst_fresh) +apply(finite_guess | simp add: abs_fresh ty_vrs_fresh type_subst_fresh)+ apply(fresh_guess add: ty_vrs_fresh abs_fresh)+ done @@ -564,7 +540,7 @@ fixes X::"tyvrs" shows "X\t \ X\T \ X\t[Y \\<^sub>\ T]" by (nominal_induct t avoiding: Y T rule: trm.strong_induct) - (auto simp add: abs_fresh type_subst_fresh) + (auto simp: abs_fresh type_subst_fresh) lemma subst_trm_ty_fresh': "X\T \ X\t[X \\<^sub>\ T]" @@ -633,7 +609,7 @@ have "S closed_in \ \ T closed_in \" by fact hence "T closed_in \" by force ultimately show "(Tvar X) closed_in \ \ T closed_in \" by simp -qed (auto simp add: closed_in_def ty.supp supp_atm abs_supp) +qed (auto simp: closed_in_def ty.supp supp_atm abs_supp) lemma subtype_implies_fresh: fixes X::"tyvrs" @@ -647,7 +623,7 @@ from a1 have "S closed_in \ \ T closed_in \" by (rule subtype_implies_closed) hence "supp S \ ((supp (ty_dom \))::tyvrs set)" and "supp T \ ((supp (ty_dom \))::tyvrs set)" by (simp_all add: ty_dom_supp closed_in_def) - ultimately show "X\S \ X\T" by (force simp add: supp_prod fresh_def) + ultimately show "X\S \ X\T" by (force simp: supp_prod fresh_def) qed lemma valid_ty_dom_fresh: @@ -655,19 +631,26 @@ assumes valid: "\ \ ok" shows "X\(ty_dom \) = X\\" using valid - apply induct - apply (simp add: fresh_list_nil fresh_set_empty) - apply (simp_all add: binding.fresh fresh_list_cons - fresh_fin_insert [OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst] finite_doms fresh_atm) - apply (auto simp add: closed_in_fresh) - done +proof induct + case valid_nil + then show ?case by auto +next + case (valid_consT \ X T) + then show ?case + by (auto simp: fresh_list_cons closed_in_fresh + fresh_fin_insert [OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst] finite_doms) +next + case (valid_cons \ x T) + then show ?case + using fresh_atm by (auto simp: fresh_list_cons closed_in_fresh) +qed equivariance subtype_of nominal_inductive subtype_of apply (simp_all add: abs_fresh) - apply (fastforce simp add: valid_ty_dom_fresh dest: subtype_implies_ok) - apply (force simp add: closed_in_fresh dest: subtype_implies_closed subtype_implies_ok)+ + apply (fastforce simp: valid_ty_dom_fresh dest: subtype_implies_ok) + apply (force simp: closed_in_fresh dest: subtype_implies_closed subtype_implies_ok)+ done section \Reflexivity of Subtyping\ @@ -685,7 +668,7 @@ hence fresh_ty_dom: "X\(ty_dom \)" by (simp add: fresh_dom) have "(\X<:T\<^sub>2. T\<^sub>1) closed_in \" by fact hence closed\<^sub>T2: "T\<^sub>2 closed_in \" and closed\<^sub>T1: "T\<^sub>1 closed_in ((TVarB X T\<^sub>2)#\)" - by (auto simp add: closed_in_def ty.supp abs_supp) + by (auto simp: closed_in_def ty.supp abs_supp) have ok: "\ \ ok" by fact hence ok': "\ ((TVarB X T\<^sub>2)#\) ok" using closed\<^sub>T2 fresh_ty_dom by simp have "\ \ T\<^sub>2 <: T\<^sub>2" using ih_T\<^sub>2 closed\<^sub>T2 ok by simp @@ -693,7 +676,7 @@ have "((TVarB X T\<^sub>2)#\) \ T\<^sub>1 <: T\<^sub>1" using ih_T\<^sub>1 closed\<^sub>T1 ok' by simp ultimately show "\ \ (\X<:T\<^sub>2. T\<^sub>1) <: (\X<:T\<^sub>2. T\<^sub>1)" using fresh_cond by (simp add: subtype_of.SA_all) -qed (auto simp add: closed_in_def ty.supp supp_atm) +qed (auto simp: closed_in_def ty.supp supp_atm) lemma subtype_reflexivity_semiautomated: assumes a: "\ \ ok" @@ -701,7 +684,7 @@ shows "\ \ T <: T" using a b apply(nominal_induct T avoiding: \ rule: ty.strong_induct) -apply(auto simp add: ty.supp abs_supp supp_atm closed_in_def) +apply(auto simp: ty.supp abs_supp supp_atm closed_in_def) \ \Too bad that this instantiation cannot be found automatically by \isakeyword{auto}; \isakeyword{blast} would find it if we had not used an explicit definition for \closed_in_def\.\ @@ -719,20 +702,15 @@ "\ extends \ \ \X Q. (TVarB X Q)\set \ \ (TVarB X Q)\set \" lemma extends_ty_dom: - assumes a: "\ extends \" + assumes "\ extends \" shows "ty_dom \ \ ty_dom \" - using a - apply (auto simp add: extends_def) - apply (drule ty_dom_existence) - apply (force simp add: ty_dom_inclusion) - done + using assms + by (meson extends_def subsetI ty_dom_existence ty_dom_inclusion) lemma extends_closed: - assumes a1: "T closed_in \" - and a2: "\ extends \" + assumes "T closed_in \" and "\ extends \" shows "T closed_in \" - using a1 a2 - by (auto dest: extends_ty_dom simp add: closed_in_def) + by (meson assms closed_in_def extends_ty_dom order.trans) lemma extends_memb: assumes a: "\ extends \" @@ -787,7 +765,7 @@ have "T\<^sub>1 closed_in \" using ext closed\<^sub>T1 by (simp only: extends_closed) hence "\ ((TVarB X T\<^sub>1)#\) ok" using fresh_dom ok by force moreover - have "((TVarB X T\<^sub>1)#\) extends ((TVarB X T\<^sub>1)#\)" using ext by (force simp add: extends_def) + have "((TVarB X T\<^sub>1)#\) extends ((TVarB X T\<^sub>1)#\)" using ext by (force simp: extends_def) ultimately have "((TVarB X T\<^sub>1)#\) \ S\<^sub>2 <: T\<^sub>2" using ih\<^sub>2 by simp moreover have "\ \ T\<^sub>1 <: S\<^sub>1" using ok ext ih\<^sub>1 by simp @@ -815,7 +793,7 @@ have "T\<^sub>1 closed_in \" using ext closed\<^sub>T1 by (simp only: extends_closed) hence "\ ((TVarB X T\<^sub>1)#\) ok" using fresh_dom ok by force moreover - have "((TVarB X T\<^sub>1)#\) extends ((TVarB X T\<^sub>1)#\)" using ext by (force simp add: extends_def) + have "((TVarB X T\<^sub>1)#\) extends ((TVarB X T\<^sub>1)#\)" using ext by (force simp: extends_def) ultimately have "((TVarB X T\<^sub>1)#\) \ S\<^sub>2 <: T\<^sub>2" using ih\<^sub>2 by simp moreover have "\ \ T\<^sub>1 <: S\<^sub>1" using ok ext ih\<^sub>1 by simp @@ -836,9 +814,8 @@ lemma S_ForallE_left: shows "\\ \ (\X<:S\<^sub>1. S\<^sub>2) <: T; X\\; X\S\<^sub>1; X\T\ \ T = Top \ (\T\<^sub>1 T\<^sub>2. T = (\X<:T\<^sub>1. T\<^sub>2) \ \ \ T\<^sub>1 <: S\<^sub>1 \ ((TVarB X T\<^sub>1)#\) \ S\<^sub>2 <: T\<^sub>2)" -apply(erule subtype_of.strong_cases[where X="X"]) -apply(auto simp add: abs_fresh ty.inject alpha) -done + using subtype_of.strong_cases[where X="X"] + by(force simp: abs_fresh ty.inject alpha) text \Next we prove the transitivity and narrowing for the subtyping-relation. The POPLmark-paper says the following: @@ -942,7 +919,7 @@ moreover have "S\<^sub>1 closed_in \" and "S\<^sub>2 closed_in ((TVarB X Q\<^sub>1)#\)" using lh_drv_prm\<^sub>1 lh_drv_prm\<^sub>2 by (simp_all add: subtype_implies_closed) - then have "(\X<:S\<^sub>1. S\<^sub>2) closed_in \" by (force simp add: closed_in_def ty.supp abs_supp) + then have "(\X<:S\<^sub>1. S\<^sub>2) closed_in \" by (force simp: closed_in_def ty.supp abs_supp) moreover have "\ \ ok" using rh_drv by (rule subtype_implies_ok) moreover @@ -1060,7 +1037,7 @@ by (rule exists_fresh) (rule fin_supp) then have "Y \ (\, t\<^sub>1, T2)" by simp moreover from Y have "(\X<:T11. T12) = (\Y<:T11. [(Y, X)] \ T12)" - by (auto simp add: ty.inject alpha' fresh_prod fresh_atm) + by (auto simp: ty.inject alpha' fresh_prod fresh_atm) with H1 have "\ \ t\<^sub>1 : (\Y<:T11. [(Y, X)] \ T12)" by simp ultimately have "\ \ t\<^sub>1 \\<^sub>\ T2 : (([(Y, X)] \ T12)[Y \ T2]\<^sub>\)" using H2 by (rule T_TApp) @@ -1079,7 +1056,7 @@ lemma ok_imp_VarB_closed_in: assumes ok: "\ \ ok" shows "VarB x T \ set \ \ T closed_in \" using ok - by induct (auto simp add: binding.inject closed_in_def) + by induct (auto simp: binding.inject closed_in_def) lemma tyvrs_of_subst: "tyvrs_of (B[X \ T]\<^sub>b) = tyvrs_of B" by (nominal_induct B rule: binding.strong_induct) simp_all @@ -1095,18 +1072,26 @@ lemma subst_closed_in: "T closed_in (\ @ TVarB X S # \) \ U closed_in \ \ T[X \ U]\<^sub>\ closed_in (\[X \ U]\<^sub>e @ \)" - apply (nominal_induct T avoiding: X U \ rule: ty.strong_induct) - apply (simp add: closed_in_def ty.supp supp_atm doms_append ty_dom_subst) - apply blast - apply (simp add: closed_in_def ty.supp) - apply (simp add: closed_in_def ty.supp) - apply (simp add: closed_in_def ty.supp abs_supp) - apply (drule_tac x = X in meta_spec) - apply (drule_tac x = U in meta_spec) - apply (drule_tac x = "(TVarB tyvrs ty2) # \" in meta_spec) - apply (simp add: doms_append ty_dom_subst) - apply blast - done +proof (nominal_induct T avoiding: X U \ rule: ty.strong_induct) + case (Tvar tyvrs) + then show ?case + by (auto simp add: closed_in_def ty.supp supp_atm doms_append ty_dom_subst) +next + case Top + then show ?case + using closed_in_def fresh_def by fastforce +next + case (Arrow ty1 ty2) + then show ?case + by (simp add: closed_in_def ty.supp) +next + case (Forall tyvrs ty1 ty2) + then have "supp (ty1[X \ U]\<^sub>\) \ ty_dom (\[X \ U]\<^sub>e @ TVarB tyvrs ty2 # \)" + apply (simp add: closed_in_def ty.supp abs_supp) + by (metis Diff_subset_conv Un_left_commute doms_append(1) le_supI2 ty_dom.simps(2) tyvrs_of.simps(2)) + with Forall show ?case + by (auto simp add: closed_in_def ty.supp abs_supp doms_append) +qed lemmas subst_closed_in' = subst_closed_in [where \="[]", simplified] @@ -1120,12 +1105,12 @@ show ?case by (rule ok_imp_VarB_closed_in) next case (T_App \ t\<^sub>1 T\<^sub>1 T\<^sub>2 t\<^sub>2) - then show ?case by (auto simp add: ty.supp closed_in_def) + then show ?case by (auto simp: ty.supp closed_in_def) next case (T_Abs x T\<^sub>1 \ t\<^sub>2 T\<^sub>2) from \VarB x T\<^sub>1 # \ \ t\<^sub>2 : T\<^sub>2\ have "T\<^sub>1 closed_in \" by (auto dest: typing_ok) - with T_Abs show ?case by (auto simp add: ty.supp closed_in_def) + with T_Abs show ?case by (auto simp: ty.supp closed_in_def) next case (T_Sub \ t S T) from \\ \ S <: T\ show ?case by (simp add: subtype_implies_closed) @@ -1133,11 +1118,11 @@ case (T_TAbs X T\<^sub>1 \ t\<^sub>2 T\<^sub>2) from \TVarB X T\<^sub>1 # \ \ t\<^sub>2 : T\<^sub>2\ have "T\<^sub>1 closed_in \" by (auto dest: typing_ok) - with T_TAbs show ?case by (auto simp add: ty.supp closed_in_def abs_supp) + with T_TAbs show ?case by (auto simp: ty.supp closed_in_def abs_supp) next case (T_TApp X \ t\<^sub>1 T2 T11 T12) then have "T12 closed_in (TVarB X T11 # \)" - by (auto simp add: closed_in_def ty.supp abs_supp) + by (auto simp: closed_in_def ty.supp abs_supp) moreover from T_TApp have "T2 closed_in \" by (simp add: subtype_implies_closed) ultimately show ?case by (rule subst_closed_in') @@ -1177,7 +1162,7 @@ then have "(\y:T11. [(y, x)] \ t12) \ v2 \ ([(y, x)] \ t12)[y \ v2]" using H by (rule E_Abs) moreover from y have "(\x:T11. t12) \ v2 = (\y:T11. [(y, x)] \ t12) \ v2" - by (auto simp add: trm.inject alpha' fresh_prod fresh_atm) + by (auto simp: trm.inject alpha' fresh_prod fresh_atm) ultimately have "(\x:T11. t12) \ v2 \ ([(y, x)] \ t12)[y \ v2]" by simp with y show ?thesis by (simp add: subst_trm_rename) @@ -1190,7 +1175,7 @@ then have "(\Y<:T11. [(Y, X)] \ t12) \\<^sub>\ T2 \ ([(Y, X)] \ t12)[Y \\<^sub>\ T2]" by (rule E_TAbs) moreover from Y have "(\X<:T11. t12) \\<^sub>\ T2 = (\Y<:T11. [(Y, X)] \ t12) \\<^sub>\ T2" - by (auto simp add: trm.inject alpha' fresh_prod fresh_atm) + by (auto simp: trm.inject alpha' fresh_prod fresh_atm) ultimately have "(\X<:T11. t12) \\<^sub>\ T2 \ ([(Y, X)] \ t12)[Y \\<^sub>\ T2]" by simp with Y show ?thesis by (simp add: subst_trm_ty_rename) @@ -1217,42 +1202,33 @@ using assms ty_dom_cons closed_in_def by auto lemma closed_in_weaken: "T closed_in (\ @ \) \ T closed_in (\ @ B # \)" - by (auto simp add: closed_in_def doms_append) + by (auto simp: closed_in_def doms_append) lemma closed_in_weaken': "T closed_in \ \ T closed_in (\ @ \)" - by (auto simp add: closed_in_def doms_append) + by (auto simp: closed_in_def doms_append) lemma valid_subst: assumes ok: "\ (\ @ TVarB X Q # \) ok" and closed: "P closed_in \" shows "\ (\[X \ P]\<^sub>e @ \) ok" using ok closed - apply (induct \) - apply simp_all - apply (erule validE) - apply assumption - apply (erule validE) - apply simp - apply (rule valid_consT) - apply assumption - apply (simp add: doms_append ty_dom_subst) - apply (simp add: fresh_fin_insert [OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst] finite_doms) - apply (rule_tac S=Q in subst_closed_in') - apply (simp add: closed_in_def doms_append ty_dom_subst) - apply (simp add: closed_in_def doms_append) - apply blast - apply simp - apply (rule valid_cons) - apply assumption - apply (simp add: doms_append trm_dom_subst) - apply (rule_tac S=Q in subst_closed_in') - apply (simp add: closed_in_def doms_append ty_dom_subst) - apply (simp add: closed_in_def doms_append) - apply blast - done +proof (induct \) + case Nil + then show ?case + by auto +next + case (Cons a \) + then have *: "\ (a # \ @ TVarB X Q # \) ok" + by fastforce + then show ?case + apply (rule validE) + using Cons + apply (simp add: at_tyvrs_inst closed doms_append(1) finite_doms(1) fresh_fin_insert fs_tyvrs_inst pt_tyvrs_inst subst_closed_in ty_dom_subst) + by (simp add: doms_append(2) subst_closed_in Cons.hyps closed trm_dom_subst) +qed lemma ty_dom_vrs: shows "ty_dom (G @ [VarB x Q] @ D) = ty_dom (G @ D)" -by (induct G) (auto) + by (induct G) (auto) lemma valid_cons': assumes "\ (\ @ VarB x Q # \) ok" @@ -1321,15 +1297,7 @@ then have closed: "T\<^sub>1 closed_in (\ @ \)" by (auto dest: typing_ok) have "\ (VarB y T\<^sub>1 # \ @ B # \) ok" - apply (rule valid_cons) - apply (rule T_Abs) - apply (simp add: doms_append - fresh_fin_insert [OF pt_vrs_inst at_vrs_inst fs_vrs_inst] - fresh_fin_union [OF pt_vrs_inst at_vrs_inst fs_vrs_inst] - finite_doms finite_vrs fresh_vrs_of T_Abs fresh_trm_dom) - apply (rule closed_in_weaken) - apply (rule closed) - done + by (simp add: T_Abs closed closed_in_weaken fresh_list_append fresh_list_cons fresh_trm_dom) then have "\ ((VarB y T\<^sub>1 # \) @ B # \) ok" by simp with _ have "(VarB y T\<^sub>1 # \) @ B # \ \ t\<^sub>2 : T\<^sub>2" by (rule T_Abs) simp @@ -1349,15 +1317,8 @@ have closed: "T\<^sub>1 closed_in (\ @ \)" by (auto dest: typing_ok) have "\ (TVarB X T\<^sub>1 # \ @ B # \) ok" - apply (rule valid_consT) - apply (rule T_TAbs) - apply (simp add: doms_append - fresh_fin_insert [OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst] - fresh_fin_union [OF pt_tyvrs_inst at_tyvrs_inst fs_tyvrs_inst] - finite_doms finite_vrs tyvrs_fresh T_TAbs fresh_dom) - apply (rule closed_in_weaken) - apply (rule closed) - done + by (simp add: T_TAbs at_tyvrs_inst closed closed_in_weaken doms_append finite_doms finite_vrs + fresh_dom fresh_fin_union fs_tyvrs_inst pt_tyvrs_inst tyvrs_fresh) then have "\ ((TVarB X T\<^sub>1 # \) @ B # \) ok" by simp with _ have "(TVarB X T\<^sub>1 # \) @ B # \ \ t\<^sub>2 : T\<^sub>2" by (rule T_TAbs) simp @@ -1375,12 +1336,14 @@ lemma type_weaken': "\ \ t : T \ \ (\@\) ok \ (\@\) \ t : T" - apply (induct \) - apply simp_all - apply (erule validE) - apply (insert type_weaken [of "[]", simplified]) - apply simp_all - done +proof (induct \) + case Nil + then show ?case by auto +next + case (Cons a \) + then show ?case + by (metis append_Cons append_Nil type_weaken validE(3)) +qed text \A.6\ @@ -1458,7 +1421,7 @@ next assume "x\y" then show ?case using T_Var - by (auto simp add:binding.inject dest: valid_cons') + by (auto simp:binding.inject dest: valid_cons') qed next case (T_App t1 T1 T2 t2 x u D) @@ -1474,13 +1437,13 @@ next case (T_TAbs X T1 t2 T2 x u D) from \TVarB X T1 # D @ VarB x U # \ \ t2 : T2\ have "X \ T1" - by (auto simp add: valid_ty_dom_fresh dest: typing_ok intro!: closed_in_fresh) + by (auto simp: valid_ty_dom_fresh dest: typing_ok intro!: closed_in_fresh) with \X \ u\ and T_TAbs show ?case by fastforce next case (T_TApp X t1 T2 T11 T12 x u D) then have "(D@\) \T2<:T11" using T_TApp by (auto dest: strengthening) then show "((D @ \) \ ((t1 \\<^sub>\ T2)[x \ u]) : (T12[X \ T2]\<^sub>\))" using T_TApp - by (force simp add: fresh_prod fresh_list_append fresh_list_cons subst_trm_fresh_tyvar) + by (force simp: fresh_prod fresh_list_append fresh_list_cons subst_trm_fresh_tyvar) qed subsubsection \Type Substitution Preserves Subtyping\ @@ -1538,7 +1501,7 @@ from X_\_ok have "X \ ty_dom \" and "\ \ ok" by auto then have "X \ \" by (simp add: valid_ty_dom_fresh) with Y have "X \ S" - by (induct \) (auto simp add: fresh_list_nil fresh_list_cons) + by (induct \) (auto simp: fresh_list_nil fresh_list_cons) with ST have "(D[X \ P]\<^sub>e @ \)\S<:T[X \ P]\<^sub>\" by (simp add: type_subst_identity) moreover from Y have "TVarB Y S \ set (D[X \ P]\<^sub>e @ \)" by simp @@ -1642,7 +1605,7 @@ from T_TApp have "D'[X \ P]\<^sub>e @ G \ t1[X \\<^sub>\ P] : (\X'<:T11. T12)[X \ P]\<^sub>\" by simp with X' and T_TApp show ?case - by (auto simp add: fresh_atm type_substitution_lemma + by (auto simp: fresh_atm type_substitution_lemma fresh_list_append fresh_list_cons ctxt_subst_fresh' type_subst_fresh subst_trm_ty_fresh intro: substT_subtype) @@ -1822,7 +1785,7 @@ case (T_Sub t S) from \[] \ S <: T\<^sub>1 \ T\<^sub>2\ obtain S\<^sub>1 S\<^sub>2 where S: "S = S\<^sub>1 \ S\<^sub>2" - by cases (auto simp add: T_Sub) + by cases (auto simp: T_Sub) then show ?case using \val t\ by (rule T_Sub) qed (auto) @@ -1834,7 +1797,7 @@ case (T_Sub t S) from \[] \ S <: (\X<:T\<^sub>1. T\<^sub>2)\ obtain X S\<^sub>1 S\<^sub>2 where S: "S = (\X<:S\<^sub>1. S\<^sub>2)" - by cases (auto simp add: T_Sub) + by cases (auto simp: T_Sub) then show ?case using T_Sub by auto qed (auto) diff -r 5e64a54f6790 -r d9b8831a6a99 src/HOL/Nominal/Examples/Lam_Funs.thy --- a/src/HOL/Nominal/Examples/Lam_Funs.thy Thu May 02 15:40:05 2024 +0200 +++ b/src/HOL/Nominal/Examples/Lam_Funs.thy Fri May 03 00:07:51 2024 +0200 @@ -22,11 +22,7 @@ "depth (Var x) = 1" | "depth (App t1 t2) = (max (depth t1) (depth t2)) + 1" | "depth (Lam [a].t) = (depth t) + 1" - apply(finite_guess)+ - apply(rule TrueI)+ - apply(simp add: fresh_nat) - apply(fresh_guess)+ - done + by(finite_guess | simp | fresh_guess)+ text \ The free variables of a lambda-term. A complication in this @@ -41,12 +37,8 @@ "frees (Var a) = {a}" | "frees (App t1 t2) = (frees t1) \ (frees t2)" | "frees (Lam [a].t) = (frees t) - {a}" -apply(finite_guess)+ -apply(simp)+ -apply(simp add: fresh_def) -apply(simp add: supp_of_fin_sets[OF pt_name_inst, OF at_name_inst, OF fs_at_inst[OF at_name_inst]]) -apply(simp add: supp_atm) -apply(blast) +apply(finite_guess | simp add: fresh_def | fresh_guess)+ + apply (simp add: at_fin_set_supp at_name_inst) apply(fresh_guess)+ done @@ -81,11 +73,7 @@ "\<(Var x)> = (lookup \ x)" | "\<(App e\<^sub>1 e\<^sub>2)> = App (\1>) (\2>)" | "x\\ \ \<(Lam [x].e)> = Lam [x].(\)" -apply(finite_guess)+ -apply(rule TrueI)+ -apply(simp add: abs_fresh)+ -apply(fresh_guess)+ -done + by (finite_guess | simp add: abs_fresh | fresh_guess)+ lemma psubst_eqvt[eqvt]: fixes pi::"name prm" @@ -107,10 +95,19 @@ lemma subst_supp: shows "supp(t1[a::=t2]) \ (((supp(t1)-{a})\supp(t2))::name set)" -apply(nominal_induct t1 avoiding: a t2 rule: lam.strong_induct) -apply(auto simp add: lam.supp supp_atm fresh_prod abs_supp) -apply(blast)+ -done +proof (nominal_induct t1 avoiding: a t2 rule: lam.strong_induct) + case (Var name) + then show ?case + by (simp add: lam.supp(1) supp_atm) +next + case (App lam1 lam2) + then show ?case + using lam.supp(2) by fastforce +next + case (Lam name lam) + then show ?case + using frees.simps(3) frees_equals_support by auto +qed text \ Contexts - lambda-terms with a single hole. diff -r 5e64a54f6790 -r d9b8831a6a99 src/HOL/Nominal/Examples/Pattern.thy --- a/src/HOL/Nominal/Examples/Pattern.thy Thu May 02 15:40:05 2024 +0200 +++ b/src/HOL/Nominal/Examples/Pattern.thy Fri May 03 00:07:51 2024 +0200 @@ -62,21 +62,23 @@ by (simp add: supp_def Collect_disj_eq del: disj_not1) instance pat :: pt_name -proof (standard, goal_cases) - case (1 x) - show ?case by (induct x) simp_all -next - case (2 _ _ x) - show ?case by (induct x) (simp_all add: pt_name2) -next - case (3 _ _ x) - then show ?case by (induct x) (simp_all add: pt_name3) +proof + fix x :: pat + show "([]::(name \ _) list) \ x = x" + by (induct x) simp_all + fix pi1 pi2 :: "(name \ name) list" + show "(pi1 @ pi2) \ x = pi1 \ pi2 \ x" + by (induct x) (simp_all add: pt_name2) + assume "pi1 \ pi2" + then show "pi1 \ x = pi2 \ x" + by (induct x) (simp_all add: pt_name3) qed instance pat :: fs_name -proof (standard, goal_cases) - case (1 x) - show ?case by (induct x) (simp_all add: fin_supp) +proof + fix x :: pat + show "finite (supp x::name set)" + by (induct x) (simp_all add: fin_supp) qed (* the following function cannot be defined using nominal_primrec, *) @@ -255,21 +257,15 @@ lemma weakening: assumes "\\<^sub>1 \ t : T" and "valid \\<^sub>2" and "\\<^sub>1 \ \\<^sub>2" shows "\\<^sub>2 \ t : T" using assms - apply (nominal_induct \\<^sub>1 t T avoiding: \\<^sub>2 rule: typing.strong_induct) - apply auto - apply (drule_tac x="(x, T) # \\<^sub>2" in meta_spec) - apply (auto intro: valid_typing) - apply (drule_tac x="\\<^sub>2" in meta_spec) - apply (drule_tac x="\ @ \\<^sub>2" in meta_spec) - apply (auto intro: valid_typing) - apply (rule typing.Let) - apply assumption+ - apply (drule meta_mp) - apply (rule valid_app_mono) - apply (rule valid_typing) - apply assumption - apply (auto simp add: pat_freshs) - done +proof (nominal_induct \\<^sub>1 t T avoiding: \\<^sub>2 rule: typing.strong_induct) + case (Abs x T \ t U) + then show ?case + by (simp add: typing.Abs valid.Cons) +next + case (Let p t \ T \ u U) + then show ?case + by (smt (verit, ccfv_threshold) Un_iff pat_freshs set_append typing.simps valid_app_mono valid_typing) +qed auto inductive match :: "pat \ trm \ (name \ trm) list \ bool" ("\ _ \ _ \ _" [50, 50, 50] 50) @@ -301,27 +297,20 @@ | "x \ \ \ \\\x:T. t\ = (\x:T. \\t\)" | "\\Base t\\<^sub>b = Base (\\t\)" | "x \ \ \ \\Bind T x t\\<^sub>b = Bind T x (\\t\\<^sub>b)" - apply finite_guess+ - apply (simp add: abs_fresh | fresh_guess)+ - done + by (finite_guess | simp add: abs_fresh | fresh_guess)+ lemma lookup_fresh: "x = y \ x \ set (map fst \) \ \(y, t)\set \. x \ t \ x \ lookup \ y" - apply (induct \) - apply (simp_all add: split_paired_all fresh_atm) - apply (case_tac "x = y") - apply (auto simp add: fresh_atm) - done + by (induct \) (use fresh_atm in force)+ lemma psubst_fresh: assumes "x \ set (map fst \)" and "\(y, t)\set \. x \ t" shows "x \ \\t\" and "x \ \\t'\\<^sub>b" using assms - apply (nominal_induct t and t' avoiding: \ rule: trm_btrm.strong_inducts) - apply simp - apply (rule lookup_fresh) - apply (rule impI) - apply (simp_all add: abs_fresh) - done +proof (nominal_induct t and t' avoiding: \ rule: trm_btrm.strong_inducts) + case (Var name) + then show ?case + by (metis lookup_fresh simps(1)) +qed (auto simp: abs_fresh) lemma psubst_eqvt[eqvt]: fixes pi :: "name prm" @@ -350,25 +339,23 @@ lemma psubst_forget: "(supp (map fst \)::name set) \* t \ \\t\ = t" "(supp (map fst \)::name set) \* t' \ \\t'\\<^sub>b = t'" - apply (nominal_induct t and t' avoiding: \ rule: trm_btrm.strong_inducts) - apply (auto simp add: fresh_star_def lookup_forget abs_fresh) - apply (drule_tac x=\ in meta_spec) - apply (drule meta_mp) - apply (rule ballI) - apply (drule_tac x=x in bspec) - apply assumption - apply (drule supp_fst) - apply (auto simp add: fresh_def) - apply (drule_tac x=\ in meta_spec) - apply (drule meta_mp) - apply (rule ballI) - apply (drule_tac x=x in bspec) - apply assumption - apply (drule supp_fst) - apply (auto simp add: fresh_def) - done +proof (nominal_induct t and t' avoiding: \ rule: trm_btrm.strong_inducts) + case (Var name) + then show ?case + by (simp add: fresh_star_set lookup_forget) +next + case (Abs ty name trm) + then show ?case + apply (simp add: fresh_def) + by (metis abs_fresh(1) fresh_star_set supp_fst trm.fresh(3)) +next + case (Bind ty name btrm) + then show ?case + apply (simp add: fresh_def) + by (metis abs_fresh(1) btrm.fresh(2) fresh_star_set supp_fst) +qed (auto simp: fresh_star_set) -lemma psubst_nil: "[]\t\ = t" "[]\t'\\<^sub>b = t'" +lemma psubst_nil[simp]: "[]\t\ = t" "[]\t'\\<^sub>b = t'" by (induct t and t' rule: trm_btrm.inducts) (simp_all add: fresh_list_nil) lemma psubst_cons: @@ -380,9 +367,20 @@ lemma psubst_append: "(supp (map fst (\\<^sub>1 @ \\<^sub>2))::name set) \* map snd (\\<^sub>1 @ \\<^sub>2) \ (\\<^sub>1 @ \\<^sub>2)\t\ = \\<^sub>2\\\<^sub>1\t\\" - by (induct \\<^sub>1 arbitrary: t) - (simp_all add: psubst_nil split_paired_all supp_list_cons psubst_cons fresh_star_def - fresh_list_cons fresh_list_append supp_list_append) +proof (induct \\<^sub>1 arbitrary: t) + case Nil + then show ?case + by (auto simp: psubst_nil) +next + case (Cons a \\<^sub>1) + then show ?case + proof (cases a) + case (Pair a b) + with Cons show ?thesis + apply (simp add: supp_list_cons fresh_star_set fresh_list_cons) + by (metis Un_iff fresh_star_set map_append psubst_cons(1) supp_list_append) + qed +qed lemma abs_pat_psubst [simp]: "(supp p::name set) \* \ \ \\\[p]. t\\<^sub>b = (\[p]. \\t\\<^sub>b)" @@ -557,17 +555,15 @@ Abs: "{x}" | Beta: "{x}" | Let: "(supp p)::name set" - apply (simp_all add: fresh_star_def abs_fresh fin_supp) - apply (rule psubst_fresh) - apply simp - apply simp - apply (rule ballI) - apply (rule psubst_fresh) - apply (rule match_vars) - apply assumption+ - apply (rule match_fresh_mono) - apply auto - done +proof (simp_all add: fresh_star_def abs_fresh fin_supp) + show "x \ t[x\u]" if "x \ u" for x t u + by (simp add: \x \ u\ psubst_fresh(1)) +next + show "\x\supp p. (x::name) \ \\u\" + if "\x\supp p. (x::name) \ t" and "\ p \ t \ \" + for p t \ u + by (meson that match_fresh_mono match_vars psubst_fresh(1)) +qed lemma typing_case_Abs: assumes ty: "\ \ (\x:T. t) : S" @@ -636,13 +632,16 @@ proof fix x assume "x \ pi \ B" then show "x \ A \ B" using pi - apply (induct pi arbitrary: x B rule: rev_induct) - apply simp - apply (simp add: split_paired_all supp_eqvt) - apply (drule perm_mem_left) - apply (simp add: calc_atm split: if_split_asm) - apply (auto dest: perm_mem_right) - done + proof (induct pi arbitrary: x B rule: rev_induct) + case Nil + then show ?case + by simp + next + case (snoc y xs) + then show ?case + apply simp + by (metis SigmaE perm_mem_left perm_pi_simp(2) pt_name2 swap_simps(3)) + qed qed lemma abs_pat_alpha': diff -r 5e64a54f6790 -r d9b8831a6a99 src/HOL/Nominal/Examples/SN.thy --- a/src/HOL/Nominal/Examples/SN.thy Thu May 02 15:40:05 2024 +0200 +++ b/src/HOL/Nominal/Examples/SN.thy Fri May 03 00:07:51 2024 +0200 @@ -88,9 +88,8 @@ assumes a: "t\\<^sub>\ s" shows "a\t \ a\s" using a -apply(nominal_induct t s avoiding: a rule: Beta.strong_induct) -apply(auto simp add: abs_fresh fresh_fact fresh_atm) -done + by (nominal_induct t s avoiding: a rule: Beta.strong_induct) + (auto simp add: abs_fresh fresh_fact fresh_atm) lemma beta_abs: assumes a: "Lam [a].t\\<^sub>\ t'" @@ -200,13 +199,8 @@ thus ?case proof (induct b rule: SN.SN.induct) case (SN_intro y) - show ?case - apply (rule hyp) - apply (erule SNI') - apply (erule SNI') - apply (rule SN.SN_intro) - apply (erule SN_intro)+ - done + with SNI' show ?case + by (metis SN.simps hyp) qed qed from b show ?thesis by (rule r) @@ -218,10 +212,7 @@ and c: "\x z. \\y. x \\<^sub>\ y \ P y z; \u. z \\<^sub>\ u \ P x u\ \ P x z" shows "P a b" using a b c -apply(rule_tac double_SN_aux) -apply(assumption)+ -apply(blast) -done + by (smt (verit, best) double_SN_aux) section \Candidates\ @@ -249,11 +240,7 @@ "fst_app_aux (Var a) = None" | "fst_app_aux (App t1 t2) = Some t1" | "fst_app_aux (Lam [x].t) = None" -apply(finite_guess)+ -apply(rule TrueI)+ -apply(simp add: fresh_none) -apply(fresh_guess)+ -done +by (finite_guess | simp add: fresh_none | fresh_guess)+ definition fst_app_def[simp]: "fst_app t = the (fst_app_aux t)" @@ -425,66 +412,40 @@ fix t u assume ih1: "\t'. \t \\<^sub>\ t'; u\RED \; \s\RED \. t'[x::=s]\RED \\ \ App (Lam [x].t') u \ RED \" assume ih2: "\u'. \u \\<^sub>\ u'; u'\RED \; \s\RED \. t[x::=s]\RED \\ \ App (Lam [x].t) u' \ RED \" - assume as1: "u \ RED \" - assume as2: "\s\RED \. t[x::=s]\RED \" + assume u: "u \ RED \" + assume t: "\s\RED \. t[x::=s]\RED \" have "CR3_RED (App (Lam [x].t) u) \" unfolding CR3_RED_def proof(intro strip) fix r assume red: "App (Lam [x].t) u \\<^sub>\ r" moreover { assume "\t'. t \\<^sub>\ t' \ r = App (Lam [x].t') u" - then obtain t' where a1: "t \\<^sub>\ t'" and a2: "r = App (Lam [x].t') u" by blast - have "App (Lam [x].t') u\RED \" using ih1 a1 as1 as2 - apply(auto) - apply(drule_tac x="t'" in meta_spec) - apply(simp) - apply(drule meta_mp) - prefer 2 - apply(auto)[1] - apply(rule ballI) - apply(drule_tac x="s" in bspec) - apply(simp) - apply(subgoal_tac "CR2 \")(*A*) - apply(unfold CR2_def)[1] - apply(drule_tac x="t[x::=s]" in spec) - apply(drule_tac x="t'[x::=s]" in spec) - apply(simp add: beta_subst) - (*A*) - apply(simp add: RED_props) - done - then have "r\RED \" using a2 by simp + then obtain t' where "t \\<^sub>\ t'" and t': "r = App (Lam [x].t') u" + by blast + then have "\s\RED \. t'[x::=s] \ RED \" + using CR2_def RED_props(2) t beta_subst by blast + then have "App (Lam [x].t') u\RED \" + using \t \\<^sub>\ t'\ u ih1 by blast + then have "r\RED \" using t' by simp } moreover { assume "\u'. u \\<^sub>\ u' \ r = App (Lam [x].t) u'" - then obtain u' where b1: "u \\<^sub>\ u'" and b2: "r = App (Lam [x].t) u'" by blast - have "App (Lam [x].t) u'\RED \" using ih2 b1 as1 as2 - apply(auto) - apply(drule_tac x="u'" in meta_spec) - apply(simp) - apply(drule meta_mp) - apply(subgoal_tac "CR2 \") - apply(unfold CR2_def)[1] - apply(drule_tac x="u" in spec) - apply(drule_tac x="u'" in spec) - apply(simp) - apply(simp add: RED_props) - apply(simp) - done - then have "r\RED \" using b2 by simp + then obtain u' where "u \\<^sub>\ u'" and u': "r = App (Lam [x].t) u'" by blast + have "CR2 \" + by (simp add: RED_props(2)) + then + have "App (Lam [x].t) u'\RED \" + using CR2_def ih2 \u \\<^sub>\ u'\ t u by blast + then have "r\RED \" using u' by simp } moreover { assume "r = t[x::=u]" - then have "r\RED \" using as1 as2 by auto + then have "r\RED \" using u t by auto } - ultimately show "r \ RED \" + ultimately show "r \ RED \" + by cases (auto simp: alpha subst_rename lam.inject dest: beta_abs) (* one wants to use the strong elimination principle; for this one has to know that x\u *) - apply(cases) - apply(auto simp add: lam.inject) - apply(drule beta_abs) - apply(auto)[1] - apply(auto simp add: alpha subst_rename) - done qed moreover have "NEUT (App (Lam [x].t) u)" unfolding NEUT_def by (auto) @@ -545,18 +506,9 @@ lemma id_closes: shows "(id \) closes \" -apply(auto) -apply(simp add: id_maps) -apply(subgoal_tac "CR3 T") \ \A\ -apply(drule CR3_implies_CR4) -apply(simp add: CR4_def) -apply(drule_tac x="Var x" in spec) -apply(force simp add: NEUT_def NORMAL_Var) -\ \A\ -apply(rule RED_props) -done + by (metis CR3_implies_CR4 CR4_def NEUT_def NORMAL_Var RED_props(3) id_maps) -lemma typing_implies_RED: +lemma typing_implies_RED: assumes a: "\ \ t : \" shows "t \ RED \" proof - diff -r 5e64a54f6790 -r d9b8831a6a99 src/HOL/Nominal/Examples/SOS.thy --- a/src/HOL/Nominal/Examples/SOS.thy Thu May 02 15:40:05 2024 +0200 +++ b/src/HOL/Nominal/Examples/SOS.thy Fri May 03 00:07:51 2024 +0200 @@ -67,11 +67,7 @@ "\<(Var x)> = (lookup \ x)" | "\<(App e\<^sub>1 e\<^sub>2)> = App (\1>) (\2>)" | "x\\ \ \<(Lam [x].e)> = Lam [x].(\)" -apply(finite_guess)+ -apply(rule TrueI)+ -apply(simp add: abs_fresh)+ -apply(fresh_guess)+ -done + by (finite_guess | simp add: abs_fresh | fresh_guess)+ lemma psubst_eqvt[eqvt]: fixes pi::"name prm" @@ -357,24 +353,30 @@ lemma V_eqvt: fixes pi::"name prm" - assumes a: "x\V T" - shows "(pi\x)\V T" -using a -apply(nominal_induct T arbitrary: pi x rule: ty.strong_induct) -apply(auto simp add: trm.inject) -apply(simp add: eqvts) -apply(rule_tac x="pi\xa" in exI) -apply(rule_tac x="pi\e" in exI) -apply(simp) -apply(auto) -apply(drule_tac x="(rev pi)\v" in bspec) -apply(force) -apply(auto) -apply(rule_tac x="pi\v'" in exI) -apply(auto) -apply(drule_tac pi="pi" in big.eqvt) -apply(perm_simp add: eqvts) -done + assumes "x \ V T" + shows "(pi\x) \ V T" +using assms +proof (nominal_induct T arbitrary: pi x rule: ty.strong_induct) + case (TVar nat) + then show ?case + by (simp add: val.eqvt) +next + case (Arrow T\<^sub>1 T\<^sub>2 pi x) + obtain a e where ae: "x = Lam [a]. e" "\v\V T\<^sub>1. \v'. e[a::=v] \ v' \ v' \ V T\<^sub>2" + using Arrow.prems by auto + have "\v'. (pi \ e)[(pi \ a)::=v] \ v' \ v' \ V T\<^sub>2" if v: "v \ V T\<^sub>1" for v + proof - + have "rev pi \ v \ V T\<^sub>1" + by (simp add: Arrow.hyps(1) v) + then obtain v' where "e[a::=(rev pi \ v)] \ v'" "v' \ V T\<^sub>2" + using ae(2) by blast + then have "(pi \ e)[(pi \ a)::=v] \ pi \ v'" + by (metis (no_types, lifting) big.eqvt cons_eqvt nil_eqvt perm_pi_simp(1) perm_prod.simps psubst_eqvt) + then show ?thesis + using Arrow.hyps \v' \ V T\<^sub>2\ by blast + qed + with ae show ?case by force +qed lemma V_arrow_elim_weak: assumes h:"u \ V (T\<^sub>1 \ T\<^sub>2)" @@ -385,33 +387,28 @@ fixes c::"'a::fs_name" assumes h: "u \ V (T\<^sub>1 \ T\<^sub>2)" obtains a t where "a\c" "u = Lam [a].t" "\v \ (V T\<^sub>1). \ v'. t[a::=v] \ v' \ v' \ V T\<^sub>2" -using h -apply - -apply(erule V_arrow_elim_weak) -apply(subgoal_tac "\a'::name. a'\(a,t,c)") (*A*) -apply(erule exE) -apply(drule_tac x="a'" in meta_spec) -apply(drule_tac x="[(a,a')]\t" in meta_spec) -apply(drule meta_mp) -apply(simp) -apply(drule meta_mp) -apply(simp add: trm.inject alpha fresh_left fresh_prod calc_atm fresh_atm) -apply(perm_simp) -apply(force) -apply(drule meta_mp) -apply(rule ballI) -apply(drule_tac x="[(a,a')]\v" in bspec) -apply(simp add: V_eqvt) -apply(auto) -apply(rule_tac x="[(a,a')]\v'" in exI) -apply(auto) -apply(drule_tac pi="[(a,a')]" in big.eqvt) -apply(perm_simp add: eqvts calc_atm) -apply(simp add: V_eqvt) -(*A*) -apply(rule exists_fresh') -apply(simp add: fin_supp) -done +proof - + obtain a t where "u = Lam [a].t" + and at: "\v. v \ (V T\<^sub>1) \ \ v'. t[a::=v] \ v' \ v' \ V T\<^sub>2" + using V_arrow_elim_weak [OF assms] by metis + obtain a'::name where a': "a'\(a,t,c)" + by (meson exists_fresh fs_name_class.axioms) + then have "u = Lam [a'].([(a, a')] \ t)" + unfolding \u = Lam [a].t\ + by (smt (verit) alpha fresh_atm fresh_prod perm_swap trm.inject(2)) + moreover + have "\ v'. ([(a, a')] \ t)[a'::=v] \ v' \ v' \ V T\<^sub>2" if "v \ (V T\<^sub>1)" for v + proof - + obtain v' where v': "t[a::=([(a, a')] \ v)] \ v' \ v' \ V T\<^sub>2" + using V_eqvt \v \ V T\<^sub>1\ at by blast + then have "([(a, a')] \ t[a::=([(a, a')] \ v)]) \ [(a, a')] \ v'" + by (simp add: big.eqvt) + then show ?thesis + by (smt (verit) V_eqvt cons_eqvt nil_eqvt perm_prod.simps perm_swap(1) psubst_eqvt swap_simps(1) v') + qed + ultimately show thesis + by (metis fresh_prod that a') +qed lemma Vs_are_values: assumes a: "e \ V T" diff -r 5e64a54f6790 -r d9b8831a6a99 src/HOL/Nominal/Examples/Standardization.thy --- a/src/HOL/Nominal/Examples/Standardization.thy Thu May 02 15:40:05 2024 +0200 +++ b/src/HOL/Nominal/Examples/Standardization.thy Fri May 03 00:07:51 2024 +0200 @@ -31,11 +31,7 @@ "size (Var n) = 0" | "size (t \ u) = size t + size u + 1" | "size (Lam [x].t) = size t + 1" - apply finite_guess+ - apply (rule TrueI)+ - apply (simp add: fresh_nat) - apply fresh_guess+ - done + by (finite_guess | simp add: fresh_nat | fresh_guess)+ instance .. @@ -47,40 +43,36 @@ subst_Var: "(Var x)[y::=s] = (if x=y then s else (Var x))" | subst_App: "(t\<^sub>1 \ t\<^sub>2)[y::=s] = t\<^sub>1[y::=s] \ t\<^sub>2[y::=s]" | subst_Lam: "x \ (y, s) \ (Lam [x].t)[y::=s] = (Lam [x].(t[y::=s]))" - apply(finite_guess)+ - apply(rule TrueI)+ - apply(simp add: abs_fresh) - apply(fresh_guess)+ - done + by (finite_guess | simp add: abs_fresh | fresh_guess)+ lemma subst_eqvt [eqvt]: "(pi::name prm) \ (t[x::=u]) = (pi \ t)[(pi \ x)::=(pi \ u)]" by (nominal_induct t avoiding: x u rule: lam.strong_induct) - (perm_simp add: fresh_bij)+ + (perm_simp add: fresh_bij)+ lemma subst_rename: "y \ t \ ([(y, x)] \ t)[y::=u] = t[x::=u]" by (nominal_induct t avoiding: x y u rule: lam.strong_induct) - (simp_all add: fresh_atm calc_atm abs_fresh) + (simp_all add: fresh_atm calc_atm abs_fresh) lemma fresh_subst: "(x::name) \ t \ x \ u \ x \ t[y::=u]" by (nominal_induct t avoiding: x y u rule: lam.strong_induct) - (auto simp add: abs_fresh fresh_atm) + (auto simp add: abs_fresh fresh_atm) lemma fresh_subst': "(x::name) \ u \ x \ t[x::=u]" by (nominal_induct t avoiding: x u rule: lam.strong_induct) - (auto simp add: abs_fresh fresh_atm) + (auto simp add: abs_fresh fresh_atm) lemma subst_forget: "(x::name) \ t \ t[x::=u] = t" by (nominal_induct t avoiding: x u rule: lam.strong_induct) - (auto simp add: abs_fresh fresh_atm) + (auto simp add: abs_fresh fresh_atm) lemma subst_subst: "x \ y \ x \ v \ t[y::=v][x::=u[y::=v]] = t[x::=u][y::=v]" by (nominal_induct t avoiding: x y u v rule: lam.strong_induct) - (auto simp add: fresh_subst subst_forget) + (auto simp add: fresh_subst subst_forget) declare subst_Var [simp del] @@ -132,20 +124,18 @@ lemma Var_apps_eq_Var_apps_conv [iff]: "(Var m \\ rs = Var n \\ ss) = (m = n \ rs = ss)" - apply (induct rs arbitrary: ss rule: rev_induct) - apply (simp add: lam.inject) - apply blast - apply (induct_tac ss rule: rev_induct) - apply (auto simp add: lam.inject) - done +proof (induct rs arbitrary: ss rule: rev_induct) + case Nil then show ?case by (auto simp add: lam.inject) +next + case (snoc x xs) then show ?case + by (induct ss rule: rev_induct) (auto simp add: lam.inject) +qed lemma App_eq_foldl_conv: "(r \ s = t \\ ts) = (if ts = [] then r \ s = t else (\ss. ts = ss @ [s] \ r = t \\ ss))" - apply (rule_tac xs = ts in rev_exhaust) - apply (auto simp add: lam.inject) - done + by (rule rev_exhaust [of ts]) (auto simp: lam.inject) lemma Abs_eq_apps_conv [iff]: "((Lam [x].r) = s \\ ss) = ((Lam [x].r) = s \ ss = [])" @@ -160,19 +150,20 @@ lemma Var_apps_neq_Abs_apps [iff]: "Var n \\ ts \ (Lam [x].r) \\ ss" - apply (induct ss arbitrary: ts rule: rev_induct) - apply simp - apply (induct_tac ts rule: rev_induct) - apply (auto simp add: lam.inject) - done +proof (induct ss arbitrary: ts rule: rev_induct) + case Nil then show ?case by simp +next + case (snoc x xs) then show ?case + by (induct ts rule: rev_induct) (auto simp add: lam.inject) +qed lemma ex_head_tail: "\ts h. t = h \\ ts \ ((\n. h = Var n) \ (\x u. h = (Lam [x].u)))" - apply (induct t rule: lam.induct) - apply (metis foldl_Nil) - apply (metis foldl_Cons foldl_Nil foldl_append) - apply (metis foldl_Nil) - done +proof (induct t rule: lam.induct) + case (App lam1 lam2) + then show ?case + by (metis foldl_Cons foldl_Nil foldl_append) +qed auto lemma size_apps [simp]: "size (r \\ rs) = size r + foldl (+) 0 (map size rs) + length rs" @@ -198,56 +189,67 @@ text \A customized induction schema for \\\\.\ -lemma lem: +lemma Apps_lam_induct_aux: assumes "\n ts (z::'a::fs_name). (\z. \t \ set ts. P z t) \ P z (Var n \\ ts)" and "\x u ts z. x \ z \ (\z. P z u) \ (\z. \t \ set ts. P z t) \ P z ((Lam [x].u) \\ ts)" shows "size t = n \ P z t" - apply (induct n arbitrary: t z rule: nat_less_induct) - apply (cut_tac t = t in ex_head_tail) - apply clarify - apply (erule disjE) - apply clarify - apply (rule assms) - apply clarify - apply (erule allE, erule impE) - prefer 2 - apply (erule allE, erule impE, rule refl, erule spec) - apply simp - apply (simp only: foldl_conv_fold add.commute fold_plus_sum_list_rev) - apply (fastforce simp add: sum_list_map_remove1) - apply clarify - apply (subgoal_tac "\y::name. y \ (x, u, z)") - prefer 2 - apply (blast intro: exists_fresh' fin_supp) - apply (erule exE) - apply (subgoal_tac "(Lam [x].u) = (Lam [y].([(y, x)] \ u))") - prefer 2 - apply (auto simp add: lam.inject alpha' fresh_prod fresh_atm)[] - apply (simp (no_asm_simp)) - apply (rule assms) - apply (simp add: fresh_prod) - apply (erule allE, erule impE) - prefer 2 - apply (erule allE, erule impE, rule refl, erule spec) - apply simp - apply clarify - apply (erule allE, erule impE) - prefer 2 - apply blast - apply simp - apply (simp only: foldl_conv_fold add.commute fold_plus_sum_list_rev) - apply (fastforce simp add: sum_list_map_remove1) - done +proof (induct n arbitrary: t z rule: less_induct) + case (less n) + obtain ts h where t: "t = h \\ ts" and D: "(\a. h = Var a) \ (\x u. h = (Lam [x].u))" + using ex_head_tail [of t] by metis + show ?case + using D + proof (elim exE disjE) + fix a :: name + assume h: "h = Var a" + have "P z t" if "t \ set ts" for z t + proof - + have "size t < length ts + fold (+) (map size ts) 0" + using that + by (fastforce simp add: sum_list_map_remove1 fold_plus_sum_list_rev) + then have "size t < size (Var a \\ ts)" + by simp (simp add: add.commute foldl_conv_fold) + then show ?thesis + using h less.hyps less.prems t by blast + qed + then show "P z t" + unfolding t h by (blast intro: assms) + next + fix x u + assume h: "h = (Lam [x].u)" + obtain y::name where y: "y \ (x, u, z)" + by (metis exists_fresh' fin_supp) + then have eq: "(Lam [x].u) = (Lam [y].([(y, x)] \ u))" + by (metis alpha' fresh_prod lam.inject(3) perm_fresh_fresh) + show "P z t" + unfolding t h eq + proof (intro assms strip) + show "y \ z" + by (simp add: y) + have "size ([(y, x)] \ u) < size ((Lam [x].u) \\ ts)" + by (simp add: eq) + then show "P z ([(y, x)] \ u)" for z + using h less.hyps less.prems t by blast + show "P z t" if "t\set ts" for z t + proof - + have 2: "size t < size ((Lam [x].u) \\ ts)" + using that + apply (simp add: eq) + apply (simp only: foldl_conv_fold add.commute fold_plus_sum_list_rev) + apply (fastforce simp add: sum_list_map_remove1) + done + then show ?thesis + using h less.hyps less.prems t by blast + qed + qed + qed +qed theorem Apps_lam_induct: assumes "\n ts (z::'a::fs_name). (\z. \t \ set ts. P z t) \ P z (Var n \\ ts)" and "\x u ts z. x \ z \ (\z. P z u) \ (\z. \t \ set ts. P z t) \ P z ((Lam [x].u) \\ ts)" shows "P z t" - apply (rule_tac t = t and z = z in lem) - prefer 3 - apply (rule refl) - using assms apply blast+ - done + using Apps_lam_induct_aux [of P] assms by blast subsection \Congruence rules\ @@ -277,66 +279,40 @@ definition step1 :: "('a \ 'a \ bool) \ 'a list \ 'a list \ bool" where - "step1 r = + "step1 r \ (\ys xs. \us z z' vs. xs = us @ z # vs \ r z' z \ ys = us @ z' # vs)" lemma not_Nil_step1 [iff]: "\ step1 r [] xs" - apply (unfold step1_def) - apply blast - done + by (simp add: step1_def) lemma not_step1_Nil [iff]: "\ step1 r xs []" - apply (unfold step1_def) - apply blast - done + by (simp add: step1_def) lemma Cons_step1_Cons [iff]: - "(step1 r (y # ys) (x # xs)) = - (r y x \ xs = ys \ x = y \ step1 r ys xs)" - apply (unfold step1_def) - apply (rule iffI) - apply (erule exE) - apply (rename_tac ts) - apply (case_tac ts) - apply fastforce - apply force - apply (erule disjE) - apply blast - apply (blast intro: Cons_eq_appendI) - done - -lemma append_step1I: - "step1 r ys xs \ vs = us \ ys = xs \ step1 r vs us - \ step1 r (ys @ vs) (xs @ us)" - apply (unfold step1_def) - apply auto - apply blast - apply (blast intro: append_eq_appendI) - done + "step1 r (y # ys) (x # xs) \ r y x \ xs = ys \ x = y \ step1 r ys xs" + apply (rule ) + apply (smt (verit, ccfv_SIG) append_eq_Cons_conv list.inject step1_def) + by (metis append_Cons append_Nil step1_def) lemma Cons_step1E [elim!]: assumes "step1 r ys (x # xs)" and "\y. ys = y # xs \ r y x \ R" and "\zs. ys = x # zs \ step1 r zs xs \ R" shows R - using assms - apply (cases ys) - apply (simp add: step1_def) - apply blast - done + by (metis Cons_step1_Cons assms list.exhaust not_Nil_step1) + +lemma append_step1I: + "step1 r ys xs \ vs = us \ ys = xs \ step1 r vs us + \ step1 r (ys @ vs) (xs @ us)" + by (smt (verit) append_Cons append_assoc step1_def) lemma Snoc_step1_SnocD: - "step1 r (ys @ [y]) (xs @ [x]) - \ (step1 r ys xs \ y = x \ ys = xs \ r y x)" - apply (unfold step1_def) - apply (clarify del: disjCI) - apply (rename_tac vs) - apply (rule_tac xs = vs in rev_exhaust) - apply force - apply simp - apply blast - done + assumes "step1 r (ys @ [y]) (xs @ [x])" + shows "(step1 r ys xs \ y = x \ ys = xs \ r y x)" + using assms + apply (clarsimp simp: step1_def) + by (metis butlast.simps(2) butlast_append butlast_snoc last.simps last_appendR list.simps(3)) subsection \Lifting beta-reduction to lists\ @@ -347,15 +323,15 @@ lemma head_Var_reduction: "Var n \\ rs \\<^sub>\ v \ \ss. rs [\\<^sub>\]\<^sub>1 ss \ v = Var n \\ ss" - apply (induct u \ "Var n \\ rs" v arbitrary: rs set: beta) - apply simp - apply (rule_tac xs = rs in rev_exhaust) - apply simp - apply (atomize, force intro: append_step1I iff: lam.inject) - apply (rule_tac xs = rs in rev_exhaust) - apply simp - apply (auto 0 3 intro: disjI2 [THEN append_step1I] simp add: lam.inject) - done +proof (induct u \ "Var n \\ rs" v arbitrary: rs set: beta) + case (appL s t u) + then show ?case + by (smt (verit, best) App_eq_foldl_conv app_last append_step1I lam.distinct(1)) +next + case (appR s t u) + then show ?case + by (smt (verit, ccfv_SIG) App_eq_foldl_conv Cons_step1_Cons app_last append_step1I lam.distinct(1)) +qed auto lemma apps_betasE [case_names appL appR beta, consumes 1]: assumes major: "r \\ rs \\<^sub>\ s" @@ -364,58 +340,48 @@ "\t u us. (x \ r \ r = (Lam [x].t) \ rs = u # us \ s = t[x::=u] \\ us) \ R" shows R proof - + note [[simproc del: defined_all]] from major have "(\r'. r \\<^sub>\ r' \ s = r' \\ rs) \ (\rs'. rs [\\<^sub>\]\<^sub>1 rs' \ s = r \\ rs') \ (\t u us. x \ r \ r = (Lam [x].t) \ rs = u # us \ s = t[x::=u] \\ us)" - supply [[simproc del: defined_all]] - apply (nominal_induct u \ "r \\ rs" s avoiding: x r rs rule: beta.strong_induct) - apply (simp add: App_eq_foldl_conv) - apply (split if_split_asm) - apply simp - apply blast - apply simp - apply (rule impI)+ - apply (rule disjI2) - apply (rule disjI2) - apply (subgoal_tac "r = [(xa, x)] \ (Lam [x].s)") - prefer 2 - apply (simp add: perm_fresh_fresh) - apply (drule conjunct1) - apply (subgoal_tac "r = (Lam [xa]. [(xa, x)] \ s)") - prefer 2 - apply (simp add: calc_atm) - apply (thin_tac "r = _") - apply simp - apply (rule exI) - apply (rule conjI) - apply (rule refl) - apply (simp add: abs_fresh fresh_atm fresh_left calc_atm subst_rename) - apply (drule App_eq_foldl_conv [THEN iffD1]) - apply (split if_split_asm) - apply simp + proof (nominal_induct u \ "r \\ rs" s avoiding: x r rs rule: beta.strong_induct) + case (beta y t s) + then show ?case + apply (simp add: App_eq_foldl_conv split: if_split_asm) apply blast - apply (force intro!: disjI1 [THEN append_step1I] simp add: fresh_list_append) - apply (drule App_eq_foldl_conv [THEN iffD1]) - apply (split if_split_asm) - apply simp + by (metis (no_types, lifting) abs_fresh(1) alpha' lam.fresh(3) lam.inject(3) subst_rename) + next + case (appL s t u) + then show ?case + apply (simp add: App_eq_foldl_conv split: if_split_asm) apply blast - apply (clarify, auto 0 3 intro!: exI intro: append_step1I) - done + by (smt (verit) append_Cons append_step1I snoc_eq_iff_butlast) + next + case (appR s t u) + then show ?case + apply (simp add: App_eq_foldl_conv step1_def split: if_split_asm) + apply force + by (metis snoc_eq_iff_butlast) + next + case (abs s t x) + then show ?case + by blast + qed with cases show ?thesis by blast qed lemma apps_preserves_betas [simp]: "rs [\\<^sub>\]\<^sub>1 ss \ r \\ rs \\<^sub>\ r \\ ss" - apply (induct rs arbitrary: ss rule: rev_induct) - apply simp - apply simp - apply (rule_tac xs = ss in rev_exhaust) - apply simp - apply simp - apply (drule Snoc_step1_SnocD) - apply blast - done +proof (induct rs arbitrary: ss rule: rev_induct) + case Nil + then show ?case by simp +next + case (snoc x ts) + then show ?case + apply (simp add: step1_def) + by (smt (verit) appR app_last apps_preserves_beta foldl_Cons foldl_append) +qed subsection \Standard reduction relation\ @@ -653,14 +619,7 @@ lemma listsp_eqvt [eqvt]: assumes xs: "listsp p (xs::'a::pt_name list)" shows "listsp ((pi::name prm) \ p) (pi \ xs)" using xs - apply induct - apply simp - apply simp - apply (rule listsp.intros) - apply (drule_tac pi=pi in perm_boolI) - apply perm_simp - apply assumption - done +by induction (use perm_app in force)+ inductive NF :: "lam \ bool" where @@ -674,13 +633,7 @@ lemma Abs_NF: assumes NF: "NF ((Lam [x].t) \\ ts)" shows "ts = []" using NF -proof cases - case (App us i) - thus ?thesis by (simp add: Var_apps_neq_Abs_apps [THEN not_sym]) -next - case (Abs u) - thus ?thesis by simp -qed + by (metis Abs_eq_apps_conv NF.cases Var_apps_neq_Abs_apps) text \ \<^term>\NF\ characterizes exactly the terms that are in normal form. diff -r 5e64a54f6790 -r d9b8831a6a99 src/HOL/Nominal/Examples/W.thy --- a/src/HOL/Nominal/Examples/W.thy Thu May 02 15:40:05 2024 +0200 +++ b/src/HOL/Nominal/Examples/W.thy Fri May 03 00:07:51 2024 +0200 @@ -13,17 +13,15 @@ lemma difference_eqvt_tvar[eqvt]: fixes pi::"tvar prm" - and Xs Ys::"tvar list" + and Xs Ys::"tvar list" shows "pi\(Xs - Ys) = (pi\Xs) - (pi\Ys)" -by (induct Xs) (simp_all add: eqvts) + by (induct Xs) (simp_all add: eqvts) -lemma difference_fresh: +lemma difference_fresh [simp]: fixes X::"tvar" - and Xs Ys::"tvar list" - assumes a: "X\set Ys" - shows "X\(Xs - Ys)" -using a -by (induct Xs) (auto simp add: fresh_list_nil fresh_list_cons fresh_atm) + and Xs Ys::"tvar list" + shows "X\(Xs - Ys) \ X\Xs \ X\set Ys" + by (induct Xs) (auto simp add: fresh_list_nil fresh_list_cons fresh_atm) lemma difference_supset: fixes xs::"'a list" @@ -114,10 +112,8 @@ where "ftv_tyS (Ty T) = ((ftv (T::ty))::tvar list)" | "ftv_tyS (\[X].S) = (ftv_tyS S) - [X]" -apply(finite_guess add: ftv_ty_eqvt fs_tvar1)+ -apply(rule TrueI)+ -apply(rule difference_fresh) -apply(simp) +apply(finite_guess add: ftv_ty_eqvt fs_tvar1 | rule TrueI)+ + apply (metis difference_fresh list.set_intros(1)) apply(fresh_guess add: ftv_ty_eqvt fs_tvar1)+ done @@ -127,18 +123,21 @@ fixes pi::"tvar prm" and S::"tyS" shows "pi\(ftv S) = ftv (pi\S)" -apply(nominal_induct S rule: tyS.strong_induct) -apply(simp add: eqvts) -apply(simp only: ftv_tyS.simps) -apply(simp only: eqvts) -apply(simp add: eqvts) -done +proof (nominal_induct S rule: tyS.strong_induct) + case (Ty ty) + then show ?case + by (simp add: ftv_ty_eqvt) +next + case (ALL tvar tyS) + then show ?case + by (metis difference_eqvt_tvar ftv_ty.simps(1) ftv_tyS.simps(2) ftv_ty_eqvt ty.perm(3) tyS.perm(4)) +qed lemma ftv_Ctxt_eqvt[eqvt]: fixes pi::"tvar prm" - and \::"Ctxt" + and \::"Ctxt" shows "pi\(ftv \) = ftv (pi\\)" -by (induct \) (auto simp add: eqvts) + by (induct \) (auto simp add: eqvts) text \Valid\ inductive @@ -157,7 +156,7 @@ lemma gen_eqvt[eqvt]: fixes pi::"tvar prm" shows "pi\(gen T Xs) = gen (pi\T) (pi\Xs)" -by (induct Xs) (simp_all add: eqvts) + by (induct Xs) (simp_all add: eqvts) @@ -169,7 +168,7 @@ lemma close_eqvt[eqvt]: fixes pi::"tvar prm" shows "pi\(close \ T) = close (pi\\) (pi\T)" -by (simp_all only: eqvts) + by (simp_all only: eqvts) text \Substitution\ @@ -183,8 +182,7 @@ where "smth[X::=T] \ ([(X,T)])" -fun - lookup :: "Subst \ tvar \ ty" +fun lookup :: "Subst \ tvar \ ty" where "lookup [] X = TVar X" | "lookup ((Y,T)#\) X = (if X=Y then T else lookup \ X)" @@ -192,15 +190,15 @@ lemma lookup_eqvt[eqvt]: fixes pi::"tvar prm" shows "pi\(lookup \ X) = lookup (pi\\) (pi\X)" -by (induct \) (auto simp add: eqvts) + by (induct \) (auto simp add: eqvts) lemma lookup_fresh: fixes X::"tvar" assumes a: "X\\" shows "lookup \ X = TVar X" -using a -by (induct \) - (auto simp add: fresh_list_cons fresh_prod fresh_atm) + using a + by (induct \) + (auto simp add: fresh_list_cons fresh_prod fresh_atm) overloading psubst_ty \ "psubst :: Subst \ ty \ ty" @@ -231,9 +229,7 @@ where "\<(Ty T)> = Ty (\)" | "X\\ \ \<(\[X].S)> = \[X].(\)" -apply(finite_guess add: psubst_ty_eqvt fs_tvar1)+ -apply(rule TrueI)+ -apply(simp add: abs_fresh) +apply(finite_guess add: psubst_ty_eqvt fs_tvar1 | simp add: abs_fresh)+ apply(fresh_guess add: psubst_ty_eqvt fs_tvar1)+ done @@ -330,26 +326,26 @@ assumes a: "X\\" shows "\ = (\)[X::=\]" using a -apply(nominal_induct T avoiding: \ X T' rule: ty.strong_induct) -apply(auto simp add: ty.inject lookup_fresh) -apply(rule sym) -apply(rule subst_forget_ty) -apply(rule fresh_lookup) -apply(simp_all add: fresh_atm) -done +proof (nominal_induct T avoiding: \ X T' rule: ty.strong_induct) + case (TVar tvar) + then show ?case + by (simp add: fresh_atm(1) fresh_lookup lookup_fresh subst_forget_ty) +qed auto lemma general_preserved: fixes \::"Subst" assumes a: "T \ S" shows "\ \ \" -using a -apply(nominal_induct T S avoiding: \ rule: inst.strong_induct) -apply(auto)[1] -apply(simp add: psubst_ty_lemma) -apply(rule_tac I_All) -apply(simp add: fresh_psubst_ty) -apply(simp) -done + using a +proof (nominal_induct T S avoiding: \ rule: inst.strong_induct) + case (I_Ty T) + then show ?case + by (simp add: inst.I_Ty) +next + case (I_All X T' T S) + then show ?case + by (simp add: fresh_psubst_ty inst.I_All psubst_ty_lemma) +qed text\typing judgements\ @@ -386,74 +382,72 @@ lemma ftv_Ctxt: fixes \::"Ctxt" shows "supp \ = set (ftv \)" -apply (induct \) -apply (simp_all add: supp_list_nil supp_list_cons) -apply (rename_tac a \') -apply (case_tac a) -apply (simp add: supp_prod supp_atm ftv_tyS) -done +proof (induct \) + case Nil + then show ?case + by (simp add: supp_list_nil) +next + case (Cons c \) + show ?case + proof (cases c) + case (Pair a b) + with Cons show ?thesis + by (simp add: ftv_tyS supp_atm(3) supp_list_cons supp_prod) + qed +qed -lemma ftv_tvars: +lemma ftv_tvars: fixes Tvs::"tvar list" shows "supp Tvs = set Tvs" -by (induct Tvs) - (simp_all add: supp_list_nil supp_list_cons supp_atm) + by (induct Tvs) (simp_all add: supp_list_nil supp_list_cons supp_atm) lemma difference_supp: fixes xs ys::"tvar list" shows "((supp (xs - ys))::tvar set) = supp xs - supp ys" -by (induct xs) - (auto simp add: supp_list_nil supp_list_cons ftv_tvars) + by (induct xs) (auto simp add: supp_list_nil supp_list_cons ftv_tvars) lemma set_supp_eq: fixes xs::"tvar list" shows "set xs = supp xs" -by (induct xs) - (simp_all add: supp_list_nil supp_list_cons supp_atm) + by (induct xs) (simp_all add: supp_list_nil supp_list_cons supp_atm) nominal_inductive2 typing avoids T_LET: "set (ftv T\<^sub>1 - ftv \)" -apply (simp add: fresh_star_def fresh_def ftv_Ctxt) -apply (simp add: fresh_star_def fresh_tvar_trm) -apply assumption -apply simp -done + apply (simp_all add: fresh_star_def fresh_def ftv_Ctxt fresh_tvar_trm) + by (meson fresh_def fresh_tvar_trm) + lemma perm_fresh_fresh_aux: "\(x,y)\set (pi::tvar prm). x \ z \ y \ z \ pi \ (z::'a::pt_tvar) = z" - apply (induct pi rule: rev_induct) - apply simp - apply (simp add: split_paired_all pt_tvar2) - apply (frule_tac x="(a, b)" in bspec) - apply simp - apply (simp add: perm_fresh_fresh) - done +proof (induct pi rule: rev_induct) + case Nil + then show ?case + by simp +next + case (snoc x xs) + then show ?case + unfolding split_paired_all pt_tvar2 + using perm_fresh_fresh(1) by fastforce +qed lemma freshs_mem: fixes S::"tvar set" assumes "x \ S" - and "S \* z" + and "S \* z" shows "x \ z" -using assms by (simp add: fresh_star_def) + using assms by (simp add: fresh_star_def) lemma fresh_gen_set: fixes X::"tvar" and Xs::"tvar list" - assumes asm: "X\set Xs" + assumes "X\set Xs" shows "X\gen T Xs" -using asm -apply(induct Xs) -apply(simp) -apply(rename_tac a Xs') -apply(case_tac "X=a") -apply(simp add: abs_fresh) -apply(simp add: abs_fresh) -done + using assms by (induct Xs) (auto simp: abs_fresh) lemma close_fresh: fixes \::"Ctxt" shows "\(X::tvar)\set ((ftv T) - (ftv \)). X\(close \ T)" -by (simp add: fresh_gen_set) + by (meson fresh_gen_set) lemma gen_supp: shows "(supp (gen T Xs)::tvar set) = supp T - supp Xs" @@ -466,9 +460,7 @@ lemma close_supp: shows "supp (close \ T) = set (ftv T) \ set (ftv \)" - apply (simp add: gen_supp difference_supp ftv_ty ftv_Ctxt) - apply (simp only: set_supp_eq minus_Int_eq) - done + using difference_supp ftv_ty gen_supp set_supp_eq by auto lemma better_T_LET: assumes x: "x\\" @@ -483,22 +475,12 @@ from pi1 have pi1': "(pi \ set (ftv T\<^sub>1 - ftv \)) \* \" by (simp add: fresh_star_prod) have Gamma_fresh: "\(x,y)\set pi. x \ \ \ y \ \" - apply (rule ballI) - apply (simp add: split_paired_all) - apply (drule subsetD [OF pi2]) - apply (erule SigmaE) - apply (drule freshs_mem [OF _ pi1']) - apply (simp add: ftv_Ctxt [symmetric] fresh_def) - done - have close_fresh': "\(x, y)\set pi. x \ close \ T\<^sub>1 \ y \ close \ T\<^sub>1" - apply (rule ballI) - apply (simp add: split_paired_all) - apply (drule subsetD [OF pi2]) - apply (erule SigmaE) - apply (drule bspec [OF close_fresh]) - apply (drule freshs_mem [OF _ pi1']) - apply (simp add: fresh_def close_supp ftv_Ctxt) - done + using freshs_mem [OF _ pi1'] subsetD [OF pi2] ftv_Ctxt fresh_def by fastforce + have "\x y. \x \ set (ftv T\<^sub>1 - ftv \); y \ pi \ set (ftv T\<^sub>1 - ftv \)\ + \ x \ close \ T\<^sub>1 \ y \ close \ T\<^sub>1" + by (metis DiffE filter_set fresh_def fresh_gen_set freshs_mem ftv_Ctxt ftv_ty gen_supp member_filter pi1') + then have close_fresh': "\(x, y)\set pi. x \ close \ T\<^sub>1 \ y \ close \ T\<^sub>1" + using pi2 by auto note x moreover from Gamma_fresh perm_boolI [OF t1, of pi] have "\ \ t\<^sub>1 : pi \ T\<^sub>1" @@ -518,96 +500,81 @@ fixes T::"ty" and \::"Subst" and X Y ::"tvar" - assumes a1: "X \ set (ftv T)" - and a2: "Y \ set (ftv (lookup \ X))" + assumes "X \ set (ftv T)" + and "Y \ set (ftv (lookup \ X))" shows "Y \ set (ftv (\))" -using a1 a2 -by (nominal_induct T rule: ty.strong_induct) (auto) + using assms + by (nominal_induct T rule: ty.strong_induct) (auto) lemma ftv_tyS_subst: fixes S::"tyS" and \::"Subst" and X Y::"tvar" - assumes a1: "X \ set (ftv S)" - and a2: "Y \ set (ftv (lookup \ X))" + assumes "X \ set (ftv S)" + and "Y \ set (ftv (lookup \ X))" shows "Y \ set (ftv (\))" -using a1 a2 + using assms by (nominal_induct S avoiding: \ Y rule: tyS.strong_induct) (auto simp add: ftv_ty_subst fresh_atm) lemma ftv_Ctxt_subst: fixes \::"Ctxt" and \::"Subst" - assumes a1: "X \ set (ftv \)" - and a2: "Y \ set (ftv (lookup \ X))" +assumes "X \ set (ftv \)" + and "Y \ set (ftv (lookup \ X))" shows "Y \ set (ftv (\<\>))" -using a1 a2 -by (induct \) - (auto simp add: ftv_tyS_subst) + using assms by (induct \) (auto simp add: ftv_tyS_subst) lemma gen_preserved1: - assumes asm: "Xs \* \" + assumes "Xs \* \" shows "\ = gen (\) Xs" -using asm -by (induct Xs) - (auto simp add: fresh_star_def) + using assms by (induct Xs) (auto simp add: fresh_star_def) lemma gen_preserved2: fixes T::"ty" and \::"Ctxt" - assumes asm: "((ftv T) - (ftv \)) \* \" + assumes "((ftv T) - (ftv \)) \* \" shows "((ftv (\)) - (ftv (\<\>))) = ((ftv T) - (ftv \))" -using asm -apply(nominal_induct T rule: ty.strong_induct) -apply(auto simp add: fresh_star_def) -apply(simp add: lookup_fresh) -apply(simp add: ftv_Ctxt[symmetric]) -apply(fold fresh_def) -apply(rule fresh_psubst_Ctxt) -apply(assumption) -apply(assumption) -apply(rule difference_supset) -apply(auto) -apply(simp add: ftv_Ctxt_subst) -done + using assms +proof(nominal_induct T rule: ty.strong_induct) + case (TVar tvar) + then show ?case + apply(auto simp add: fresh_star_def ftv_Ctxt_subst) + by (metis filter.simps fresh_def fresh_psubst_Ctxt ftv_Ctxt ftv_ty.simps(1) lookup_fresh) +next + case (Fun ty1 ty2) + then show ?case + by (simp add: fresh_star_list) +qed lemma close_preserved: fixes \::"Ctxt" - assumes asm: "((ftv T) - (ftv \)) \* \" + assumes "((ftv T) - (ftv \)) \* \" shows "\ T> = close (\<\>) (\)" -using asm -by (simp add: gen_preserved1 gen_preserved2) + using assms by (simp add: gen_preserved1 gen_preserved2) lemma var_fresh_for_ty: fixes x::"var" - and T::"ty" + and T::"ty" shows "x\T" -by (nominal_induct T rule: ty.strong_induct) - (simp_all add: fresh_atm) + by (nominal_induct T rule: ty.strong_induct) (simp_all add: fresh_atm) lemma var_fresh_for_tyS: - fixes x::"var" - and S::"tyS" + fixes x::"var" and S::"tyS" shows "x\S" -by (nominal_induct S rule: tyS.strong_induct) - (simp_all add: abs_fresh var_fresh_for_ty) + by (nominal_induct S rule: tyS.strong_induct) (simp_all add: abs_fresh var_fresh_for_ty) lemma psubst_fresh_Ctxt: - fixes x::"var" - and \::"Ctxt" - and \::"Subst" + fixes x::"var" and \::"Ctxt" and \::"Subst" shows "x\\<\> = x\\" -by (induct \) - (auto simp add: fresh_list_cons fresh_list_nil fresh_prod var_fresh_for_tyS) + by (induct \) (auto simp add: fresh_list_cons fresh_list_nil fresh_prod var_fresh_for_tyS) lemma psubst_valid: fixes \::Subst - and \::Ctxt - assumes a: "valid \" + and \::Ctxt + assumes "valid \" shows "valid (\<\>)" -using a -by (induct) - (auto simp add: psubst_fresh_Ctxt) + using assms by (induct) (auto simp add: psubst_fresh_Ctxt) lemma psubst_in: fixes \::"Ctxt" @@ -616,17 +583,14 @@ and S::"tyS" assumes a: "(x,S)\set \" shows "(x,\)\set (\<\>)" -using a -by (induct \) - (auto simp add: calc_atm) - + using a by (induct \) (auto simp add: calc_atm) lemma typing_preserved: fixes \::"Subst" - and pi::"tvar prm" - assumes a: "\ \ t : T" + and pi::"tvar prm" + assumes "\ \ t : T" shows "(\<\>) \ t : (\)" -using a + using assms proof (nominal_induct \ t T avoiding: \ rule: typing.strong_induct) case (T_VAR \ x S T) have a1: "valid \" by fact @@ -662,14 +626,7 @@ have a2: "\<\> \ t1 : \" by fact have a3: "\<((x, close \ T1)#\)> \ t2 : \" by fact from a2 a3 show "\<\> \ Let x be t1 in t2 : \" - apply - - apply(rule better_T_LET) - apply(rule a1) - apply(rule a2) - apply(simp add: close_preserved vc) - done + by (simp add: a1 better_T_LET close_preserved vc) qed - - end diff -r 5e64a54f6790 -r d9b8831a6a99 src/HOL/Nominal/Nominal.thy --- a/src/HOL/Nominal/Nominal.thy Thu May 02 15:40:05 2024 +0200 +++ b/src/HOL/Nominal/Nominal.thy Fri May 03 00:07:51 2024 +0200 @@ -1,5 +1,6 @@ theory Nominal -imports "HOL-Library.Infinite_Set" "HOL-Library.Old_Datatype" + imports "HOL-Library.Infinite_Set" "HOL-Library.Old_Datatype" + keywords "atom_decl" :: thy_decl and "nominal_datatype" :: thy_defn and @@ -138,7 +139,7 @@ by (simp add: perm_bool) (* permutation on sets *) -lemma empty_eqvt: +lemma empty_eqvt[simp]: shows "pi\{} = {}" by (simp add: perm_set_def) @@ -209,11 +210,11 @@ shows "(supp x) = {a::'x. \a\x}" by (simp add: fresh_def) -lemma supp_unit: +lemma supp_unit[simp]: shows "supp () = {}" by (simp add: supp_def) -lemma supp_set_empty: +lemma supp_set_empty[simp]: shows "supp {} = {}" by (force simp add: supp_def empty_eqvt) @@ -229,7 +230,7 @@ shows "(supp (nPair x y)) = (supp x)\(supp y)" by (force simp add: supp_def Collect_imp_eq Collect_neg_eq) -lemma supp_list_nil: +lemma supp_list_nil[simp]: shows "supp [] = {}" by (simp add: supp_def) @@ -250,47 +251,47 @@ shows "supp (rev xs) = (supp xs)" by (induct xs, auto simp add: supp_list_append supp_list_cons supp_list_nil) -lemma supp_bool: +lemma supp_bool[simp]: fixes x :: "bool" shows "supp x = {}" by (cases "x") (simp_all add: supp_def) -lemma supp_some: +lemma supp_some[simp]: fixes x :: "'a" shows "supp (Some x) = (supp x)" by (simp add: supp_def) -lemma supp_none: +lemma supp_none[simp]: fixes x :: "'a" shows "supp (None) = {}" by (simp add: supp_def) -lemma supp_int: +lemma supp_int[simp]: fixes i::"int" shows "supp (i) = {}" by (simp add: supp_def perm_int_def) -lemma supp_nat: +lemma supp_nat[simp]: fixes n::"nat" shows "(supp n) = {}" by (simp add: supp_def perm_nat_def) -lemma supp_char: +lemma supp_char[simp]: fixes c::"char" shows "(supp c) = {}" by (simp add: supp_def perm_char_def) -lemma supp_string: +lemma supp_string[simp]: fixes s::"string" shows "(supp s) = {}" by (simp add: supp_def perm_string) (* lemmas about freshness *) -lemma fresh_set_empty: +lemma fresh_set_empty[simp]: shows "a\{}" by (simp add: fresh_def supp_set_empty) -lemma fresh_unit: +lemma fresh_unit[simp]: shows "a\()" by (simp add: fresh_def supp_unit) @@ -301,7 +302,7 @@ shows "a\(x,y) = (a\x \ a\y)" by (simp add: fresh_def supp_prod) -lemma fresh_list_nil: +lemma fresh_list_nil[simp]: fixes a :: "'x" shows "a\[]" by (simp add: fresh_def supp_list_nil) @@ -320,48 +321,48 @@ shows "a\(xs@ys) = (a\xs \ a\ys)" by (simp add: fresh_def supp_list_append) -lemma fresh_list_rev: +lemma fresh_list_rev[simp]: fixes a :: "'x" and xs :: "'a list" shows "a\(rev xs) = a\xs" by (simp add: fresh_def supp_list_rev) -lemma fresh_none: +lemma fresh_none[simp]: fixes a :: "'x" shows "a\None" by (simp add: fresh_def supp_none) -lemma fresh_some: +lemma fresh_some[simp]: fixes a :: "'x" and x :: "'a" shows "a\(Some x) = a\x" by (simp add: fresh_def supp_some) -lemma fresh_int: +lemma fresh_int[simp]: fixes a :: "'x" and i :: "int" shows "a\i" by (simp add: fresh_def supp_int) -lemma fresh_nat: +lemma fresh_nat[simp]: fixes a :: "'x" and n :: "nat" shows "a\n" by (simp add: fresh_def supp_nat) -lemma fresh_char: +lemma fresh_char[simp]: fixes a :: "'x" and c :: "char" shows "a\c" by (simp add: fresh_def supp_char) -lemma fresh_string: +lemma fresh_string[simp]: fixes a :: "'x" and s :: "string" shows "a\s" by (simp add: fresh_def supp_string) -lemma fresh_bool: +lemma fresh_bool[simp]: fixes a :: "'x" and b :: "bool" shows "a\b" @@ -726,17 +727,13 @@ and b :: "'x" assumes at: "at TYPE('x)" shows "(pi1@pi2) \ ((pi1\pi2)@pi1)" -apply(induct_tac pi2) -apply(simp add: prm_eq_def) -apply(auto simp add: prm_eq_def) -apply(simp add: at2[OF at]) -apply(drule_tac x="aa" in spec) -apply(drule sym) -apply(simp) -apply(simp add: at_append[OF at]) -apply(simp add: at2[OF at]) -apply(simp add: at_ds8_aux[OF at]) -done +proof(induct_tac pi2) + show "(pi1 @ []) \ (pi1 \ [] @ pi1)" + by(simp add: prm_eq_def) + show "\a l. (pi1 @ l) \ (pi1 \ l @ pi1) \ + (pi1 @ a # l) \ (pi1 \ (a # l) @ pi1) " + by(auto simp add: prm_eq_def at at2 at_append at_ds8_aux) +qed lemma at_ds9: fixes pi1 :: "'x prm" @@ -745,32 +742,16 @@ and b :: "'x" assumes at: "at TYPE('x)" shows " ((rev pi2)@(rev pi1)) \ ((rev pi1)@(rev (pi1\pi2)))" -apply(induct_tac pi2) -apply(simp add: prm_eq_def) -apply(auto simp add: prm_eq_def) -apply(simp add: at_append[OF at]) -apply(simp add: at2[OF at] at1[OF at]) -apply(drule_tac x="swap(pi1\a,pi1\b) aa" in spec) -apply(drule sym) -apply(simp) -apply(simp add: at_ds8_aux[OF at]) -apply(simp add: at_rev_pi[OF at]) -done + using at at_ds8 at_prm_rev_eq1 rev_append by fastforce lemma at_ds10: fixes pi :: "'x prm" and a :: "'x" and b :: "'x" - assumes at: "at TYPE('x)" - and a: "b\(rev pi)" + assumes "at TYPE('x)" + and "b\(rev pi)" shows "([(pi\a,b)]@pi) \ (pi@[(a,b)])" -using a -apply - -apply(rule at_prm_eq_trans) -apply(rule at_ds2[OF at]) -apply(simp add: at_prm_fresh[OF at] at_rev_pi[OF at]) -apply(rule at_prm_eq_refl) -done + by (metis assms at_bij1 at_ds2 at_prm_fresh) \ \there always exists an atom that is not being in a finite set\ lemma ex_in_inf: @@ -778,14 +759,7 @@ assumes at: "at TYPE('x)" and fs: "finite A" obtains c::"'x" where "c\A" -proof - - from fs at4[OF at] have "infinite ((UNIV::'x set) - A)" - by (simp add: Diff_infinite_finite) - hence "((UNIV::'x set) - A) \ ({}::'x set)" by (force simp only:) - then obtain c::"'x" where "c\((UNIV::'x set) - A)" by force - then have "c\A" by simp - then show ?thesis .. -qed + using at at4 ex_new_if_finite fs by blast text \there always exists a fresh name for an object with finite support\ lemma at_exists_fresh': @@ -806,16 +780,8 @@ fixes S::"'a set" assumes a: "at TYPE('a)" and b: "finite S" - shows "\x. x \ S" - using a b - apply(drule_tac S="UNIV::'a set" in Diff_infinite_finite) - apply(simp add: at_def) - apply(subgoal_tac "UNIV - S \ {}") - apply(simp only: ex_in_conv [symmetric]) - apply(blast) - apply(rule notI) - apply(simp) - done + shows "\x. x \ S" + by (meson a b ex_in_inf) lemma at_different: assumes at: "at TYPE('x)" @@ -823,12 +789,8 @@ proof - have "infinite (UNIV::'x set)" by (rule at4[OF at]) hence inf2: "infinite (UNIV-{a})" by (rule infinite_remove) - have "(UNIV-{a}) \ ({}::'x set)" - proof (rule_tac ccontr, drule_tac notnotD) - assume "UNIV-{a} = ({}::'x set)" - with inf2 have "infinite ({}::'x set)" by simp - then show "False" by auto - qed + have "(UNIV-{a}) \ ({}::'x set)" + by (metis finite.emptyI inf2) hence "\(b::'x). b\(UNIV-{a})" by blast then obtain b::"'x" where mem2: "b\(UNIV-{a})" by blast from mem2 have "a\b" by blast @@ -839,11 +801,7 @@ lemma at_pt_inst: assumes at: "at TYPE('x)" shows "pt TYPE('x) TYPE('x)" -apply(auto simp only: pt_def) -apply(simp only: at1[OF at]) -apply(simp only: at_append[OF at]) -apply(simp only: prm_eq_def) -done + using at at_append at_def prm_eq_def pt_def by fastforce section \finite support properties\ (*===================================*) @@ -858,56 +816,36 @@ fixes a :: "'x" assumes at: "at TYPE('x)" shows "fs TYPE('x) TYPE('x)" -apply(simp add: fs_def) -apply(simp add: at_supp[OF at]) -done + by (simp add: at at_supp fs_def) lemma fs_unit_inst: shows "fs TYPE(unit) TYPE('x)" -apply(simp add: fs_def) -apply(simp add: supp_unit) -done + by(simp add: fs_def supp_unit) lemma fs_prod_inst: assumes fsa: "fs TYPE('a) TYPE('x)" and fsb: "fs TYPE('b) TYPE('x)" shows "fs TYPE('a\'b) TYPE('x)" -apply(unfold fs_def) -apply(auto simp add: supp_prod) -apply(rule fs1[OF fsa]) -apply(rule fs1[OF fsb]) -done + by (simp add: assms fs1 supp_prod fs_def) + lemma fs_nprod_inst: assumes fsa: "fs TYPE('a) TYPE('x)" and fsb: "fs TYPE('b) TYPE('x)" shows "fs TYPE(('a,'b) nprod) TYPE('x)" -apply(unfold fs_def, rule allI) -apply(case_tac x) -apply(auto simp add: supp_nprod) -apply(rule fs1[OF fsa]) -apply(rule fs1[OF fsb]) -done + unfolding fs_def by (metis assms finite_Un fs_def nprod.exhaust supp_nprod) lemma fs_list_inst: - assumes fs: "fs TYPE('a) TYPE('x)" + assumes "fs TYPE('a) TYPE('x)" shows "fs TYPE('a list) TYPE('x)" -apply(simp add: fs_def, rule allI) -apply(induct_tac x) -apply(simp add: supp_list_nil) -apply(simp add: supp_list_cons) -apply(rule fs1[OF fs]) -done + unfolding fs_def + by (clarify, induct_tac x) (auto simp: fs1 assms supp_list_nil supp_list_cons) lemma fs_option_inst: assumes fs: "fs TYPE('a) TYPE('x)" shows "fs TYPE('a option) TYPE('x)" -apply(simp add: fs_def, rule allI) -apply(case_tac x) -apply(simp add: supp_none) -apply(simp add: supp_some) -apply(rule fs1[OF fs]) -done + unfolding fs_def + by (metis assms finite.emptyI fs1 option.exhaust supp_none supp_some) section \Lemmas about the permutation properties\ (*=================================================*) @@ -954,13 +892,10 @@ using cp by (simp add: cp_def) lemma cp_pt_inst: - assumes pt: "pt TYPE('a) TYPE('x)" - and at: "at TYPE('x)" + assumes "pt TYPE('a) TYPE('x)" + and "at TYPE('x)" shows "cp TYPE('a) TYPE('x) TYPE('x)" -apply(auto simp add: cp_def pt2[OF pt,symmetric]) -apply(rule pt3[OF pt]) -apply(rule at_ds8[OF at]) -done + using assms at_ds8 cp_def pt2 pt3 by fastforce section \disjointness properties\ (*=================================*) @@ -998,8 +933,7 @@ fixes a::"'x" assumes dj: "disjoint TYPE('x) TYPE('y)" shows "(supp a) = ({}::'y set)" -apply(simp add: supp_def dj_perm_forget[OF dj]) -done + by (simp add: supp_def dj_perm_forget[OF dj]) lemma at_fresh_ineq: fixes a :: "'x" @@ -1016,53 +950,35 @@ and ptb: "pt TYPE('b) TYPE('x)" and at: "at TYPE('x)" shows "pt TYPE('a\'b) TYPE('x)" -apply(auto simp only: pt_def) -apply(simp_all add: perm_fun_def) -apply(simp add: pt1[OF pta] pt1[OF ptb]) -apply(simp add: pt2[OF pta] pt2[OF ptb]) -apply(subgoal_tac "(rev pi1) \ (rev pi2)")(*A*) -apply(simp add: pt3[OF pta] pt3[OF ptb]) -(*A*) -apply(simp add: at_prm_rev_eq[OF at]) -done - -lemma pt_bool_inst: + unfolding pt_def using assms + by (auto simp add: perm_fun_def pt1 pt2 ptb pt3 pt3_rev) + +lemma pt_bool_inst[simp]: shows "pt TYPE(bool) TYPE('x)" by (simp add: pt_def perm_bool_def) lemma pt_set_inst: assumes pt: "pt TYPE('a) TYPE('x)" shows "pt TYPE('a set) TYPE('x)" -apply(simp add: pt_def) -apply(simp_all add: perm_set_def) -apply(simp add: pt1[OF pt]) -apply(force simp add: pt2[OF pt] pt3[OF pt]) -done - -lemma pt_unit_inst: + unfolding pt_def + by(auto simp add: perm_set_def pt1[OF pt] pt2[OF pt] pt3[OF pt]) + +lemma pt_unit_inst[simp]: shows "pt TYPE(unit) TYPE('x)" by (simp add: pt_def) lemma pt_prod_inst: assumes pta: "pt TYPE('a) TYPE('x)" and ptb: "pt TYPE('b) TYPE('x)" - shows "pt TYPE('a \ 'b) TYPE('x)" - apply(auto simp add: pt_def) - apply(rule pt1[OF pta]) - apply(rule pt1[OF ptb]) - apply(rule pt2[OF pta]) - apply(rule pt2[OF ptb]) - apply(rule pt3[OF pta],assumption) - apply(rule pt3[OF ptb],assumption) - done +shows "pt TYPE('a \ 'b) TYPE('x)" + using assms pt1 pt2 pt3 + by(auto simp add: pt_def) lemma pt_list_nil: fixes xs :: "'a list" assumes pt: "pt TYPE('a) TYPE ('x)" shows "([]::'x prm)\xs = xs" -apply(induct_tac xs) -apply(simp_all add: pt1[OF pt]) -done + by (induct_tac xs) (simp_all add: pt1[OF pt]) lemma pt_list_append: fixes pi1 :: "'x prm" @@ -1070,9 +986,7 @@ and xs :: "'a list" assumes pt: "pt TYPE('a) TYPE ('x)" shows "(pi1@pi2)\xs = pi1\(pi2\xs)" -apply(induct_tac xs) -apply(simp_all add: pt2[OF pt]) -done + by (induct_tac xs) (simp_all add: pt2[OF pt]) lemma pt_list_prm_eq: fixes pi1 :: "'x prm" @@ -1080,55 +994,67 @@ and xs :: "'a list" assumes pt: "pt TYPE('a) TYPE ('x)" shows "pi1 \ pi2 \ pi1\xs = pi2\xs" -apply(induct_tac xs) -apply(simp_all add: prm_eq_def pt3[OF pt]) -done + by (induct_tac xs) (simp_all add: pt3[OF pt]) lemma pt_list_inst: assumes pt: "pt TYPE('a) TYPE('x)" shows "pt TYPE('a list) TYPE('x)" -apply(auto simp only: pt_def) -apply(rule pt_list_nil[OF pt]) -apply(rule pt_list_append[OF pt]) -apply(rule pt_list_prm_eq[OF pt],assumption) -done + by (simp add: pt pt_def pt_list_append pt_list_nil pt_list_prm_eq) lemma pt_option_inst: assumes pta: "pt TYPE('a) TYPE('x)" shows "pt TYPE('a option) TYPE('x)" -apply(auto simp only: pt_def) -apply(case_tac "x") -apply(simp_all add: pt1[OF pta]) -apply(case_tac "x") -apply(simp_all add: pt2[OF pta]) -apply(case_tac "x") -apply(simp_all add: pt3[OF pta]) -done +proof - + have "([]::('x \ _) list) \ x = x" for x :: "'a option" + by (metis assms none_eqvt not_None_eq pt1 some_eqvt) + moreover have "(pi1 @ pi2) \ x = pi1 \ pi2 \ x" + for pi1 pi2 :: "('x \ 'x) list" and x :: "'a option" + by (metis assms none_eqvt option.collapse pt2 some_eqvt) + moreover have "pi1 \ x = pi2 \ x" + if "pi1 \ pi2" for pi1 pi2 :: "('x \ 'x) list" and x :: "'a option" + using that pt3[OF pta] by (metis none_eqvt not_Some_eq some_eqvt) + ultimately show ?thesis + by (auto simp add: pt_def) +qed lemma pt_noption_inst: assumes pta: "pt TYPE('a) TYPE('x)" shows "pt TYPE('a noption) TYPE('x)" -apply(auto simp only: pt_def) -apply(case_tac "x") -apply(simp_all add: pt1[OF pta]) -apply(case_tac "x") -apply(simp_all add: pt2[OF pta]) -apply(case_tac "x") -apply(simp_all add: pt3[OF pta]) -done +proof - + have "([]::('x \ _) list) \ x = x" for x :: "'a noption" + by (metis assms nnone_eqvt noption.exhaust nsome_eqvt pt1) + moreover have "(pi1 @ pi2) \ x = pi1 \ pi2 \ x" + for pi1 pi2 :: "('x \ 'x) list" and x :: "'a noption" + using pt2[OF pta] + by (metis nnone_eqvt noption.exhaust nsome_eqvt) + moreover have "pi1 \ x = pi2 \ x" + if "pi1 \ pi2" + for pi1 pi2 :: "('x \ 'x) list" + and x :: "'a noption" + using that pt3[OF pta] by (metis nnone_eqvt noption.exhaust nsome_eqvt) + ultimately show ?thesis + by (auto simp add: pt_def) +qed lemma pt_nprod_inst: assumes pta: "pt TYPE('a) TYPE('x)" and ptb: "pt TYPE('b) TYPE('x)" shows "pt TYPE(('a,'b) nprod) TYPE('x)" - apply(auto simp add: pt_def) - apply(case_tac x) - apply(simp add: pt1[OF pta] pt1[OF ptb]) - apply(case_tac x) - apply(simp add: pt2[OF pta] pt2[OF ptb]) - apply(case_tac x) - apply(simp add: pt3[OF pta] pt3[OF ptb]) - done +proof - + have "([]::('x \ _) list) \ x = x" + for x :: "('a, 'b) nprod" + by (metis assms(1) nprod.exhaust perm_nprod.simps pt1 ptb) + moreover have "(pi1 @ pi2) \ x = pi1 \ pi2 \ x" + for pi1 pi2 :: "('x \ 'x) list" and x :: "('a, 'b) nprod" + using pt2[OF pta] pt2[OF ptb] + by (metis nprod.exhaust perm_nprod.simps) + moreover have "pi1 \ x = pi2 \ x" + if "pi1 \ pi2" for pi1 pi2 :: "('x \ 'x) list" and x :: "('a, 'b) nprod" + using that pt3[OF pta] pt3[OF ptb] by (smt (verit) nprod.exhaust perm_nprod.simps) + ultimately show ?thesis + by (auto simp add: pt_def) +qed + section \further lemmas for permutation types\ (*==============================================*) @@ -1236,12 +1162,7 @@ assumes pt: "pt TYPE('a) TYPE('x)" and at: "at TYPE('x)" shows "[(a,b)]\([(b,a)]\x) = x" -apply(simp add: pt2[OF pt,symmetric]) -apply(rule trans) -apply(rule pt3[OF pt]) -apply(rule at_ds5'[OF at]) -apply(rule pt1[OF pt]) -done + by (metis assms at_ds5 pt_def pt_swap_bij) lemma pt_swap_bij'': fixes a :: "'x" @@ -1249,11 +1170,7 @@ assumes pt: "pt TYPE('a) TYPE('x)" and at: "at TYPE('x)" shows "[(a,a)]\x = x" -apply(rule trans) -apply(rule pt3[OF pt]) -apply(rule at_ds1[OF at]) -apply(rule pt1[OF pt]) -done + by (metis assms at_ds1 pt_def) lemma supp_singleton: shows "supp {x} = supp x" @@ -1320,14 +1237,6 @@ shows "(pi\x)\X" using a by (simp add: pt_set_bij1[OF pt, OF at]) -(* FIXME: is this lemma needed anywhere? *) -lemma pt_set_bij3: - fixes pi :: "'x prm" - and x :: "'a" - and X :: "'a set" - shows "pi\(x\X) = (x\X)" -by (simp add: perm_bool) - lemma pt_subseteq_eqvt: fixes pi :: "'x prm" and Y :: "'a set" @@ -1351,10 +1260,13 @@ 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 +proof - + have "\y. x = pi \ y \ P y" + if "P (rev pi \ x)" for x + using that by (metis at pt pt_pi_rev) + then show ?thesis + by (auto simp add: perm_set_def pt_rev_pi [OF assms]) +qed \ \some helper lemmas for the pt_perm_supp_ineq lemma\ lemma Collect_permI: @@ -1459,14 +1371,7 @@ assumes pt: "pt TYPE('a) TYPE('x)" and at: "at TYPE('x)" shows "(pi\((supp x)::'x set)) = supp (pi\x)" -apply(rule pt_perm_supp_ineq) -apply(rule pt) -apply(rule at_pt_inst) -apply(rule at)+ -apply(rule cp_pt_inst) -apply(rule pt) -apply(rule at) -done + by (rule pt_perm_supp_ineq) (auto simp: assms at_pt_inst cp_pt_inst) lemma pt_supp_finite_pi: fixes pi :: "'x prm" @@ -1475,10 +1380,7 @@ and at: "at TYPE('x)" and f: "finite ((supp x)::'x set)" shows "finite ((supp (pi\x))::'x set)" -apply(simp add: pt_perm_supp[OF pt, OF at, symmetric]) -apply(simp add: pt_set_finite_ineq[OF at_pt_inst[OF at], OF at]) -apply(rule f) -done + by (metis at at_pt_inst f pt pt_perm_supp pt_set_finite_ineq) lemma pt_fresh_left_ineq: fixes pi :: "'x prm" @@ -1489,10 +1391,8 @@ and at: "at TYPE('x)" and cp: "cp TYPE('a) TYPE('x) TYPE('y)" shows "a\(pi\x) = ((rev pi)\a)\x" -apply(simp add: fresh_def) -apply(simp add: pt_set_bij1[OF ptb, OF at]) -apply(simp add: pt_perm_supp_ineq[OF pta, OF ptb, OF at, OF cp]) -done + using pt_perm_supp_ineq[OF pta, OF ptb, OF at, OF cp] pt_set_bij1[OF ptb, OF at] + by (simp add: fresh_def) lemma pt_fresh_right_ineq: fixes pi :: "'x prm" @@ -1503,10 +1403,7 @@ and at: "at TYPE('x)" and cp: "cp TYPE('a) TYPE('x) TYPE('y)" shows "(pi\a)\x = a\((rev pi)\x)" -apply(simp add: fresh_def) -apply(simp add: pt_set_bij1[OF ptb, OF at]) -apply(simp add: pt_perm_supp_ineq[OF pta, OF ptb, OF at, OF cp]) -done + by (simp add: assms pt_fresh_left_ineq) lemma pt_fresh_bij_ineq: fixes pi :: "'x prm" @@ -1517,9 +1414,7 @@ and at: "at TYPE('x)" and cp: "cp TYPE('a) TYPE('x) TYPE('y)" shows "(pi\a)\(pi\x) = a\x" -apply(simp add: pt_fresh_left_ineq[OF pta, OF ptb, OF at, OF cp]) -apply(simp add: pt_rev_pi[OF ptb, OF at]) -done + using assms pt_bij1 pt_fresh_right_ineq by fastforce lemma pt_fresh_left: fixes pi :: "'x prm" @@ -1528,14 +1423,7 @@ assumes pt: "pt TYPE('a) TYPE('x)" and at: "at TYPE('x)" shows "a\(pi\x) = ((rev pi)\a)\x" -apply(rule pt_fresh_left_ineq) -apply(rule pt) -apply(rule at_pt_inst) -apply(rule at)+ -apply(rule cp_pt_inst) -apply(rule pt) -apply(rule at) -done + by (simp add: assms at_pt_inst cp_pt_inst pt_fresh_left_ineq) lemma pt_fresh_right: fixes pi :: "'x prm" @@ -1544,14 +1432,7 @@ assumes pt: "pt TYPE('a) TYPE('x)" and at: "at TYPE('x)" shows "(pi\a)\x = a\((rev pi)\x)" -apply(rule pt_fresh_right_ineq) -apply(rule pt) -apply(rule at_pt_inst) -apply(rule at)+ -apply(rule cp_pt_inst) -apply(rule pt) -apply(rule at) -done + by (simp add: assms at_pt_inst cp_pt_inst pt_fresh_right_ineq) lemma pt_fresh_bij: fixes pi :: "'x prm" @@ -1560,14 +1441,7 @@ assumes pt: "pt TYPE('a) TYPE('x)" and at: "at TYPE('x)" shows "(pi\a)\(pi\x) = a\x" -apply(rule pt_fresh_bij_ineq) -apply(rule pt) -apply(rule at_pt_inst) -apply(rule at)+ -apply(rule cp_pt_inst) -apply(rule pt) -apply(rule at) -done + by (metis assms pt_bij1 pt_fresh_right) lemma pt_fresh_bij1: fixes pi :: "'x prm" @@ -1796,21 +1670,20 @@ assumes pt: "pt TYPE('a) TYPE('x)" and at: "at TYPE('x)" shows "pi\(\(x::'a). P x) = (\(x::'a). pi\(P ((rev pi)\x)))" -apply(auto simp add: perm_bool perm_fun_def) -apply(drule_tac x="pi\x" in spec) -apply(simp add: pt_rev_pi[OF pt, OF at]) -done + by (smt (verit, ccfv_threshold) assms pt_bij1 true_eqvt) lemma pt_ex_eqvt: fixes pi :: "'x prm" and x :: "'a" assumes pt: "pt TYPE('a) TYPE('x)" and at: "at TYPE('x)" - shows "pi\(\(x::'a). P x) = (\(x::'a). pi\(P ((rev pi)\x)))" -apply(auto simp add: perm_bool perm_fun_def) -apply(rule_tac x="pi\x" in exI) -apply(simp add: pt_rev_pi[OF pt, OF at]) -done +shows "pi\(\(x::'a). P x) = (\(x::'a). pi\(P ((rev pi)\x)))" +proof - + have "\x. P x \ P (rev pi \ pi \ x)" + by (simp add: assms(1) at pt_rev_pi) + then show ?thesis + by(auto simp add: perm_bool perm_fun_def) +qed lemma pt_ex1_eqvt: fixes pi :: "'x prm" @@ -1828,12 +1701,7 @@ and at: "at TYPE('x)" and unique: "\!x. P x" shows "pi\(THE(x::'a). P x) = (THE(x::'a). pi\(P ((rev pi)\x)))" - apply(rule the1_equality [symmetric]) - apply(simp add: pt_ex1_eqvt[OF pt at,symmetric]) - apply(simp add: perm_bool unique) - apply(simp add: perm_bool pt_rev_pi [OF pt at]) - apply(rule theI'[OF unique]) - done + by (smt (verit, best) assms perm_bool_def pt_rev_pi theI_unique unique) section \facts about supports\ (*==============================*) @@ -1934,13 +1802,17 @@ assumes pt: "pt TYPE('a) TYPE('x)" and at: "at TYPE ('x)" and a: "\x\X. (\(a::'x) (b::'x). a\S\b\S \ ([(a,b)]\x)\X)" - shows "S supports X" -using a -apply(auto simp add: supports_def) -apply(simp add: pt_set_bij1a[OF pt, OF at]) -apply(force simp add: pt_swap_bij[OF pt, OF at]) -apply(simp add: pt_set_bij1a[OF pt, OF at]) -done +shows "S supports X" +proof - + have "x \ X" + if "a \ S" "b \ S" and "x \ [(a, b)] \ X" for a b x + using that by (metis a assms(1) at pt_pi_rev pt_set_bij1a) + moreover have "x \ [(a, b)] \ X" + if "a \ S" "b \ S" and "x \ X" for a b x + using that by (simp add: a assms(1) at pt_set_bij1a) + ultimately show ?thesis + by (meson subsetI subset_antisym supports_def) +qed lemma supports_fresh: fixes S :: "'x set" @@ -1950,10 +1822,7 @@ and a2: "finite S" and a3: "a\S" shows "a\x" -proof (simp add: fresh_def) - have "(supp x)\S" using a1 a2 by (rule supp_is_subset) - thus "a\(supp x)" using a3 by force -qed + by (meson assms fresh_def in_mono supp_is_subset) lemma at_fin_set_supports: fixes X::"'x set" @@ -1969,12 +1838,7 @@ assumes a1:"infinite X" and a2:"\b\X. P(b)" shows "infinite {b\X. P(b)}" - using a1 a2 - apply auto - apply (subgoal_tac "infinite (X - {b\X. P b})") - apply (simp add: set_diff_eq) - apply (simp add: Diff_infinite_finite) - done + using assms rev_finite_subset by fastforce lemma at_fin_set_supp: fixes X::"'x set" @@ -2206,32 +2070,15 @@ proof - have pt_x: "pt TYPE('x) TYPE('x)" by (force intro: at_pt_inst at) show ?thesis - proof (rule equalityI) - show "pi\(\x\X. f x) \ (\x\(pi\X). (pi\f) x)" - apply(auto simp add: perm_set_def) - apply(rule_tac x="pi\xb" in exI) - apply(rule conjI) - apply(rule_tac x="xb" in exI) - apply(simp) - apply(subgoal_tac "(pi\f) (pi\xb) = pi\(f xb)")(*A*) - apply(simp) - apply(rule pt_set_bij2[OF pt_x, OF at]) - apply(assumption) - (*A*) - apply(rule sym) - apply(rule pt_fun_app_eq[OF pt, OF at]) - done - next - show "(\x\(pi\X). (pi\f) x) \ pi\(\x\X. f x)" - apply(auto simp add: perm_set_def) - apply(rule_tac x="(rev pi)\x" in exI) - apply(rule conjI) - apply(simp add: pt_pi_rev[OF pt_x, OF at]) - apply(rule_tac x="xb" in bexI) - apply(simp add: pt_set_bij1[OF pt_x, OF at]) - apply(simp add: pt_fun_app_eq[OF pt, OF at]) - apply(assumption) - done + proof - + have "\x. (\u. x = pi \ u \ u \ X) \ pi \ z \ (pi \ f) x" + if "y \ X" and "z \ f y" for y z + using that by (metis assms at_pt_inst pt_fun_app_eq pt_set_bij) + moreover have "\u. x = pi \ u \ (\x\X. u \ f x)" + if "x \ (pi \ f) (pi \ y)" and "y \ X" for x y + using that by (metis at at_pi_rev pt pt_fun_app_eq pt_set_bij1a pt_x) + ultimately show ?thesis + by (auto simp: perm_set_def) qed qed @@ -2241,34 +2088,14 @@ assumes pt: "pt TYPE('a) TYPE('x)" and at: "at TYPE('x)" shows "pi\(X_to_Un_supp X) = ((X_to_Un_supp (pi\X))::'x set)" - apply(simp add: X_to_Un_supp_def) - apply(simp add: UNION_f_eqvt[OF pt, OF at] perm_fun_def) - apply(simp add: pt_perm_supp[OF pt, OF at]) - apply(simp add: pt_pi_rev[OF pt, OF at]) - done + by (metis UNION_f_eqvt X_to_Un_supp_def assms pt_fun_eq pt_perm_supp) lemma Union_supports_set: fixes X::"('a set)" assumes pt: "pt TYPE('a) TYPE('x)" and at: "at TYPE('x)" shows "(\x\X. ((supp x)::'x set)) supports X" - apply(simp add: supports_def fresh_def[symmetric]) - apply(rule allI)+ - apply(rule impI) - apply(erule conjE) - apply(simp add: perm_set_def) - apply(auto) - apply(subgoal_tac "[(a,b)]\xa = xa")(*A*) - apply(simp) - apply(rule pt_fresh_fresh[OF pt, OF at]) - apply(force) - apply(force) - apply(rule_tac x="x" in exI) - apply(simp) - apply(rule sym) - apply(rule pt_fresh_fresh[OF pt, OF at]) - apply(force)+ - done + by (simp add: assms fresh_def pt_fresh_fresh supports_set) lemma Union_of_fin_supp_sets: fixes X::"('a set)" @@ -2286,19 +2113,14 @@ shows "(\x\X. ((supp x)::'x set)) \ supp X" proof - have "supp ((X_to_Un_supp X)::'x set) \ ((supp X)::'x set)" - apply(rule pt_empty_supp_fun_subset) - apply(force intro: pt_set_inst at_pt_inst pt at)+ - apply(rule pt_eqvt_fun2b) - apply(force intro: pt_set_inst at_pt_inst pt at)+ - apply(rule allI)+ - apply(rule X_to_Un_supp_eqvt[OF pt, OF at]) - done + proof (rule pt_empty_supp_fun_subset) + show "supp (\a. X_to_Un_supp (a::'a set)::'x set) = ({}::'x set)" + by (simp add: X_to_Un_supp_eqvt at at_pt_inst pt pt_eqvt_fun2b pt_set_inst) + qed (use assms at_pt_inst pt_set_inst in auto) hence "supp (\x\X. ((supp x)::'x set)) \ ((supp X)::'x set)" by (simp add: X_to_Un_supp_def) moreover have "supp (\x\X. ((supp x)::'x set)) = (\x\X. ((supp x)::'x set))" - apply(rule at_fin_set_supp[OF at]) - apply(rule Union_of_fin_supp_sets[OF fs, OF fi]) - done + using Union_of_fin_supp_sets at at_fin_set_supp fi fs by auto ultimately show ?thesis by force qed @@ -2309,12 +2131,15 @@ and fs: "fs TYPE('a) TYPE('x)" and fi: "finite X" shows "(supp X) = (\x\X. ((supp x)::'x set))" -apply(rule equalityI) -apply(rule supp_is_subset) -apply(rule Union_supports_set[OF pt, OF at]) -apply(rule Union_of_fin_supp_sets[OF fs, OF fi]) -apply(rule Union_included_in_supp[OF pt, OF at, OF fs, OF fi]) -done +proof (rule equalityI) + have "finite (\ (supp ` X)::'x set)" + using Union_of_fin_supp_sets fi fs by blast + then show "(supp X::'x set) \ \ (supp ` X)" + by (metis Union_supports_set at pt supp_is_subset) +next + show "(\x\X. (supp x::'x set)) \ supp X" + by (simp add: Union_included_in_supp at fi fs pt) +qed lemma supp_fin_union: fixes X::"('a set)" @@ -2353,9 +2178,7 @@ and f1: "finite X" and f2: "finite Y" shows "a\(X\Y) = (a\X \ a\Y)" -apply(simp add: fresh_def) -apply(simp add: supp_fin_union[OF pt, OF at, OF fs, OF f1, OF f2]) -done + by (simp add: assms fresh_def supp_fin_union) lemma fresh_fin_insert: fixes X::"('a set)" @@ -2366,9 +2189,7 @@ and fs: "fs TYPE('a) TYPE('x)" and f: "finite X" shows "a\(insert x X) = (a\x \ a\X)" -apply(simp add: fresh_def) -apply(simp add: supp_fin_insert[OF pt, OF at, OF fs, OF f]) -done + by (simp add: assms fresh_def supp_fin_insert) lemma fresh_fin_insert1: fixes X::"('a set)" @@ -2447,11 +2268,11 @@ lemma fresh_star_Un_elim: "((S \ T) \* c \ PROP C) \ (S \* c \ T \* c \ PROP C)" - apply rule - apply (simp_all add: fresh_star_def) - apply (erule meta_mp) - apply blast - done +proof + assume \
: "(S \ T) \* c \ PROP C" and c: "S \* c" "T \* c" + show "PROP C" + using c by (intro \
) (metis Un_iff fresh_star_set) +qed (auto simp: fresh_star_def) lemma fresh_star_insert_elim: "(insert x S \* c \ PROP C) \ (x \ c \ S \* c \ PROP C)" @@ -2485,22 +2306,22 @@ and cp: "cp TYPE('a) TYPE('x) TYPE('y)" shows "(pi\a)\*(pi\x) = a\*x" and "(pi\b)\*(pi\x) = b\*x" -apply(unfold fresh_star_def) -apply(auto) -apply(drule_tac x="pi\xa" in bspec) -apply(erule pt_set_bij2[OF ptb, OF at]) -apply(simp add: fresh_star_def pt_fresh_bij_ineq[OF pta, OF ptb, OF at, OF cp]) -apply(drule_tac x="(rev pi)\xa" in bspec) -apply(simp add: pt_set_bij1[OF ptb, OF at]) -apply(simp add: pt_fresh_left_ineq[OF pta, OF ptb, OF at, OF cp]) -apply(drule_tac x="pi\xa" in bspec) -apply(simp add: pt_set_bij1[OF ptb, OF at]) -apply(simp add: set_eqvt pt_rev_pi[OF pt_list_inst[OF ptb], OF at]) -apply(simp add: pt_fresh_bij_ineq[OF pta, OF ptb, OF at, OF cp]) -apply(drule_tac x="(rev pi)\xa" in bspec) -apply(simp add: pt_set_bij1[OF ptb, OF at] set_eqvt) -apply(simp add: pt_fresh_left_ineq[OF pta, OF ptb, OF at, OF cp]) -done + unfolding fresh_star_def +proof - + have "y \ x" if "\z\pi \ a. z \ pi \ x" and "y \ a" for y + using that by (meson assms at pt_fresh_bij_ineq pt_set_bij2) + moreover have "y \ pi \ x" if "\z\a. z \ x" and "y \ pi \ a" for y + using that by (simp add: assms pt_fresh_left_ineq pt_set_bij1a) + moreover have "y \ x" + if "\z\set (pi \ b). z \ pi \ x" and "y \ set b" for y + using that by (metis at cp pt_fresh_bij_ineq pt_set_bij pta ptb set_eqvt) + moreover have "y \ pi \ x" + if "\z\set b. z \ x" and "y \ set (pi \ b)" for y + using that by (metis at cp pt_fresh_left_ineq pt_set_bij1a pta ptb set_eqvt) + ultimately show "(\xa\pi \ a. xa \ pi \ x) = (\xa\a. xa \ x)" "(\xa\set (pi \ b). xa \ pi \ x) = (\xa\set b. xa \ x)" + by blast+ +qed + lemma pt_fresh_star_bij: fixes pi :: "'x prm" @@ -2511,21 +2332,10 @@ and at: "at TYPE('x)" shows "(pi\a)\*(pi\x) = a\*x" and "(pi\b)\*(pi\x) = b\*x" -apply(rule pt_fresh_star_bij_ineq(1)) -apply(rule pt) -apply(rule at_pt_inst) -apply(rule at)+ -apply(rule cp_pt_inst) -apply(rule pt) -apply(rule at) -apply(rule pt_fresh_star_bij_ineq(2)) -apply(rule pt) -apply(rule at_pt_inst) -apply(rule at)+ -apply(rule cp_pt_inst) -apply(rule pt) -apply(rule at) -done +proof (rule pt_fresh_star_bij_ineq) + show "(pi \ b) \* (pi \ x) = b \* x" + by (simp add: at at_pt_inst cp_pt_inst pt pt_fresh_star_bij_ineq) +qed (auto simp: at pt at_pt_inst cp_pt_inst) lemma pt_fresh_star_eqvt: fixes pi :: "'x prm" @@ -2582,16 +2392,9 @@ shows "pi\x = x" using a apply(induct pi) + apply (simp add: assms(1) pt1) apply(auto simp add: fresh_star_def fresh_list_cons fresh_prod pt1[OF pt]) -apply(subgoal_tac "((a,b)#pi)\x = ([(a,b)]@pi)\x") -apply(simp only: pt2[OF pt]) -apply(rule pt_fresh_fresh[OF pt at]) -apply(simp add: fresh_def at_supp[OF at]) -apply(blast) -apply(simp add: fresh_def at_supp[OF at]) -apply(blast) -apply(simp add: pt2[OF pt]) -done + by (metis Cons_eq_append_conv append_self_conv2 assms(1) at at_fresh fresh_def pt2 pt_fresh_fresh) section \Infrastructure lemmas for strong rule inductions\ (*==========================================================*) @@ -2685,63 +2488,43 @@ assumes c1: "cp TYPE ('a) TYPE('x) TYPE('y)" shows "cp TYPE ('a list) TYPE('x) TYPE('y)" using c1 -apply(simp add: cp_def) -apply(auto) -apply(induct_tac x) -apply(auto) -done +apply(clarsimp simp add: cp_def) + by (induct_tac x) auto lemma cp_set_inst: assumes c1: "cp TYPE ('a) TYPE('x) TYPE('y)" shows "cp TYPE ('a set) TYPE('x) TYPE('y)" -using c1 -apply(simp add: cp_def) -apply(auto) -apply(auto simp add: perm_set_def) -apply(rule_tac x="pi2\xc" in exI) -apply(auto) -done + using c1 + unfolding cp_def perm_set_def + by (smt (verit) Collect_cong mem_Collect_eq) + lemma cp_option_inst: assumes c1: "cp TYPE ('a) TYPE('x) TYPE('y)" shows "cp TYPE ('a option) TYPE('x) TYPE('y)" -using c1 -apply(simp add: cp_def) -apply(auto) -apply(case_tac x) -apply(auto) -done + using c1 unfolding cp_def by (metis none_eqvt not_None_eq some_eqvt) lemma cp_noption_inst: assumes c1: "cp TYPE ('a) TYPE('x) TYPE('y)" shows "cp TYPE ('a noption) TYPE('x) TYPE('y)" -using c1 -apply(simp add: cp_def) -apply(auto) -apply(case_tac x) -apply(auto) -done + using c1 unfolding cp_def + by (metis nnone_eqvt noption.exhaust nsome_eqvt) lemma cp_unit_inst: shows "cp TYPE (unit) TYPE('x) TYPE('y)" -apply(simp add: cp_def) -done + by (simp add: cp_def) lemma cp_bool_inst: shows "cp TYPE (bool) TYPE('x) TYPE('y)" -apply(simp add: cp_def) -apply(rule allI)+ -apply(induct_tac x) -apply(simp_all) -done + apply(clarsimp simp add: cp_def) + by (induct_tac x) auto lemma cp_prod_inst: assumes c1: "cp TYPE ('a) TYPE('x) TYPE('y)" and c2: "cp TYPE ('b) TYPE('x) TYPE('y)" shows "cp TYPE ('a\'b) TYPE('x) TYPE('y)" using c1 c2 -apply(simp add: cp_def) -done + by (simp add: cp_def) lemma cp_fun_inst: assumes c1: "cp TYPE ('a) TYPE('x) TYPE('y)" @@ -2749,11 +2532,8 @@ and pt: "pt TYPE ('y) TYPE('x)" and at: "at TYPE ('x)" shows "cp TYPE ('a\'b) TYPE('x) TYPE('y)" -using c1 c2 -apply(auto simp add: cp_def perm_fun_def fun_eq_iff) -apply(simp add: rev_eqvt[symmetric]) -apply(simp add: pt_rev_pi[OF pt_list_inst[OF pt_prod_inst[OF pt, OF pt]], OF at]) -done + using c1 c2 + by(auto simp add: cp_def perm_fun_def fun_eq_iff at pt pt_list_inst pt_prod_inst pt_rev_pi rev_eqvt) section \Andy's freshness lemma\ @@ -2856,9 +2636,7 @@ and f1: "finite ((supp h)::'x set)" and a: "a\h" "a\h a" shows "(fresh_fun h) = (h a)" -apply(rule fresh_fun_app[OF pt, OF at, OF f1]) -apply(auto simp add: fresh_prod intro: a) -done + by (meson assms fresh_fun_app fresh_prod pt) lemma fresh_fun_equiv_ineq: fixes h :: "'y\'a" @@ -2938,11 +2716,7 @@ and f1: "finite ((supp h)::'x set)" and a: "\(a::'x). a\(h,h a)" shows "((supp h)::'x set) supports (fresh_fun h)" - apply(simp add: supports_def fresh_def[symmetric]) - apply(auto) - apply(simp add: fresh_fun_equiv[OF pt, OF at, OF f1, OF a]) - apply(simp add: pt_fresh_fresh[OF pt_fun_inst[OF at_pt_inst[OF at], OF pt], OF at, OF at]) - done + by(simp flip: fresh_def add: supports_def assms at_pt_inst fresh_fun_equiv pt_fresh_fresh pt_fun_inst) section \Abstraction function\ (*==============================*) @@ -2951,7 +2725,7 @@ assumes pt: "pt TYPE('a) TYPE('x)" and at: "at TYPE('x)" shows "pt TYPE('x\('a noption)) TYPE('x)" - by (rule pt_fun_inst[OF at_pt_inst[OF at],OF pt_noption_inst[OF pt],OF at]) + by (simp add: at at_pt_inst pt pt_fun_inst pt_noption_inst) definition abs_fun :: "'x\'a\('x\('a noption))" (\[_]._\ [100,100] 100) where "[a].x \ (\b. (if b=a then nSome(x) else (if b\x then nSome([(a,b)]\x) else nNone)))" @@ -2973,26 +2747,22 @@ and ptb: "pt TYPE('y) TYPE('x)" and at: "at TYPE('x)" and cp: "cp TYPE('a) TYPE('x) TYPE('y)" - shows "pi\([a].x) = [(pi\a)].(pi\x)" - apply(simp add: abs_fun_def perm_fun_def abs_fun_if) - apply(simp only: fun_eq_iff) - apply(rule allI) - apply(subgoal_tac "(((rev pi)\(xa::'y)) = (a::'y)) = (xa = pi\a)")(*A*) - apply(subgoal_tac "(((rev pi)\xa)\x) = (xa\(pi\x))")(*B*) - apply(subgoal_tac "pi\([(a,(rev pi)\xa)]\x) = [(pi\a,xa)]\(pi\x)")(*C*) - apply(simp) -(*C*) - apply(simp add: cp1[OF cp]) - apply(simp add: pt_pi_rev[OF ptb, OF at]) -(*B*) - apply(simp add: pt_fresh_left_ineq[OF pta, OF ptb, OF at, OF cp]) -(*A*) - apply(rule iffI) - apply(rule pt_bij2[OF ptb, OF at, THEN sym]) - apply(simp) - apply(rule pt_bij2[OF ptb, OF at]) - apply(simp) -done +shows "pi\([a].x) = [(pi\a)].(pi\x)" + unfolding fun_eq_iff +proof + fix y + have "(((rev pi)\y) = a) = (y = pi\a)" + by (metis at pt_rev_pi ptb) + moreover + have "(((rev pi)\y)\x) = (y\(pi\x))" + by (simp add: assms pt_fresh_left_ineq) + moreover + have "pi\([(a,(rev pi)\y)]\x) = [(pi\a,y)]\(pi\x)" + using assms cp1[OF cp] by (simp add: pt_pi_rev) + ultimately + show "(pi \ [a].x) y = ([(pi \ a)].(pi \ x)) y" + by (simp add: abs_fun_def perm_fun_def) +qed lemma abs_fun_pi: fixes a :: "'x" @@ -3001,25 +2771,14 @@ assumes pt: "pt TYPE('a) TYPE('x)" and at: "at TYPE('x)" shows "pi\([a].x) = [(pi\a)].(pi\x)" -apply(rule abs_fun_pi_ineq) -apply(rule pt) -apply(rule at_pt_inst) -apply(rule at)+ -apply(rule cp_pt_inst) -apply(rule pt) -apply(rule at) -done + by (simp add: abs_fun_pi_ineq at at_pt_inst cp_pt_inst pt) lemma abs_fun_eq1: fixes x :: "'a" and y :: "'a" and a :: "'x" shows "([a].x = [a].y) = (x = y)" -apply(auto simp add: abs_fun_def) -apply(auto simp add: fun_eq_iff) -apply(drule_tac x="a" in spec) -apply(simp) -done + by (metis abs_fun_def noption.inject) lemma abs_fun_eq2: fixes x :: "'a" @@ -3210,10 +2969,7 @@ and as: "[a].x = [b].y" and fr: "c\a" "c\b" "c\x" "c\y" shows "x = [(a,c)]\[(b,c)]\y" -using as fr -apply(drule_tac sym) -apply(simp add: abs_fun_fresh[OF pt, OF at] pt_swap_bij[OF pt, OF at]) -done + using assms by (metis abs_fun_fresh pt_swap_bij) lemma abs_fun_supp_approx: fixes x :: "'a" @@ -3359,12 +3115,10 @@ and at: "at TYPE('x)" and cp: "cp TYPE('a) TYPE('x) TYPE('y)" and dj: "disjoint TYPE('y) TYPE('x)" - shows "((supp ([a].x))::'x set) = (supp x)" -apply(auto simp add: supp_def) -apply(auto simp add: abs_fun_pi_ineq[OF pta, OF ptb, OF at, OF cp]) -apply(auto simp add: dj_perm_forget[OF dj]) -apply(auto simp add: abs_fun_eq1) -done +shows "((supp ([a].x))::'x set) = (supp x)" +unfolding supp_def + using abs_fun_pi_ineq[OF pta, OF ptb, OF at, OF cp] dj_perm_forget[OF dj] + by (smt (verit, ccfv_threshold) Collect_cong abs_fun_eq1) lemma fresh_abs_fun_iff_ineq: fixes a :: "'y" @@ -3453,27 +3207,17 @@ have pt_prm: "pt TYPE('x prm) TYPE('x)" by (rule pt_list_inst[OF pt_prod_inst[OF pt, OF pt]]) from a show ?thesis - apply - - apply(auto simp add: prm_eq_def) - apply(rule_tac pi="rev pi3" in pt_bij4[OF pt, OF at]) - apply(rule trans) - apply(rule pt_perm_compose[OF pt, OF at]) - apply(simp add: pt_rev_pi[OF pt_prm, OF at]) - apply(rule sym) - apply(rule trans) - apply(rule pt_perm_compose[OF pt, OF at]) - apply(simp add: pt_rev_pi[OF pt_prm, OF at]) - done + by (auto simp add: prm_eq_def at pt pt_perm_compose') qed (************************) (* Various eqvt-lemmas *) -lemma Zero_nat_eqvt: +lemma Zero_nat_eqvt[simp]: shows "pi\(0::nat) = 0" by (auto simp add: perm_nat_def) -lemma One_nat_eqvt: +lemma One_nat_eqvt[simp]: shows "pi\(1::nat) = 1" by (simp add: perm_nat_def) @@ -3515,41 +3259,41 @@ shows "pi\(x div y) = (pi\x) div (pi\y)" by (simp add:perm_nat_def) -lemma Zero_int_eqvt: +lemma Zero_int_eqvt[simp]: shows "pi\(0::int) = 0" by (auto simp add: perm_int_def) -lemma One_int_eqvt: +lemma One_int_eqvt[simp]: shows "pi\(1::int) = 1" by (simp add: perm_int_def) -lemma numeral_int_eqvt: - shows "pi\((numeral n)::int) = numeral n" -by (simp add: perm_int_def perm_int_def) - -lemma neg_numeral_int_eqvt: - shows "pi\((- numeral n)::int) = - numeral n" -by (simp add: perm_int_def perm_int_def) +lemma numeral_int_eqvt[simp]: + shows "pi\((numeral n)::int) = numeral n" + using perm_int_def by blast + +lemma neg_numeral_int_eqvt[simp]: + shows "pi\((- numeral n)::int) = - numeral n" + by (simp add: perm_int_def) lemma max_int_eqvt: fixes x::"int" shows "pi\(max (x::int) y) = max (pi\x) (pi\y)" -by (simp add:perm_int_def) + by (simp add:perm_int_def) lemma min_int_eqvt: fixes x::"int" shows "pi\(min x y) = min (pi\x) (pi\y)" -by (simp add:perm_int_def) + by (simp add:perm_int_def) lemma plus_int_eqvt: fixes x::"int" shows "pi\(x + y) = (pi\x) + (pi\y)" -by (simp add:perm_int_def) + by (simp add:perm_int_def) lemma minus_int_eqvt: fixes x::"int" shows "pi\(x - y) = (pi\x) - (pi\y)" -by (simp add:perm_int_def) + by (simp add:perm_int_def) lemma mult_int_eqvt: fixes x::"int" @@ -3559,7 +3303,7 @@ lemma div_int_eqvt: fixes x::"int" shows "pi\(x div y) = (pi\x) div (pi\y)" -by (simp add:perm_int_def) + by (simp add:perm_int_def) (*******************************************************) (* Setup of the theorem attributes eqvt and eqvt_force *)