# HG changeset patch # User nipkow # Date 1722259616 -7200 # Node ID fbb38db0435d13354429867812267d38212fba27 # Parent 424b90ba7b6ff24d6e7ef09011ff8e558a412a47# Parent 9f8034d29365c3e50f96e7de0e4b2efc2f0d48d5 merged diff -r 9f8034d29365 -r fbb38db0435d src/Doc/Implementation/Logic.thy --- a/src/Doc/Implementation/Logic.thy Mon Jul 29 15:26:03 2024 +0200 +++ b/src/Doc/Implementation/Logic.thy Mon Jul 29 15:26:56 2024 +0200 @@ -1293,7 +1293,7 @@ @{define_ML Proofterm.reconstruct_proof: "theory -> term -> proof -> proof"} \\ @{define_ML Proofterm.expand_proof: "theory -> - (Proofterm.thm_header -> Thm_Name.T option) -> proof -> proof"} \\ + (Proofterm.thm_header -> Thm_Name.P option) -> proof -> proof"} \\ @{define_ML Proof_Checker.thm_of_proof: "theory -> proof -> thm"} \\ @{define_ML Proof_Syntax.read_proof: "theory -> bool -> bool -> string -> proof"} \\ @{define_ML Proof_Syntax.pretty_proof: "Proof.context -> proof -> Pretty.T"} \\ diff -r 9f8034d29365 -r fbb38db0435d src/HOL/Analysis/Improper_Integral.thy --- a/src/HOL/Analysis/Improper_Integral.thy Mon Jul 29 15:26:03 2024 +0200 +++ b/src/HOL/Analysis/Improper_Integral.thy Mon Jul 29 15:26:56 2024 +0200 @@ -959,9 +959,8 @@ by (simp add: dist_norm norm_minus_commute) also have "... \ \ * (b \ i - a \ i) / \v \ i - u \ i\ / (4 * content (cbox a b))" proof (intro mult_right_mono divide_left_mono divide_right_mono uvi) - show "norm (v - u) * \v \ i - u \ i\ > 0" - using u_less_v [OF \i \ Basis\] - by (auto simp: less_eq_real_def zero_less_mult_iff that) + show "\v \ i - u \ i\ > 0" + using u_less_v [OF \i \ Basis\] by force show "\ * (b \ i - a \ i) \ 0" using a_less_b \0 < \\ \i \ Basis\ by force qed auto diff -r 9f8034d29365 -r fbb38db0435d src/HOL/Decision_Procs/Approximation_Bounds.thy --- a/src/HOL/Decision_Procs/Approximation_Bounds.thy Mon Jul 29 15:26:03 2024 +0200 +++ b/src/HOL/Decision_Procs/Approximation_Bounds.thy Mon Jul 29 15:26:56 2024 +0200 @@ -956,13 +956,14 @@ using bnds_sqrt'[of ?sxx prec] by auto finally have "sqrt (1 + x*x) \ ub_sqrt prec ?sxx" . - hence "?R \ ?fR" by (auto simp: float_plus_up.rep_eq plus_up_def intro!: truncate_up_le) + hence \: "?R \ ?fR" by (auto simp: float_plus_up.rep_eq plus_up_def intro!: truncate_up_le) hence "0 < ?fR" and "0 < real_of_float ?fR" using \0 < ?R\ by auto have monotone: "?DIV \ x / ?R" proof - have "?DIV \ real_of_float x / ?fR" by (rule float_divl) - also have "\ \ x / ?R" by (rule divide_left_mono[OF \?R \ ?fR\ \0 \ real_of_float x\ mult_pos_pos[OF order_less_le_trans[OF divisor_gt0 \?R \ real_of_float ?fR\] divisor_gt0]]) + also have "\ \ x / ?R" + by (simp add: \ assms divide_left_mono divisor_gt0) finally show ?thesis . qed @@ -1081,16 +1082,16 @@ also have "\ \ sqrt (1 + x*x)" by (auto simp: float_plus_down.rep_eq plus_down_def float_round_down.rep_eq truncate_down_le) finally have "lb_sqrt prec ?sxx \ sqrt (1 + x*x)" . - hence "?fR \ ?R" + hence \: "?fR \ ?R" by (auto simp: float_plus_down.rep_eq plus_down_def truncate_down_le) + have monotone: "x / ?R \ (float_divr prec x ?fR)" + proof - have "0 < real_of_float ?fR" by (auto simp: float_plus_down.rep_eq plus_down_def float_round_down.rep_eq intro!: truncate_down_ge1 lb_sqrt_lower_bound order_less_le_trans[OF zero_less_one] truncate_down_nonneg add_nonneg_nonneg) - have monotone: "x / ?R \ (float_divr prec x ?fR)" - proof - - from divide_left_mono[OF \?fR \ ?R\ \0 \ real_of_float x\ mult_pos_pos[OF divisor_gt0 \0 < real_of_float ?fR\]] - have "x / ?R \ x / ?fR" . + then have "x / ?R \ x / ?fR" + using \ assms divide_left_mono by blast also have "\ \ ?DIV" by (rule float_divr) finally show ?thesis . qed diff -r 9f8034d29365 -r fbb38db0435d src/HOL/Deriv.thy --- a/src/HOL/Deriv.thy Mon Jul 29 15:26:03 2024 +0200 +++ b/src/HOL/Deriv.thy Mon Jul 29 15:26:56 2024 +0200 @@ -1656,34 +1656,6 @@ by (blast dest: DERIV_unique order_less_imp_le) qed -lemma pos_deriv_imp_strict_mono: - assumes "\x. (f has_real_derivative f' x) (at x)" - assumes "\x. f' x > 0" - shows "strict_mono f" -proof (rule strict_monoI) - fix x y :: real assume xy: "x < y" - from assms and xy have "\z>x. z < y \ f y - f x = (y - x) * f' z" - by (intro MVT2) (auto dest: connectedD_interval) - then obtain z where z: "z > x" "z < y" "f y - f x = (y - x) * f' z" by blast - note \f y - f x = (y - x) * f' z\ - also have "(y - x) * f' z > 0" using xy assms by (intro mult_pos_pos) auto - finally show "f x < f y" by simp -qed - -proposition deriv_nonneg_imp_mono: - assumes deriv: "\x. x \ {a..b} \ (g has_real_derivative g' x) (at x)" - assumes nonneg: "\x. x \ {a..b} \ g' x \ 0" - assumes ab: "a \ b" - shows "g a \ g b" -proof (cases "a < b") - assume "a < b" - from deriv have "\x. \x \ a; x \ b\ \ (g has_real_derivative g' x) (at x)" by simp - with MVT2[OF \a < b\] and deriv - obtain \ where \_ab: "\ > a" "\ < b" and g_ab: "g b - g a = (b - a) * g' \" by blast - from \_ab ab nonneg have "(b - a) * g' \ \ 0" by simp - with g_ab show ?thesis by simp -qed (insert ab, simp) - subsubsection \A function is constant if its derivative is 0 over an interval.\ @@ -1927,6 +1899,74 @@ apply (metis filterlim_at_top_mirror lim) done +proposition deriv_nonpos_imp_antimono: + assumes deriv: "\x. x \ {a..b} \ (g has_real_derivative g' x) (at x)" + assumes nonneg: "\x. x \ {a..b} \ g' x \ 0" + assumes "a \ b" + shows "g b \ g a" +proof - + have "- g a \ - g b" + proof (intro DERIV_nonneg_imp_nondecreasing [where f = "\x. - g x"] conjI exI) + fix x + assume x: "a \ x" "x \ b" + show "((\x. - g x) has_real_derivative - g' x) (at x)" + by (simp add: DERIV_minus deriv x) + show "0 \ - g' x" + by (simp add: nonneg x) + qed (rule \a\b\) + then show ?thesis by simp +qed + +lemma DERIV_nonneg_imp_increasing_open: + fixes a b :: real + and f :: "real \ real" + assumes "a \ b" + and "\x. a < x \ x < b \ (\y. DERIV f x :> y \ y \ 0)" + and con: "continuous_on {a..b} f" + shows "f a \ f b" +proof (cases "a=b") + case False + with \a\b\ have "a ?thesis" + have "\l z. a < z \ z < b \ DERIV f z :> l \ f b - f a = (b - a) * l" + by (rule MVT) (use assms \a real_differentiable_def in \force+\) + then obtain l z where z: "a < z" "z < b" "DERIV f z :> l" and "f b - f a = (b - a) * l" + by auto + with assms z f show False + by (metis DERIV_unique diff_ge_0_iff_ge zero_le_mult_iff) + qed +qed auto + +lemma DERIV_nonpos_imp_decreasing_open: + fixes a b :: real + and f :: "real \ real" + assumes "a \ b" + and "\x. a < x \ x < b \ \y. DERIV f x :> y \ y \ 0" + and con: "continuous_on {a..b} f" + shows "f a \ f b" +proof - + have "(\x. -f x) a \ (\x. -f x) b" + proof (rule DERIV_nonneg_imp_increasing_open [of a b]) + show "\x. \a < x; x < b\ \ \y. ((\x. - f x) has_real_derivative y) (at x) \ 0 \ y" + using assms + by (metis Deriv.field_differentiable_minus neg_0_le_iff_le) + show "continuous_on {a..b} (\x. - f x)" + using con continuous_on_minus by blast + qed (use assms in auto) + then show ?thesis + by simp +qed + + +proposition deriv_nonneg_imp_mono: (*DELETE*) + assumes deriv: "\x. x \ {a..b} \ (g has_real_derivative g' x) (at x)" + assumes nonneg: "\x. x \ {a..b} \ g' x \ 0" + assumes ab: "a \ b" + shows "g a \ g b" + by (metis DERIV_nonneg_imp_nondecreasing atLeastAtMost_iff assms) + text \Derivative of inverse function\ lemma DERIV_inverse_function: diff -r 9f8034d29365 -r fbb38db0435d src/HOL/Fields.thy --- a/src/HOL/Fields.thy Mon Jul 29 15:26:03 2024 +0200 +++ b/src/HOL/Fields.thy Mon Jul 29 15:26:56 2024 +0200 @@ -949,7 +949,7 @@ by (auto simp: field_simps zero_less_mult_iff mult_strict_right_mono) lemma divide_left_mono: - "\b \ a; 0 \ c; 0 < a*b\ \ c / a \ c / b" + "\b \ a; 0 \ c; 0 < b\ \ c / a \ c / b" by (auto simp: field_simps zero_less_mult_iff mult_right_mono) lemma divide_strict_left_mono_neg: @@ -957,16 +957,16 @@ by (auto simp: field_simps zero_less_mult_iff mult_strict_right_mono_neg) lemma mult_imp_div_pos_le: "0 < y \ x \ z * y \ x / y \ z" -by (subst pos_divide_le_eq, assumption+) + by (subst pos_divide_le_eq, assumption+) lemma mult_imp_le_div_pos: "0 < y \ z * y \ x \ z \ x / y" -by(simp add:field_simps) + by(simp add:field_simps) lemma mult_imp_div_pos_less: "0 < y \ x < z * y \ x / y < z" -by(simp add:field_simps) + by(simp add:field_simps) lemma mult_imp_less_div_pos: "0 < y \ z * y < x \ z < x / y" -by(simp add:field_simps) + by(simp add:field_simps) lemma frac_le: assumes "0 \ y" "x \ y" "0 < w" "w \ z" @@ -1010,6 +1010,11 @@ using assms by (force intro: mult_imp_less_div_pos mult_le_less_imp_less) qed +text \As above, with a better name\ +lemma divide_mono: + "\b \ a; c \ d; 0 < b; 0 \ c\ \ c / a \ d / b" + by (simp add: frac_le) + lemma less_half_sum: "a < b \ a < (a+b) / (1+1)" by (simp add: field_simps zero_less_two) @@ -1151,7 +1156,7 @@ lemma divide_right_mono_neg: "a \ b \ c \ 0 \ b / c \ a / c" by (auto dest: divide_right_mono [of _ _ "- c"]) -lemma divide_left_mono_neg: "a \ b \ c \ 0 \ 0 < a * b \ c / a \ c / b" +lemma divide_left_mono_neg: "a \ b \ c \ 0 \ 0 < a \ c / a \ c / b" by (auto simp add: mult.commute dest: divide_left_mono [of _ _ "- c"]) lemma inverse_le_iff: "inverse a \ inverse b \ (0 < a * b \ b \ a) \ (a * b \ 0 \ a \ b)" diff -r 9f8034d29365 -r fbb38db0435d src/HOL/Filter.thy --- a/src/HOL/Filter.thy Mon Jul 29 15:26:03 2024 +0200 +++ b/src/HOL/Filter.thy Mon Jul 29 15:26:56 2024 +0200 @@ -926,6 +926,10 @@ "eventually P sequentially \ (\N. \n\N. P n)" by (rule eventually_at_top_linorder) +lemma frequently_sequentially: + "frequently P sequentially \ (\N. \n\N. P n)" + by (simp add: frequently_def eventually_sequentially) + lemma sequentially_bot [simp, intro]: "sequentially \ bot" unfolding filter_eq_iff eventually_sequentially by auto diff -r 9f8034d29365 -r fbb38db0435d src/HOL/IMP/Compiler2.thy --- a/src/HOL/IMP/Compiler2.thy Mon Jul 29 15:26:03 2024 +0200 +++ b/src/HOL/IMP/Compiler2.thy Mon Jul 29 15:26:56 2024 +0200 @@ -194,14 +194,15 @@ lemma acomp_succs [simp]: "succs (acomp a) n = {n + 1 .. n + size (acomp a)}" by (induct a arbitrary: n) auto - -lemma acomp_size: - "(1::int) \ size (acomp a)" - by (induct a) auto - + lemma acomp_exits [simp]: "exits (acomp a) = {size (acomp a)}" - by (auto simp: exits_def acomp_size) +proof - + have "Suc 0 \ length (acomp a)" + by (induct a) auto + then show ?thesis + by (auto simp add: exits_def) +qed lemma bcomp_succs: "0 \ i \ diff -r 9f8034d29365 -r fbb38db0435d src/HOL/Int.thy --- a/src/HOL/Int.thy Mon Jul 29 15:26:03 2024 +0200 +++ b/src/HOL/Int.thy Mon Jul 29 15:26:56 2024 +0200 @@ -198,7 +198,7 @@ then have \int n \ int 1\ by (rule of_nat_mono) with \l - k = int n\ show ?Q - by simp + by (simp add: algebra_simps) qed qed diff -r 9f8034d29365 -r fbb38db0435d src/HOL/Library/Discrete.thy --- a/src/HOL/Library/Discrete.thy Mon Jul 29 15:26:03 2024 +0200 +++ b/src/HOL/Library/Discrete.thy Mon Jul 29 15:26:56 2024 +0200 @@ -63,7 +63,7 @@ then show ?thesis by (simp add: log_rec) qed -lemma log_exp [simp]: "log (2 ^ n) = n" +lemma log_power [simp]: "log (2 ^ n) = n" by (induct n) simp_all lemma log_mono: "mono log" diff -r 9f8034d29365 -r fbb38db0435d src/HOL/Nominal/Examples/Class1.thy --- a/src/HOL/Nominal/Examples/Class1.thy Mon Jul 29 15:26:03 2024 +0200 +++ b/src/HOL/Nominal/Examples/Class1.thy Mon Jul 29 15:26:56 2024 +0200 @@ -14,7 +14,7 @@ PR "string" | NOT "ty" | AND "ty" "ty" ("_ AND _" [100,100] 100) - | OR "ty" "ty" ("_ OR _" [100,100] 100) + | OR "ty" "ty" ("_ OR _" [100,100] 100) | IMP "ty" "ty" ("_ IMP _" [100,100] 100) instantiation ty :: size @@ -50,8 +50,8 @@ nominal_datatype trm = Ax "name" "coname" - | Cut "\coname\trm" "\name\trm" ("Cut <_>._ (_)._" [0,100,1000,100] 101) - | NotR "\name\trm" "coname" ("NotR (_)._ _" [100,100,100] 101) + | Cut "\coname\trm" "\name\trm" ("Cut <_>._ ('(_'))._" [0,0,0,100] 101) + | NotR "\name\trm" "coname" ("NotR ('(_'))._ _" [0,100,100] 101) | NotL "\coname\trm" "name" ("NotL <_>._ _" [0,100,100] 101) | AndR "\coname\trm" "\coname\trm" "coname" ("AndR <_>._ <_>._ _" [0,100,0,100,100] 101) | AndL1 "\name\trm" "name" ("AndL1 (_)._ _" [100,100,100] 101) @@ -73,7 +73,7 @@ text \renaming functions\ nominal_primrec (freshness_context: "(d::coname,e::coname)") - crename :: "trm \ coname \ coname \ trm" ("_[_\c>_]" [100,100,100] 100) + crename :: "trm \ coname \ coname \ trm" ("_[_\c>_]" [100,0,0] 100) where "(Ax x a)[d\c>e] = (if a=d then Ax x e else Ax x a)" | "\a\(d,e,N);x\M\ \ (Cut .M (x).N)[d\c>e] = Cut .(M[d\c>e]) (x).(N[d\c>e])" @@ -92,7 +92,7 @@ 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) + nrename :: "trm \ name \ name \ trm" ("_[_\n>_]" [100,0,0] 100) where "(Ax x a)[u\n>v] = (if x=u then Ax v a else Ax x a)" | "\a\N;x\(u,v,M)\ \ (Cut .M (x).N)[u\n>v] = Cut .(M[u\n>v]) (x).(N[u\n>v])" @@ -529,21 +529,24 @@ lemma fresh_fun_OrR1[eqvt_force]: fixes pi1::"name prm" and pi2::"coname prm" - shows "pi1\fresh_fun (\a'. Cut .(OrR1 .M a') (x).P)= - fresh_fun (pi1\(\a'. Cut .(OrR1 .M a') (x).P))" - and "pi2\fresh_fun (\a'. Cut .(OrR1 .M a') (x).P)= - fresh_fun (pi2\(\a'. Cut .(OrR1 .M a') (x).P))" - apply(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) + shows "pi1\fresh_fun (\a'. Cut .(OrR1 .M a') (x).P) = + fresh_fun (pi1\(\a'. Cut .(OrR1 .M a') (x).P))" (is "?t1") + and "pi2\fresh_fun (\a'. Cut .(OrR1 .M a') (x).P) = + fresh_fun (pi2\(\a'. Cut .(OrR1 .M a') (x).P))" (is "?t2") +proof - + obtain n::coname where "n\(P,M,b,pi1\P,pi1\M,pi1\b,pi1)" + by (meson exists_fresh(2) fs_coname1) + then show ?t1 + by perm_simp (force simp add: fresh_prod fresh_fun_simp_OrR1 calc_atm) + obtain n::coname where "n\(P,M,b,pi2\P,pi2\M,pi2\b,pi2)" + by (meson exists_fresh(2) fs_coname1) + then show ?t2 + by perm_simp + (force simp add: fresh_prod fresh_fun_simp_OrR1 at_prm_fresh[OF at_coname_inst] swap_simps) +qed lemma fresh_fun_simp_OrR2: - assumes a: "a'\P" "a'\M" "a'\b" + assumes "a'\P" "a'\M" "a'\b" shows "fresh_fun (\a'. Cut .(OrR2 .M a') (x).P) = Cut .(OrR2 .M a') (x).P" proof (rule fresh_fun_app [OF pt_coname_inst at_coname_inst]) obtain n::coname where "n\(x,P,b,M)" @@ -553,23 +556,26 @@ apply(simp add: fresh_prod abs_fresh) apply(fresh_guess) done -qed (use a in \fresh_guess|finite_guess\)+ +qed (use assms in \fresh_guess|finite_guess\)+ lemma fresh_fun_OrR2[eqvt_force]: fixes pi1::"name prm" and pi2::"coname prm" - shows "pi1\fresh_fun (\a'. Cut .(OrR2 .M a') (x).P)= - fresh_fun (pi1\(\a'. Cut .(OrR2 .M a') (x).P))" - and "pi2\fresh_fun (\a'. Cut .(OrR2 .M a') (x).P)= - fresh_fun (pi2\(\a'. Cut .(OrR2 .M a') (x).P))" - apply(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) + shows "pi1\fresh_fun (\a'. Cut .(OrR2 .M a') (x).P) = + fresh_fun (pi1\(\a'. Cut .(OrR2 .M a') (x).P))" (is "?t1") + and "pi2\fresh_fun (\a'. Cut .(OrR2 .M a') (x).P) = + fresh_fun (pi2\(\a'. Cut .(OrR2 .M a') (x).P))" (is "?t2") +proof - + obtain n::coname where "n\(P,M,b,pi1\P,pi1\M,pi1\b,pi1)" + by (meson exists_fresh(2) fs_coname1) + then show ?t1 + by perm_simp (force simp add: fresh_prod fresh_fun_simp_OrR2 calc_atm) + obtain n::coname where "n\(P,M,b,pi2\P,pi2\M,pi2\b,pi2)" + by (meson exists_fresh(2) fs_coname1) + then show ?t2 + by perm_simp + (force simp add: fresh_prod fresh_fun_simp_OrR2 at_prm_fresh[OF at_coname_inst] swap_simps) +qed lemma fresh_fun_simp_ImpR: assumes a: "a'\P" "a'\M" "a'\b" @@ -587,18 +593,21 @@ lemma fresh_fun_ImpR[eqvt_force]: fixes pi1::"name prm" and pi2::"coname prm" - shows "pi1\fresh_fun (\a'. Cut .(ImpR (y)..M a') (x).P)= - fresh_fun (pi1\(\a'. Cut .(ImpR (y)..M a') (x).P))" + shows "pi1\fresh_fun (\a'. Cut .(ImpR (y)..M a') (x).P) = + fresh_fun (pi1\(\a'. Cut .(ImpR (y)..M a') (x).P))" (is "?t1") and "pi2\fresh_fun (\a'. Cut .(ImpR (y)..M a') (x).P)= - fresh_fun (pi2\(\a'. Cut .(ImpR (y)..M a') (x).P))" - 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) + fresh_fun (pi2\(\a'. Cut .(ImpR (y)..M a') (x).P))" (is "?t2") +proof - + obtain n::coname where "n\(P,M,b,pi1\P,pi1\M,pi1\b,pi1)" + by (meson exists_fresh(2) fs_coname1) + then show ?t1 + by perm_simp (force simp add: fresh_prod fresh_fun_simp_ImpR calc_atm) + obtain n::coname where "n\(P,M,b,pi2\P,pi2\M,pi2\b,pi2)" + by (meson exists_fresh(2) fs_coname1) + then show ?t2 + by perm_simp + (force simp add: fresh_prod fresh_fun_simp_ImpR at_prm_fresh[OF at_coname_inst] swap_simps) +qed nominal_primrec (freshness_context: "(y::name,c::coname,P::trm)") substn :: "trm \ name \ coname \ trm \ trm" ("_{_:=<_>._}" [100,100,100,100] 100) @@ -664,7 +673,7 @@ done nominal_primrec (freshness_context: "(d::name,z::coname,P::trm)") - substc :: "trm \ coname \ name \ trm \ trm" ("_{_:=(_)._}" [100,100,100,100] 100) + substc :: "trm \ coname \ name \ trm \ trm" ("_{_:=('(_'))._}" [100,0,0,100] 100) where "(Ax x a){d:=(z).P} = (if d=a then Cut .(Ax x a) (z).P else Ax x a)" | "\a\(d,P,N);x\(z,P,M)\ \ (Cut .M (x).N){d:=(z).P} = @@ -1078,9 +1087,8 @@ 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: fresh_atm abs_fresh abs_supp fin_supp) -done + by (nominal_induct M avoiding: x c P rule: trm.strong_induct) + (auto simp: fresh_atm abs_fresh abs_supp fin_supp) lemma substc_rename1: assumes a: "c\(M,a)" @@ -1257,6 +1265,7 @@ unfolding eq2[symmetric] apply(auto simp: trm.inject) apply(simp_all add: alpha fresh_atm fresh_prod subst_fresh eqvts perm_fresh_fresh calc_atm) + apply (simp add: fresh_atm(2) substn_rename4) by (metis abs_fresh(2) fresh_atm(2) fresh_prod perm_fresh_fresh(2) substn_rename4) finally show ?thesis by simp qed @@ -1894,343 +1903,266 @@ (auto simp: calc_atm simp add: fresh_left abs_fresh) lemma not_fin_subst1: - assumes a: "\(fin M x)" + assumes "\(fin M x)" shows "\(fin (M{c:=(y).P}) x)" -using a [[simproc del: defined_all]] -apply(nominal_induct M avoiding: x c y P rule: trm.strong_induct) -apply(auto) -apply(drule fin_elims, simp) -apply(drule fin_elims, simp) -apply(drule fin_elims, simp) -apply(subgoal_tac "\a'::coname. a'\(trm{coname:=(y).P},P,x)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_NotR) -apply(erule fin.cases, simp_all add: trm.inject) -apply(rule exists_fresh'(2)[OF fs_coname1]) -apply(drule fin_elims, simp) -apply(drule fin_elims) -apply(auto)[1] -apply(drule freshn_after_substc) -apply(simp add: fin.intros) -apply(subgoal_tac "\a'::coname. a'\(trm1{coname3:=(y).P},trm2{coname3:=(y).P},P,coname1,coname2,coname3,x)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_AndR) -apply(erule fin.cases, simp_all add: trm.inject) -apply(subgoal_tac "\a'::coname. a'\(trm1{coname3:=(y).P},trm2{coname3:=(y).P},P,coname1,coname2,coname3,x)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_AndR) -apply(erule fin.cases, simp_all add: trm.inject) -apply(rule exists_fresh'(2)[OF fs_coname1]) -apply(erule fin.cases, simp_all add: trm.inject) -apply(drule fin_AndL1_elim) -apply(auto simp: 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: abs_fresh)[1] -apply(drule freshn_after_substc) -apply(subgoal_tac "name2\[name1].trm") -apply(simp add: fin.intros) -apply(simp add: abs_fresh) -apply(subgoal_tac "\a'::coname. a'\(trm{coname2:=(y).P},coname1,P,x)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_OrR1) -apply(erule fin.cases, simp_all add: trm.inject) -apply(rule exists_fresh'(2)[OF fs_coname1]) -apply(erule fin.cases, simp_all add: trm.inject) -apply(subgoal_tac "\a'::coname. a'\(trm{coname2:=(y).P},coname1,P,x)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_OrR2) -apply(erule fin.cases, simp_all add: trm.inject) -apply(rule exists_fresh'(2)[OF fs_coname1]) -apply(erule fin.cases, simp_all add: trm.inject) -apply(drule fin_OrL_elim) -apply(auto simp: abs_fresh)[1] -apply(drule freshn_after_substc)+ -apply(subgoal_tac "name3\[name1].trm1 \ name3\[name2].trm2") -apply(simp add: fin.intros) -apply(simp add: abs_fresh) -apply(subgoal_tac "\a'::coname. a'\(trm{coname2:=(y).P},coname1,P,x)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_ImpR) -apply(erule fin.cases, simp_all add: trm.inject) -apply(rule exists_fresh'(2)[OF fs_coname1]) -apply(erule fin.cases, simp_all add: trm.inject) -apply(drule fin_ImpL_elim) -apply(auto simp: abs_fresh)[1] -apply(drule freshn_after_substc)+ -apply(subgoal_tac "x\[name1].trm2") -apply(simp add: fin.intros) -apply(simp add: abs_fresh) -done +using assms +proof (nominal_induct M avoiding: x c y P rule: trm.strong_induct) + case (Ax name coname) + then show ?case + using fin_rest_elims(1) substc.simps(1) by presburger +next + case (Cut coname trm1 name trm2) + then show ?case + using fin_rest_elims(1) substc.simps(1) by simp presburger +next + case (NotR name trm coname) + obtain a'::coname where "a'\(trm{coname:=(y).P},P,x)" + by (meson exists_fresh(2) fs_coname1) + with NotR show ?case + apply (simp add: fresh_prod trm.inject) + using fresh_fun_simp_NotR fin_rest_elims by fastforce +next + case (NotL coname trm name) + then show ?case + using fin_NotL_elim freshn_after_substc by simp blast +next + case (AndR coname1 trm1 coname2 trm2 coname3) + obtain a'::coname where "a'\(trm1{coname3:=(y).P},trm2{coname3:=(y).P},P,coname1,coname2,coname3,x)" + by (meson exists_fresh(2) fs_coname1) + with AndR show ?case + apply (simp add: fresh_prod trm.inject) + using fresh_fun_simp_AndR fin_rest_elims by fastforce +next + case (AndL1 name1 trm name2) + then show ?case + using fin_AndL1_elim freshn_after_substc + by simp (metis abs_fresh(1) fin.intros(3)) +next + case (AndL2 name1 trm name2) + then show ?case + using fin_AndL2_elim freshn_after_substc + by simp (metis abs_fresh(1) fin.intros(4)) +next + case (OrR1 coname1 trm coname2) + obtain a'::coname where "a'\(trm{coname2:=(y).P},coname1,P,x)" + by (meson exists_fresh(2) fs_coname1) + with OrR1 show ?case + apply (simp add: fresh_prod trm.inject) + using fresh_fun_simp_OrR1 fin_rest_elims by fastforce +next + case (OrR2 coname1 trm coname2) + obtain a'::coname where "a'\(trm{coname2:=(y).P},coname1,P,x)" + by (meson exists_fresh(2) fs_coname1) + with OrR2 show ?case + apply (simp add: fresh_prod trm.inject) + using fresh_fun_simp_OrR2 fin_rest_elims by fastforce +next + case (OrL name1 trm1 name2 trm2 name3) + then show ?case + by simp (metis abs_fresh(1) fin.intros(5) fin_OrL_elim freshn_after_substc) +next + case (ImpR name coname1 trm coname2) + obtain a'::coname where "a'\(trm{coname2:=(y).P},coname1,P,x)" + by (meson exists_fresh(2) fs_coname1) + with ImpR show ?case + apply (simp add: fresh_prod trm.inject) + using fresh_fun_simp_ImpR fin_rest_elims by fastforce +next + case (ImpL coname trm1 name1 trm2 name2) + then show ?case + by simp (metis abs_fresh(1) fin.intros(6) fin_ImpL_elim freshn_after_substc) +qed + + lemma not_fin_subst2: - assumes a: "\(fin M x)" + assumes "\(fin M x)" shows "\(fin (M{y:=.P}) x)" -using a [[simproc del: defined_all]] -apply(nominal_induct M avoiding: x c y P rule: trm.strong_induct) -apply(auto) -apply(erule fin.cases, simp_all add: trm.inject) -apply(erule fin.cases, simp_all add: trm.inject) -apply(erule fin.cases, simp_all add: trm.inject) -apply(erule fin.cases, simp_all add: trm.inject) -apply(subgoal_tac "\a'::name. a'\(trm{y:=.P},P,x)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_NotL) -apply(erule fin.cases, simp_all add: trm.inject) -apply(rule exists_fresh'(1)[OF fs_name1]) -apply(drule fin_NotL_elim) -apply(auto)[1] -apply(drule freshn_after_substn) -apply(simp) -apply(simp add: fin.intros) -apply(erule fin.cases, simp_all add: trm.inject) -apply(subgoal_tac "\a'::name. a'\(trm{y:=.P},P,name1,x)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_AndL1) -apply(erule fin.cases, simp_all add: trm.inject) -apply(rule exists_fresh'(1)[OF fs_name1]) -apply(drule fin_AndL1_elim) -apply(auto simp: abs_fresh)[1] -apply(drule freshn_after_substn) -apply(simp) -apply(subgoal_tac "name2\[name1]. trm") -apply(simp add: fin.intros) -apply(simp add: abs_fresh) -apply(subgoal_tac "\a'::name. a'\(trm{y:=.P},P,name1,x)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_AndL2) -apply(erule fin.cases, simp_all add: trm.inject) -apply(rule exists_fresh'(1)[OF fs_name1]) -apply(drule fin_AndL2_elim) -apply(auto simp: abs_fresh)[1] -apply(drule freshn_after_substn) -apply(simp) -apply(subgoal_tac "name2\[name1].trm") -apply(simp add: fin.intros) -apply(simp add: abs_fresh) -apply(erule fin.cases, simp_all add: trm.inject) -apply(erule fin.cases, simp_all add: trm.inject) -apply(subgoal_tac "\a'::name. a'\(trm1{y:=.P},trm2{y:=.P},name1,name2,P,x)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_OrL) -apply(erule fin.cases, simp_all add: trm.inject) -apply(rule exists_fresh'(1)[OF fs_name1]) -apply(drule fin_OrL_elim) -apply(auto simp: abs_fresh)[1] -apply(drule freshn_after_substn) -apply(simp) -apply(drule freshn_after_substn) -apply(simp) -apply(subgoal_tac "name3\[name1].trm1 \ name3\[name2].trm2") -apply(simp add: fin.intros) -apply(simp add: abs_fresh) -apply(erule fin.cases, simp_all add: trm.inject) -apply(subgoal_tac "\a'::name. a'\(trm1{name2:=.P},trm2{name2:=.P},name1,P,x)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_ImpL) -apply(erule fin.cases, simp_all add: trm.inject) -apply(rule exists_fresh'(1)[OF fs_name1]) -apply(drule fin_ImpL_elim) -apply(auto simp: abs_fresh)[1] -apply(drule freshn_after_substn) -apply(simp) -apply(drule freshn_after_substn) -apply(simp) -apply(subgoal_tac "x\[name1].trm2") -apply(simp add: fin.intros) -apply(simp add: abs_fresh) -done +using assms +proof (nominal_induct M avoiding: x c y P rule: trm.strong_induct) + case (Ax name coname) + then show ?case + using fin_rest_elims(1) substn.simps(1) by presburger +next + case (Cut coname trm1 name trm2) + then show ?case + using fin_rest_elims(1) substc.simps(1) by simp presburger +next + case (NotR name trm coname) + with fin_rest_elims show ?case + by (fastforce simp add: fresh_prod trm.inject) +next + case (NotL coname trm name) + obtain a'::name where "a'\(trm{y:=.P},P,x)" + by (meson exists_fresh(1) fs_name1) + with NotL show ?case + apply (clarsimp simp: fresh_prod fresh_fun_simp_NotL trm.inject) + by (metis fin.intros(2) fin_NotL_elim fin_rest_elims(1) freshn_after_substn) +next + case (AndR coname1 trm1 coname2 trm2 coname3) + obtain a'::name where "a'\(trm1{coname3:=(y).P},trm2{coname3:=(y).P},P,coname1,coname2,coname3,x)" + by (meson exists_fresh(1) fs_name1) + with AndR fin_rest_elims show ?case + by (fastforce simp add: fresh_prod trm.inject) +next + case (AndL1 name1 trm name2) + obtain a'::name where "a'\(trm{y:=.P},P,name1,x)" + by (meson exists_fresh(1) fs_name1) + with AndL1 show ?case + apply (clarsimp simp: fresh_prod fresh_fun_simp_AndL1 trm.inject) + by (metis abs_fresh(1) fin.intros(3) fin_AndL1_elim fin_rest_elims(1) freshn_after_substn) +next + case (AndL2 name1 trm name2) + obtain a'::name where "a'\(trm{y:=.P},P,name1,x)" + by (meson exists_fresh(1) fs_name1) + with AndL2 show ?case + apply (clarsimp simp: fresh_prod fresh_fun_simp_AndL2 trm.inject) + by (metis abs_fresh(1) fin.intros(4) fin_AndL2_elim fin_rest_elims(1) freshn_after_substn) +next + case (OrR1 coname1 trm coname2) + then show ?case + using fin_rest_elims by (fastforce simp: fresh_prod trm.inject) +next + case (OrR2 coname1 trm coname2) + then show ?case + using fin_rest_elims by (fastforce simp: fresh_prod trm.inject) +next + case (OrL name1 trm1 name2 trm2 name3) + obtain a'::name where "a'\(trm1{y:=.P},trm2{y:=.P},name1,name2,P,x)" + by (meson exists_fresh(1) fs_name1) + with OrL show ?case + apply (simp add: fresh_prod trm.inject) + by (metis (full_types) abs_fresh(1) fin.intros(5) fin_OrL_elim fin_rest_elims(1) fresh_fun_simp_OrL freshn_after_substn) +next + case (ImpR name coname1 trm coname2) + with fin_rest_elims show ?case + by (fastforce simp add: fresh_prod trm.inject) +next + case (ImpL coname trm1 name1 trm2 name2) + obtain a'::name where "a'\(trm1{name2:=.P},trm2{name2:=.P},name1,P,x)" + by (meson exists_fresh(1) fs_name1) + with ImpL show ?case + apply (simp add: fresh_prod trm.inject) + by (metis abs_fresh(1) fin.intros(6) fin_ImpL_elim fin_rest_elims(1) fresh_fun_simp_ImpL freshn_after_substn) +qed + lemma fin_subst1: - assumes a: "fin M x" "x\y" "x\P" + assumes "fin M x" "x\y" "x\P" shows "fin (M{y:=.P}) x" -using a -apply(nominal_induct M avoiding: x y c P rule: trm.strong_induct) -apply(auto dest!: fin_elims simp add: subst_fresh abs_fresh) -apply(rule fin.intros, simp add: subst_fresh abs_fresh) -apply(rule fin.intros, simp add: subst_fresh abs_fresh) -apply(rule fin.intros, simp add: subst_fresh abs_fresh) -apply(rule fin.intros, simp add: subst_fresh abs_fresh) -apply(rule fin.intros, simp add: subst_fresh abs_fresh, simp add: subst_fresh abs_fresh) -apply(rule fin.intros, simp add: subst_fresh abs_fresh, simp add: subst_fresh abs_fresh) -apply(rule fin.intros, simp add: subst_fresh abs_fresh, simp add: subst_fresh abs_fresh) -apply(rule fin.intros, simp add: subst_fresh abs_fresh, simp add: subst_fresh abs_fresh) -apply(rule fin.intros, simp add: subst_fresh abs_fresh, simp add: subst_fresh abs_fresh) -done + using assms +proof (nominal_induct M avoiding: x y c P rule: trm.strong_induct) + case (AndL1 name1 trm name2) + then show ?case + apply (clarsimp simp add: subst_fresh abs_fresh dest!: fin_AndL1_elim) + by (metis abs_fresh(1) fin.intros(3) fresh_prod subst_fresh(3)) +next + case (AndL2 name1 trm name2) + then show ?case + apply (clarsimp simp add: subst_fresh abs_fresh dest!: fin_AndL2_elim) + by (metis abs_fresh(1) fin.intros(4) fresh_prod subst_fresh(3)) +next + case (OrR1 coname1 trm coname2) + then show ?case + by (auto simp add: subst_fresh abs_fresh dest!: fin_rest_elims) +next + case (OrR2 coname1 trm coname2) + then show ?case + by (auto simp add: subst_fresh abs_fresh dest!: fin_rest_elims) +next + case (OrL name1 trm1 name2 trm2 name3) + then show ?case + apply (clarsimp simp add: subst_fresh abs_fresh) + by (metis abs_fresh(1) fin.intros(5) fin_OrL_elim fresh_prod subst_fresh(3)) +next + case (ImpR name coname1 trm coname2) + then show ?case + by (auto simp add: subst_fresh abs_fresh dest!: fin_rest_elims) +next + case (ImpL coname trm1 name1 trm2 name2) + then show ?case + apply (clarsimp simp add: subst_fresh abs_fresh) + by (metis abs_fresh(1) fin.intros(6) fin_ImpL_elim fresh_prod subst_fresh(3)) +qed (auto dest!: fin_elims simp add: subst_fresh abs_fresh) + + +thm abs_fresh fresh_prod lemma fin_subst2: - assumes a: "fin M y" "x\y" "y\P" "M\Ax y c" + assumes "fin M y" "x\y" "y\P" "M\Ax y c" shows "fin (M{c:=(x).P}) y" -using a -apply(nominal_induct M avoiding: x y c P rule: trm.strong_induct) -apply(drule fin_elims) -apply(simp add: trm.inject) -apply(rule fin.intros) -apply(drule fin_elims, simp) -apply(drule fin_elims, simp) -apply(drule fin_elims, simp) -apply(rule fin.intros) -apply(rule subst_fresh) -apply(simp) -apply(drule fin_elims, simp) -apply(drule fin_elims, simp) -apply(rule fin.intros) -apply(simp add: abs_fresh fresh_atm) -apply(rule subst_fresh) -apply(auto)[1] -apply(drule fin_elims, simp) -apply(rule fin.intros) -apply(simp add: abs_fresh fresh_atm) -apply(rule subst_fresh) -apply(auto)[1] -apply(drule fin_elims, simp) -apply(drule fin_elims, simp) -apply(drule fin_elims, simp) -apply(rule fin.intros) -apply(simp add: abs_fresh fresh_atm) -apply(rule subst_fresh) -apply(auto)[1] -apply(simp add: abs_fresh fresh_atm) -apply(rule subst_fresh) -apply(auto)[1] -apply(drule fin_elims, simp) -apply(drule fin_elims, simp) -apply(rule fin.intros) -apply(simp add: abs_fresh fresh_atm) -apply(rule subst_fresh) -apply(auto)[1] -apply(simp add: abs_fresh fresh_atm) -apply(rule subst_fresh) -apply(auto)[1] -done +using assms +proof (nominal_induct M avoiding: x y c P rule: trm.strong_induct) + case (Ax name coname) + then show ?case + using fin_Ax_elim by force +next + case (NotL coname trm name) + then show ?case + by simp (metis fin.intros(2) fin_NotL_elim fresh_prod subst_fresh(5)) +next + case (AndL1 name1 trm name2) + then show ?case + by simp (metis abs_fresh(1) fin.intros(3) fin_AndL1_elim fresh_prod subst_fresh(5)) +next + case (AndL2 name1 trm name2) + then show ?case + by simp (metis abs_fresh(1) fin.intros(4) fin_AndL2_elim fresh_prod subst_fresh(5)) +next + case (OrL name1 trm1 name2 trm2 name3) + then show ?case + by simp (metis abs_fresh(1) fin.intros(5) fin_OrL_elim fresh_prod subst_fresh(5)) +next + case (ImpL coname trm1 name1 trm2 name2) + then show ?case + by simp (metis abs_fresh(1) fin.intros(6) fin_ImpL_elim fresh_prod subst_fresh(5)) +qed (use fin_rest_elims in force)+ + lemma fin_substn_nrename: - assumes a: "fin M x" "x\y" "x\P" + assumes "fin M x" "x\y" "x\P" shows "M[x\n>y]{y:=.P} = Cut .P (x).(M{y:=.P})" -using a [[simproc del: defined_all]] -apply(nominal_induct M avoiding: x y c P rule: trm.strong_induct) -apply(drule fin_Ax_elim) -apply(simp) -apply(simp add: trm.inject) -apply(simp add: alpha calc_atm fresh_atm) -apply(simp) -apply(drule fin_rest_elims) -apply(simp) -apply(drule fin_rest_elims) -apply(simp) -apply(drule fin_NotL_elim) -apply(simp) -apply(subgoal_tac "\z::name. z\(trm,y,x,P,trm[x\n>y]{y:=.P})") -apply(erule exE, simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_NotL) -apply(simp add: trm.inject alpha fresh_atm calc_atm abs_fresh) -apply(rule conjI) -apply(simp add: nsubst_eqvt calc_atm) -apply(simp add: perm_fresh_fresh) -apply(simp add: nrename_fresh) -apply(rule subst_fresh) -apply(simp) -apply(rule exists_fresh') -apply(rule fin_supp) -apply(drule fin_rest_elims) -apply(simp) -apply(drule fin_AndL1_elim) -apply(simp) -apply(subgoal_tac "\z::name. z\(name2,name1,P,trm[name2\n>y]{y:=.P},y,P,trm)") -apply(erule exE, simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_AndL1) -apply(simp add: trm.inject alpha fresh_atm calc_atm abs_fresh) -apply(rule conjI) -apply(simp add: nsubst_eqvt calc_atm) -apply(simp add: perm_fresh_fresh) -apply(simp add: nrename_fresh) -apply(rule subst_fresh) -apply(simp) -apply(rule exists_fresh') -apply(rule fin_supp) -apply(drule fin_AndL2_elim) -apply(simp) -apply(subgoal_tac "\z::name. z\(name2,name1,P,trm[name2\n>y]{y:=.P},y,P,trm)") -apply(erule exE, simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_AndL2) -apply(simp add: trm.inject alpha fresh_atm calc_atm abs_fresh) -apply(rule conjI) -apply(simp add: nsubst_eqvt calc_atm) -apply(simp add: perm_fresh_fresh) -apply(simp add: nrename_fresh) -apply(rule subst_fresh) -apply(simp) -apply(rule exists_fresh') -apply(rule fin_supp) -apply(drule fin_rest_elims) -apply(simp) -apply(drule fin_rest_elims) -apply(simp) -apply(drule fin_OrL_elim) -apply(simp add: abs_fresh) -apply(simp add: subst_fresh rename_fresh) -apply(subgoal_tac "\z::name. z\(name3,name2,name1,P,trm1[name3\n>y]{y:=.P},trm2[name3\n>y]{y:=.P},y,P,trm1,trm2)") -apply(erule exE, simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_OrL) -apply(simp add: trm.inject alpha fresh_atm calc_atm abs_fresh) -apply(rule conjI) -apply(simp add: nsubst_eqvt calc_atm) -apply(simp add: perm_fresh_fresh) -apply(simp add: nrename_fresh) -apply(simp add: nsubst_eqvt calc_atm) -apply(simp add: perm_fresh_fresh) -apply(simp add: nrename_fresh) -apply(rule exists_fresh') -apply(rule fin_supp) -apply(drule fin_rest_elims) -apply(simp) -apply(drule fin_ImpL_elim) -apply(simp add: abs_fresh) -apply(simp add: subst_fresh rename_fresh) -apply(subgoal_tac "\z::name. z\(name1,x,P,trm1[x\n>y]{y:=.P},trm2[x\n>y]{y:=.P},y,P,trm1,trm2)") -apply(erule exE, simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_ImpL) -apply(simp add: trm.inject alpha fresh_atm calc_atm abs_fresh) -apply(rule conjI) -apply(simp add: nsubst_eqvt calc_atm) -apply(simp add: perm_fresh_fresh) -apply(simp add: nrename_fresh) -apply(simp add: nsubst_eqvt calc_atm) -apply(simp add: perm_fresh_fresh) -apply(simp add: nrename_fresh) -apply(rule exists_fresh') -apply(rule fin_supp) -done +using assms [[simproc del: defined_all]] +proof (nominal_induct M avoiding: x y c P rule: trm.strong_induct) + case (Ax name coname) + then show ?case + by (metis fin_Ax_elim fresh_atm(1,3) fresh_prod nrename_swap subst_rename(3) substn.simps(1) trm.fresh(1)) +next + case (NotL coname trm name) + obtain z::name where "z \ (trm,y,x,P,trm[x\n>y]{y:=.P})" + by (meson exists_fresh(1) fs_name1) + with NotL show ?case + apply (clarsimp simp add: fresh_prod fresh_fun_simp_NotL trm.inject alpha fresh_atm calc_atm abs_fresh) + by (metis fin_NotL_elim nrename_fresh nrename_swap substn_nrename_comm') +next + case (AndL1 name1 trm name2) + obtain z::name where "z \ (name2,name1,P,trm[name2\n>y]{y:=.P},y,P,trm)" + by (meson exists_fresh(1) fs_name1) + with AndL1 show ?case + using fin_AndL1_elim[OF AndL1.prems(1)] + by simp (metis abs_fresh(1) fresh_atm(1) fresh_fun_simp_AndL1 fresh_prod nrename_fresh subst_fresh(3)) +next + case (AndL2 name1 trm name2) + obtain z::name where "z \ (name2,name1,P,trm[name2\n>y]{y:=.P},y,P,trm)" + by (meson exists_fresh(1) fs_name1) + with AndL2 show ?case + using fin_AndL2_elim[OF AndL2.prems(1)] + by simp (metis abs_fresh(1) fresh_atm(1) fresh_fun_simp_AndL2 fresh_prod nrename_fresh subst_fresh(3)) +next + case (OrL name1 trm1 name2 trm2 name3) + obtain z::name where "z \ (name3,name2,name1,P,trm1[name3\n>y]{y:=.P},trm2[name3\n>y]{y:=.P},y,P,trm1,trm2)" + by (meson exists_fresh(1) fs_name1) + with OrL show ?case + using fin_OrL_elim[OF OrL.prems(1)] + by (auto simp add: abs_fresh fresh_fun_simp_OrL fresh_atm(1) nrename_fresh subst_fresh(3)) +next + case (ImpL coname trm1 name1 trm2 name2) + obtain z::name where "z \ (name1,x,P,trm1[x\n>y]{y:=.P},trm2[x\n>y]{y:=.P},y,P,trm1,trm2)" + by (meson exists_fresh(1) fs_name1) + with ImpL show ?case + using fin_ImpL_elim[OF ImpL.prems(1)] + by (auto simp add: abs_fresh fresh_fun_simp_ImpL fresh_atm(1) nrename_fresh subst_fresh(3)) +qed (use fin_rest_elims in force)+ inductive fic :: "trm \ coname \ bool" @@ -2293,312 +2225,202 @@ (auto simp: calc_atm simp add: fresh_left abs_fresh) lemma not_fic_subst1: - assumes a: "\(fic M a)" + assumes "\(fic M a)" shows "\(fic (M{y:=.P}) a)" -using a -apply(nominal_induct M avoiding: a c y P rule: trm.strong_induct) -apply(auto) -apply(drule fic_elims, simp) -apply(drule fic_elims, simp) -apply(drule fic_elims, simp) -apply(drule fic_elims) -apply(auto)[1] -apply(drule freshc_after_substn) -apply(simp add: fic.intros) -apply(subgoal_tac "\x'::name. x'\(trm{y:=.P},P,a)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_NotL) -apply(drule fic_elims, simp) -apply(rule exists_fresh'(1)[OF fs_name1]) -apply(drule fic_elims, simp) -apply(drule fic_elims) -apply(auto)[1] -apply(simp add: abs_fresh fresh_atm) -apply(drule freshc_after_substn) -apply(drule freshc_after_substn) -apply(simp add: fic.intros abs_fresh) -apply(subgoal_tac "\x'::name. x'\(trm{y:=.P},P,name1,a)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_AndL1) -apply(drule fic_elims, simp) -apply(rule exists_fresh'(1)[OF fs_name1]) -apply(drule fic_elims, simp) -apply(subgoal_tac "\x'::name. x'\(trm{y:=.P},P,name1,a)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_AndL2) -apply(drule fic_elims, simp) -apply(rule exists_fresh'(1)[OF fs_name1]) -apply(drule fic_elims, simp) -apply(drule fic_elims) -apply(auto)[1] -apply(simp add: abs_fresh fresh_atm) -apply(drule freshc_after_substn) -apply(simp add: fic.intros abs_fresh) -apply(drule fic_elims) -apply(auto)[1] -apply(simp add: abs_fresh fresh_atm) -apply(drule freshc_after_substn) -apply(simp add: fic.intros abs_fresh) -apply(subgoal_tac "\x'::name. x'\(trm1{y:=.P},trm2{y:=.P},P,name1,name2,a)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_OrL) -apply(drule fic_elims, simp) -apply(rule exists_fresh'(1)[OF fs_name1]) -apply(drule fic_elims, simp) -apply(drule fic_elims, simp) -apply(simp add: abs_fresh fresh_atm) -apply(drule freshc_after_substn) -apply(simp add: fic.intros abs_fresh) -apply(subgoal_tac "\x'::name. x'\(trm1{name2:=.P},trm2{name2:=.P},P,name1,name2,a)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_ImpL) -apply(drule fic_elims, simp) -apply(rule exists_fresh'(1)[OF fs_name1]) -apply(drule fic_elims, simp) -done +using assms +proof (nominal_induct M avoiding: a c y P rule: trm.strong_induct) + case (Ax name coname) + then show ?case + using fic_rest_elims(1) substn.simps(1) by presburger +next + case (Cut coname trm1 name trm2) + then show ?case + by (metis abs_fresh(2) better_Cut_substn fic_rest_elims(1) fresh_prod) +next + case (NotR name trm coname) + then show ?case + by (metis fic.intros(2) fic_NotR_elim fresh_prod freshc_after_substn substn.simps(3)) +next + case (NotL coname trm name) + obtain x'::name where "x' \ (trm{y:=.P},P,a)" + by (meson exists_fresh(1) fs_name1) + with NotL show ?case + by (simp add: fic.intros fresh_prod) (metis fic_rest_elims(1,2) fresh_fun_simp_NotL) +next + case (AndR coname1 trm1 coname2 trm2 coname3) + then show ?case + by simp (metis abs_fresh(2) fic.intros(3) fic_AndR_elim freshc_after_substn) +next + case (AndL1 name1 trm name2) + obtain x'::name where "x' \ (trm{y:=.P},P,name1,a)" + by (meson exists_fresh(1) fs_name1) + with AndL1 fic_rest_elims show ?case + by (simp add: fic.intros fresh_prod)(metis fresh_fun_simp_AndL1) +next + case (AndL2 name1 trm name2) + obtain x'::name where "x' \ (trm{y:=.P},P,name1,a)" + by (meson exists_fresh(1) fs_name1) + with AndL2 fic_rest_elims show ?case + by (simp add: fic.intros fresh_prod) (metis fresh_fun_simp_AndL2) +next + case (OrR1 coname1 trm coname2) + then show ?case + by simp (metis abs_fresh(2) fic.simps fic_OrR1_elim freshc_after_substn) +next + case (OrR2 coname1 trm coname2) + then show ?case + by simp (metis abs_fresh(2) fic.simps fic_OrR2_elim freshc_after_substn) +next + case (OrL name1 trm1 name2 trm2 name3) + obtain x'::name where "x' \ (trm1{y:=.P},trm2{y:=.P},P,name1,name2,a)" + by (meson exists_fresh(1) fs_name1) + with OrL fic_rest_elims show ?case + by (simp add: fic.intros fresh_prod) (metis fresh_fun_simp_OrL) +next + case (ImpR name coname1 trm coname2) + then show ?case + by simp (metis abs_fresh(2) fic.intros(6) fic_ImpR_elim freshc_after_substn) +next + case (ImpL coname trm1 name1 trm2 name2) + obtain x'::name where "x' \ (trm1{name2:=.P},trm2{name2:=.P},P,name1,name2,a)" + by (meson exists_fresh(1) fs_name1) + with ImpL fic_rest_elims fresh_fun_simp_ImpL show ?case + by (simp add: fresh_prod) fastforce +qed lemma not_fic_subst2: - assumes a: "\(fic M a)" + assumes "\(fic M a)" shows "\(fic (M{c:=(y).P}) a)" -using a -apply(nominal_induct M avoiding: a c y P rule: trm.strong_induct) -apply(auto) -apply(drule fic_elims, simp) -apply(drule fic_elims, simp) -apply(drule fic_elims, simp) -apply(subgoal_tac "\c'::coname. c'\(trm{coname:=(y).P},P,a)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_NotR) -apply(drule fic_elims, simp) -apply(rule exists_fresh'(2)[OF fs_coname1]) -apply(drule fic_elims, simp) -apply(drule freshc_after_substc) -apply(simp) -apply(simp add: fic.intros abs_fresh) -apply(drule fic_elims, simp) -apply(subgoal_tac "\c'::coname. c'\(trm1{coname3:=(y).P},trm2{coname3:=(y).P},P,coname1,coname2,a)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_AndR) -apply(drule fic_elims, simp) -apply(rule exists_fresh'(2)[OF fs_coname1]) -apply(drule fic_elims, simp) -apply(auto)[1] -apply(simp add: abs_fresh fresh_atm) -apply(drule freshc_after_substc) -apply(simp) -apply(drule freshc_after_substc) -apply(simp) -apply(simp add: fic.intros abs_fresh) -apply(drule fic_elims, simp) -apply(drule fic_elims, simp) -apply(subgoal_tac "\c'::coname. c'\(trm{coname2:=(y).P},P,coname1,a)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_OrR1) -apply(drule fic_elims, simp) -apply(rule exists_fresh'(2)[OF fs_coname1]) -apply(drule fic_elims, simp) -apply(simp add: abs_fresh fresh_atm) -apply(drule freshc_after_substc) -apply(simp) -apply(simp add: fic.intros abs_fresh) -apply(subgoal_tac "\c'::coname. c'\(trm{coname2:=(y).P},P,coname1,a)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_OrR2) -apply(drule fic_elims, simp) -apply(rule exists_fresh'(2)[OF fs_coname1]) -apply(drule fic_elims, simp) -apply(simp add: abs_fresh fresh_atm) -apply(drule freshc_after_substc) -apply(simp) -apply(simp add: fic.intros abs_fresh) -apply(drule fic_elims, simp) -apply(subgoal_tac "\c'::coname. c'\(trm{coname2:=(y).P},P,coname1,a)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_ImpR) -apply(drule fic_elims, simp) -apply(rule exists_fresh'(2)[OF fs_coname1]) -apply(drule fic_elims, simp) -apply(simp add: abs_fresh fresh_atm) -apply(drule freshc_after_substc) -apply(simp) -apply(simp add: fic.intros abs_fresh) -apply(drule fic_elims, simp) -done +using assms +proof (nominal_induct M avoiding: a c y P rule: trm.strong_induct) + case (NotR name trm coname) + obtain c'::coname where "c'\(trm{coname:=(y).P},P,a)" + by (meson exists_fresh'(2) fs_coname1) + with NotR fic_rest_elims show ?case + apply (clarsimp simp: fresh_prod fresh_fun_simp_NotR) + by (metis fic.intros(2) fic_NotR_elim freshc_after_substc) +next + case (AndR coname1 trm1 coname2 trm2 coname3) + obtain c'::coname where "c'\(trm1{coname3:=(y).P},trm2{coname3:=(y).P},P,coname1,coname2,a)" + by (meson exists_fresh'(2) fs_coname1) + with AndR fic_rest_elims show ?case + apply (clarsimp simp: fresh_prod fresh_fun_simp_AndR) + by (metis abs_fresh(2) fic.intros(3) fic_AndR_elim freshc_after_substc) +next + case (OrR1 coname1 trm coname2) + obtain c'::coname where "c'\(trm{coname2:=(y).P},P,coname1,a)" + by (meson exists_fresh'(2) fs_coname1) + with OrR1 fic_rest_elims show ?case + apply (clarsimp simp: fresh_prod fresh_fun_simp_OrR1) + by (metis abs_fresh(2) fic.intros(4) fic_OrR1_elim freshc_after_substc) +next + case (OrR2 coname1 trm coname2) + obtain c'::coname where "c'\(trm{coname2:=(y).P},P,coname1,a)" + by (meson exists_fresh'(2) fs_coname1) + with OrR2 fic_rest_elims show ?case + apply (clarsimp simp: fresh_prod fresh_fun_simp_OrR2) + by (metis abs_fresh(2) fic.simps fic_OrR2_elim freshc_after_substc) +next + case (ImpR name coname1 trm coname2) + obtain c'::coname where "c'\(trm{coname2:=(y).P},P,coname1,a)" + by (meson exists_fresh'(2) fs_coname1) + with ImpR fic_rest_elims show ?case + apply (clarsimp simp: fresh_prod fresh_fun_simp_ImpR) + by (metis abs_fresh(2) fic.intros(6) fic_ImpR_elim freshc_after_substc) +qed (use fic_rest_elims in auto) lemma fic_subst1: - assumes a: "fic M a" "a\b" "a\P" + assumes "fic M a" "a\b" "a\P" shows "fic (M{b:=(x).P}) a" -using a -apply(nominal_induct M avoiding: x b a P rule: trm.strong_induct) -apply(drule fic_elims) -apply(simp add: fic.intros) -apply(drule fic_elims, simp) -apply(drule fic_elims, simp) -apply(rule fic.intros) -apply(rule subst_fresh) -apply(simp) -apply(drule fic_elims, simp) -apply(drule fic_elims, simp) -apply(rule fic.intros) -apply(simp add: abs_fresh fresh_atm) -apply(rule subst_fresh) -apply(auto)[1] -apply(simp add: abs_fresh fresh_atm) -apply(rule subst_fresh) -apply(auto)[1] -apply(drule fic_elims, simp) -apply(drule fic_elims, simp) -apply(drule fic_elims, simp) -apply(rule fic.intros) -apply(simp add: abs_fresh fresh_atm) -apply(rule subst_fresh) -apply(auto)[1] -apply(drule fic_elims, simp) -apply(rule fic.intros) -apply(simp add: abs_fresh fresh_atm) -apply(rule subst_fresh) -apply(auto)[1] -apply(drule fic_elims, simp) -apply(drule fic_elims, simp) -apply(rule fic.intros) -apply(simp add: abs_fresh fresh_atm) -apply(rule subst_fresh) -apply(auto)[1] -apply(drule fic_elims, simp) -done +using assms +proof (nominal_induct M avoiding: x b a P rule: trm.strong_induct) + case (Ax name coname) + then show ?case + using fic_Ax_elim by force +next + case (NotR name trm coname) + with fic_NotR_elim show ?case + by (fastforce simp add: fresh_prod subst_fresh) +next + case (AndR coname1 trm1 coname2 trm2 coname3) + with fic_AndR_elim show ?case + by (fastforce simp add: abs_fresh fresh_prod subst_fresh) +next + case (OrR1 coname1 trm coname2) + with fic_OrR1_elim show ?case + by (fastforce simp add: abs_fresh fresh_prod subst_fresh) +next + case (OrR2 coname1 trm coname2) + with fic_OrR2_elim show ?case + by (fastforce simp add: abs_fresh fresh_prod subst_fresh) +next + case (ImpR name coname1 trm coname2) + with fic_ImpR_elim show ?case + by (fastforce simp add: abs_fresh fresh_prod subst_fresh) +qed (use fic_rest_elims in force)+ lemma fic_subst2: - assumes a: "fic M a" "c\a" "a\P" "M\Ax x a" + assumes "fic M a" "c\a" "a\P" "M\Ax x a" shows "fic (M{x:=.P}) a" -using a -apply(nominal_induct M avoiding: x a c P rule: trm.strong_induct) -apply(drule fic_elims) -apply(simp add: trm.inject) -apply(rule fic.intros) -apply(drule fic_elims, simp) -apply(drule fic_elims, simp) -apply(rule fic.intros) -apply(rule subst_fresh) -apply(simp) -apply(drule fic_elims, simp) -apply(drule fic_elims, simp) -apply(rule fic.intros) -apply(simp add: abs_fresh fresh_atm) -apply(rule subst_fresh) -apply(auto)[1] -apply(simp add: abs_fresh fresh_atm) -apply(rule subst_fresh) -apply(auto)[1] -apply(drule fic_elims, simp) -apply(drule fic_elims, simp) -apply(drule fic_elims, simp) -apply(rule fic.intros) -apply(simp add: abs_fresh fresh_atm) -apply(rule subst_fresh) -apply(auto)[1] -apply(drule fic_elims, simp) -apply(rule fic.intros) -apply(simp add: abs_fresh fresh_atm) -apply(rule subst_fresh) -apply(auto)[1] -apply(drule fic_elims, simp) -apply(drule fic_elims, simp) -apply(rule fic.intros) -apply(simp add: abs_fresh fresh_atm) -apply(rule subst_fresh) -apply(auto)[1] -apply(drule fic_elims, simp) -done +using assms +proof (nominal_induct M avoiding: x a c P rule: trm.strong_induct) + case (Ax name coname) + then show ?case + using fic_Ax_elim by force +next + case (NotR name trm coname) + with fic_NotR_elim show ?case + by (fastforce simp add: fresh_prod subst_fresh) +next + case (AndR coname1 trm1 coname2 trm2 coname3) + with fic_AndR_elim show ?case + by simp (metis abs_fresh(2) crename_cfresh crename_fresh fic.intros(3) fresh_atm(2) substn_crename_comm') +next + case (OrR1 coname1 trm coname2) + with fic_OrR1_elim show ?case + by (fastforce simp add: abs_fresh fresh_prod subst_fresh) +next + case (OrR2 coname1 trm coname2) + with fic_OrR2_elim show ?case + by (fastforce simp add: abs_fresh fresh_prod subst_fresh) +next + case (ImpR name coname1 trm coname2) + with fic_ImpR_elim show ?case + by (fastforce simp add: abs_fresh fresh_prod subst_fresh) +qed (use fic_rest_elims in force)+ lemma fic_substc_crename: - assumes a: "fic M a" "a\b" "a\P" + assumes "fic M a" "a\b" "a\P" shows "M[a\c>b]{b:=(y).P} = Cut .(M{b:=(y).P}) (y).P" -using a -apply(nominal_induct M avoiding: a b y P rule: trm.strong_induct) -apply(drule fic_Ax_elim) -apply(simp) -apply(simp add: trm.inject) -apply(simp add: alpha calc_atm fresh_atm trm.inject) -apply(simp) -apply(drule fic_rest_elims) -apply(simp) -apply(drule fic_NotR_elim) -apply(simp) -apply(generate_fresh "coname") -apply(fresh_fun_simp) -apply(simp add: trm.inject alpha fresh_atm fresh_prod fresh_atm calc_atm abs_fresh) -apply(rule conjI) -apply(simp add: csubst_eqvt calc_atm) -apply(simp add: perm_fresh_fresh) -apply(simp add: crename_fresh) -apply(rule subst_fresh) -apply(simp) -apply(drule fic_rest_elims) -apply(simp) -apply(drule fic_AndR_elim) -apply(simp add: abs_fresh fresh_atm subst_fresh rename_fresh) -apply(generate_fresh "coname") -apply(fresh_fun_simp) -apply(simp add: trm.inject alpha fresh_atm calc_atm abs_fresh fresh_prod) -apply(rule conjI) -apply(simp add: csubst_eqvt calc_atm) -apply(simp add: perm_fresh_fresh) -apply(simp add: csubst_eqvt calc_atm) -apply(simp add: perm_fresh_fresh) -apply(simp add: subst_fresh) -apply(drule fic_rest_elims) -apply(simp) -apply(drule fic_rest_elims) -apply(simp) -apply(drule fic_OrR1_elim) -apply(simp) -apply(generate_fresh "coname") -apply(fresh_fun_simp) -apply(simp add: trm.inject alpha fresh_atm calc_atm abs_fresh fresh_prod) -apply(simp add: csubst_eqvt calc_atm) -apply(simp add: perm_fresh_fresh) -apply(simp add: subst_fresh rename_fresh) -apply(drule fic_OrR2_elim) -apply(simp add: abs_fresh fresh_atm) -apply(generate_fresh "coname") -apply(fresh_fun_simp) -apply(simp add: trm.inject alpha fresh_atm calc_atm abs_fresh fresh_prod) -apply(simp add: csubst_eqvt calc_atm) -apply(simp add: perm_fresh_fresh) -apply(simp add: subst_fresh rename_fresh) -apply(drule fic_rest_elims) -apply(simp) -apply(drule fic_ImpR_elim) -apply(simp add: abs_fresh fresh_atm) -apply(generate_fresh "coname") -apply(fresh_fun_simp) -apply(simp add: trm.inject alpha fresh_atm calc_atm abs_fresh fresh_prod) -apply(simp add: csubst_eqvt calc_atm) -apply(simp add: perm_fresh_fresh) -apply(simp add: subst_fresh rename_fresh) -apply(drule fic_rest_elims) -apply(simp) -done +using assms +proof (nominal_induct M avoiding: a b y P rule: trm.strong_induct) + case (Ax name coname) + with fic_Ax_elim show ?case + by(force simp add: trm.inject alpha(2) fresh_atm(2,4) swap_simps(4,8)) +next + case (Cut coname trm1 name trm2) + with fic_rest_elims show ?case by force +next + case (NotR name trm coname) + with fic_NotR_elim[OF NotR.prems(1)] show ?case + by (simp add: trm.inject crename_fresh fresh_fun_simp_NotR subst_fresh(6)) +next + case (AndR coname1 trm1 coname2 trm2 coname3) + with AndR fic_AndR_elim[OF AndR.prems(1)] show ?case + by (simp add: abs_fresh rename_fresh fresh_fun_simp_AndR fresh_atm(2) subst_fresh(6)) +next + case (OrR1 coname1 trm coname2) + with fic_OrR1_elim[OF OrR1.prems(1)] show ?case + by (simp add: abs_fresh rename_fresh fresh_fun_simp_OrR1 fresh_atm(2) subst_fresh(6)) +next + case (OrR2 coname1 trm coname2) + with fic_OrR2_elim[OF OrR2.prems(1)] show ?case + by (simp add: abs_fresh rename_fresh fresh_fun_simp_OrR2 fresh_atm(2) subst_fresh(6)) +next + case (ImpR name coname1 trm coname2) + with fic_ImpR_elim[OF ImpR.prems(1)] crename_fresh show ?case + by (force simp add: abs_fresh fresh_fun_simp_ImpR fresh_atm(2) subst_fresh(6)) +qed (use fic_rest_elims in force)+ inductive l_redu :: "trm \ trm \ bool" ("_ \\<^sub>l _" [100,100] 100) @@ -2626,37 +2448,42 @@ and pi2::"coname prm" shows "(pi1\M) \\<^sub>l (pi1\M') \ M \\<^sub>l M'" and "(pi2\M) \\<^sub>l (pi2\M') \ M \\<^sub>l M'" -apply - -apply(drule_tac pi="rev pi1" in l_redu.eqvt(1)) -apply(perm_simp) -apply(drule_tac pi="rev pi2" in l_redu.eqvt(2)) -apply(perm_simp) -done + using l_redu.eqvt perm_pi_simp by metis+ nominal_inductive l_redu - apply(simp_all add: abs_fresh fresh_atm rename_fresh fresh_prod abs_supp fin_supp) - apply(force)+ - done - -lemma fresh_l_redu: - fixes x::"name" - and a::"coname" - shows "M \\<^sub>l M' \ x\M \ x\M'" - and "M \\<^sub>l M' \ a\M \ a\M'" -apply - -apply(induct rule: l_redu.induct) -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: abs_fresh rename_fresh) -apply(case_tac "aa=a") -apply(simp add: rename_fresh) -apply(simp add: rename_fresh fresh_atm) -apply(simp add: fresh_prod abs_fresh abs_supp fin_supp)+ -done + by (force simp add: abs_fresh fresh_atm rename_fresh fresh_prod abs_supp fin_supp)+ + + +lemma fresh_l_redu_x: + fixes z::"name" + shows "M \\<^sub>l M' \ z\M \ z\M'" +proof (induct rule: l_redu.induct) + case (LAxL a M x y) + then show ?case + by (metis abs_fresh(1,5) nrename_nfresh nrename_rename trm.fresh(1,3)) +next + case (LImp z N y P x M b a c) + then show ?case + apply (simp add: abs_fresh fresh_prod) + by (metis abs_fresh(3,5) abs_supp(5) fs_name1) +qed (auto simp add: abs_fresh fresh_prod crename_nfresh) + +lemma fresh_l_redu_a: + fixes c::"coname" + shows "M \\<^sub>l M' \ c\M \ c\M'" +proof (induct rule: l_redu.induct) + case (LAxR x M a b) + then show ?case + apply (simp add: abs_fresh) + by (metis crename_cfresh crename_rename) +next + case (LAxL a M x y) + with nrename_cfresh show ?case + by (simp add: abs_fresh) +qed (auto simp add: abs_fresh fresh_prod crename_nfresh) + + +lemmas fresh_l_redu = fresh_l_redu_x fresh_l_redu_a lemma better_LAxR_intro[intro]: shows "fic M a \ Cut .M (x).(Ax x b) \\<^sub>l M[a\c>b]" @@ -2852,13 +2679,13 @@ \ Cut .(ImpR (x)..M b) (z).(ImpL .N (y).P z) \\<^sub>l Cut .(Cut .N (x).M) (y).P" proof - assume fs: "z\(N,[y].P)" "b\[a].M" "a\N" - obtain y'::"name" where f1: "y'\(y,M,N,P,z,x)" by (rule exists_fresh(1), rule fin_supp, blast) - obtain x'::"name" where f2: "x'\(y,M,N,P,z,x,y')" by (rule exists_fresh(1), rule fin_supp, blast) - obtain z'::"name" where f3: "z'\(y,M,N,P,z,x,y',x')" by (rule exists_fresh(1), rule fin_supp, blast) - obtain a'::"coname" where f4: "a'\(a,N,P,M,b)" by (rule exists_fresh(2), rule fin_supp, blast) - obtain b'::"coname" where f5: "b'\(a,N,P,M,b,c,a')" by (rule exists_fresh(2),rule fin_supp, blast) - obtain c'::"coname" where f6: "c'\(a,N,P,M,b,c,a',b')" by (rule exists_fresh(2),rule fin_supp, blast) - have " Cut .(ImpR (x)..M b) (z).(ImpL .N (y).P z) + obtain y'::"name" and x'::"name" and z'::"name" + where f1: "y'\(y,M,N,P,z,x)" and f2: "x'\(y,M,N,P,z,x,y')" and f3: "z'\(y,M,N,P,z,x,y',x')" + by (meson exists_fresh(1) fs_name1) + obtain a'::"coname" and b'::"coname" and c'::"coname" + where f4: "a'\(a,N,P,M,b)" and f5: "b'\(a,N,P,M,b,c,a')" and f6: "c'\(a,N,P,M,b,c,a',b')" + by (meson exists_fresh(2) fs_coname1) + have "Cut .(ImpR (x)..M b) (z).(ImpL .N (y).P z) = Cut .(ImpR (x)..M b') (z').(ImpL .N (y).P z')" using f1 f2 f3 f4 f5 fs apply(rule_tac sym) @@ -2869,16 +2696,7 @@ (z').(ImpL .([(c',c)]\N) (y').([(y',y)]\P) z')" using f1 f2 f3 f4 f5 f6 fs apply(rule_tac sym) - apply(simp add: trm.inject) - apply(simp add: alpha) - apply(rule conjI) - apply(simp add: trm.inject) - apply(simp add: alpha fresh_prod fresh_atm abs_perm calc_atm fresh_left abs_fresh) - apply(simp add: trm.inject) - apply(simp add: alpha) - apply(rule conjI) - apply(simp add: alpha fresh_prod fresh_atm abs_perm calc_atm fresh_left abs_fresh) - apply(simp add: alpha fresh_prod fresh_atm abs_perm calc_atm fresh_left abs_fresh) + apply(simp add: trm.inject alpha fresh_prod fresh_atm abs_perm calc_atm fresh_left abs_fresh) done also have "\ \\<^sub>l Cut .(Cut .([(c',c)]\N) (x').([(a',a)]\[(x',x)]\M)) (y').([(y',y)]\P)" using f1 f2 f3 f4 f5 f6 fs @@ -2889,67 +2707,41 @@ also have "\ = Cut .(Cut .N (x).M) (y).P" using f1 f2 f3 f4 f5 f6 fs apply(simp add: trm.inject) - apply(rule conjI) - apply(simp add: alpha) - apply(rule disjI2) - apply(simp add: trm.inject) - apply(rule conjI) - apply(simp add: fresh_prod fresh_atm) - apply(rule conjI) - apply(perm_simp add: calc_atm) - apply(auto simp: fresh_prod fresh_atm)[1] - apply(perm_simp add: alpha) - apply(perm_simp add: alpha) - apply(perm_simp add: alpha) - apply(rule conjI) - apply(perm_simp add: calc_atm) - apply(rule_tac pi="[(a',a)]" in pt_bij4[OF pt_coname_inst, OF at_coname_inst]) - apply(perm_simp add: abs_perm calc_atm) apply(perm_simp add: alpha fresh_prod fresh_atm) - apply(simp add: abs_fresh) - apply(perm_simp add: alpha fresh_prod fresh_atm) + apply(simp add: trm.inject abs_fresh alpha(1) fresh_perm_app(4) perm_compose(4) perm_dj(4)) + apply (metis alpha'(2) crename_fresh crename_swap swap_simps(2,4,6)) done finally show ?thesis by simp qed lemma alpha_coname: fixes M::"trm" - and a::"coname" - assumes a: "[a].M = [b].N" "c\(a,b,M,N)" + and a::"coname" + assumes "[a].M = [b].N" "c\(a,b,M,N)" shows "M = [(a,c)]\[(b,c)]\N" -using a -apply(auto simp: alpha_fresh fresh_prod fresh_atm) -apply(drule sym) -apply(perm_simp) -done + by (metis alpha_fresh'(2) assms fresh_atm(2) fresh_prod) lemma alpha_name: fixes M::"trm" - and x::"name" - assumes a: "[x].M = [y].N" "z\(x,y,M,N)" + and x::"name" + assumes "[x].M = [y].N" "z\(x,y,M,N)" shows "M = [(x,z)]\[(y,z)]\N" -using a -apply(auto simp: alpha_fresh fresh_prod fresh_atm) -apply(drule sym) -apply(perm_simp) -done + by (metis alpha_fresh'(1) assms fresh_atm(1) fresh_prod) lemma alpha_name_coname: fixes M::"trm" and x::"name" and a::"coname" - assumes a: "[x].[b].M = [y].[c].N" "z\(x,y,M,N)" "a\(b,c,M,N)" +assumes "[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: alpha_fresh fresh_prod fresh_atm - abs_supp fin_supp abs_fresh abs_perm fresh_left calc_atm) -apply(drule sym) -apply(simp) -apply(perm_simp) -done + using assms + apply(clarsimp simp: alpha_fresh fresh_prod fresh_atm + abs_supp fin_supp abs_fresh abs_perm fresh_left calc_atm) + by (metis perm_swap(1,2)) + lemma Cut_l_redu_elim: - assumes a: "Cut .M (x).N \\<^sub>l R" + assumes "Cut .M (x).N \\<^sub>l R" shows "(\b. R = M[a\c>b]) \ (\y. R = N[x\n>y]) \ (\y M' b N'. M = NotR (y).M' a \ N = NotL .N' x \ R = Cut .N' (y).M' \ fic M a \ fin N x) \ (\b M1 c M2 y N'. M = AndR .M1 .M2 a \ N = AndL1 (y).N' x \ R = Cut .M1 (y).N' @@ -2962,344 +2754,294 @@ \ fic M a \ fin N x) \ (\z b M' c N1 y N2. M = ImpR (z)..M' a \ N = ImpL .N1 (y).N2 x \ R = Cut .(Cut .N1 (z).M') (y).N2 \ b\(c,N1) \ fic M a \ fin N x)" -using a -apply(erule_tac l_redu.cases) -apply(rule disjI1) -(* ax case *) -apply(simp add: trm.inject) -apply(rule_tac x="b" in exI) -apply(erule conjE) -apply(simp add: alpha) -apply(erule disjE) -apply(simp) -apply(simp) -apply(simp add: rename_fresh) -apply(rule disjI2) -apply(rule disjI1) -(* ax case *) -apply(simp add: trm.inject) -apply(rule_tac x="y" in exI) -apply(erule conjE) -apply(thin_tac "[a].M = [aa].Ax y aa") -apply(simp add: alpha) -apply(erule disjE) -apply(simp) -apply(simp) -apply(simp add: rename_fresh) -apply(rule disjI2) -apply(rule disjI2) -apply(rule disjI1) -(* not case *) -apply(simp add: trm.inject) -apply(erule conjE)+ -apply(generate_fresh "coname") -apply(simp add: abs_fresh fresh_prod fresh_atm) -apply(auto)[1] -apply(drule_tac c="c" in alpha_coname) -apply(simp add: fresh_prod fresh_atm abs_fresh) -apply(simp add: calc_atm) -apply(rule exI)+ -apply(rule conjI) -apply(rule refl) -apply(generate_fresh "name") -apply(simp add: calc_atm abs_fresh fresh_prod fresh_atm fresh_left) -apply(auto)[1] -apply(drule_tac z="ca" in alpha_name) -apply(simp add: fresh_prod fresh_atm abs_fresh) -apply(simp add: calc_atm) -apply(rule exI)+ -apply(rule conjI) -apply(rule refl) -apply(auto simp: calc_atm abs_fresh fresh_left)[1] -apply(case_tac "y=x") -apply(perm_simp) -apply(perm_simp) -apply(case_tac "aa=a") -apply(perm_simp) -apply(perm_simp) -(* and1 case *) -apply(rule disjI2) -apply(rule disjI2) -apply(rule disjI2) -apply(rule disjI1) -apply(simp add: trm.inject) -apply(erule conjE)+ -apply(generate_fresh "coname") -apply(simp add: abs_fresh fresh_prod fresh_atm) -apply(auto)[1] -apply(drule_tac c="c" in alpha_coname) -apply(simp add: fresh_prod fresh_atm abs_fresh) -apply(simp) -apply(rule exI)+ -apply(rule conjI) -apply(rule exI)+ -apply(rule_tac s="a" and t="[(a,c)]\[(b,c)]\b" in subst) -apply(simp add: calc_atm) -apply(rule refl) -apply(generate_fresh "name") -apply(simp add: abs_fresh fresh_prod fresh_atm) -apply(auto)[1] -apply(drule_tac z="ca" in alpha_name) -apply(simp add: fresh_prod fresh_atm abs_fresh) -apply(simp) -apply(rule exI)+ -apply(rule conjI) -apply(rule_tac s="x" and t="[(x,ca)]\[(y,ca)]\y" in subst) -apply(simp add: calc_atm) -apply(rule refl) -apply(auto simp: fresh_left calc_atm abs_fresh split: if_splits)[1] -apply(generate_fresh "name") -apply(simp add: abs_fresh fresh_prod fresh_atm) -apply(auto)[1] -apply(drule_tac z="cb" in alpha_name) -apply(simp add: fresh_prod fresh_atm abs_fresh) -apply(simp) -apply(rule exI)+ -apply(rule conjI) -apply(rule_tac s="x" and t="[(x,cb)]\[(y,cb)]\y" in subst) -apply(simp add: calc_atm) -apply(rule refl) -apply(auto simp: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] -apply(perm_simp)+ -apply(generate_fresh "name") -apply(simp add: abs_fresh fresh_prod fresh_atm) -apply(auto)[1] -apply(drule_tac z="cb" in alpha_name) -apply(simp add: fresh_prod fresh_atm abs_fresh) -apply(simp) -apply(rule exI)+ -apply(rule conjI) -apply(rule_tac s="x" and t="[(x,cb)]\[(y,cb)]\y" in subst) -apply(simp add: calc_atm) -apply(rule refl) -apply(auto simp: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] -apply(perm_simp)+ -apply(generate_fresh "name") -apply(simp add: abs_fresh fresh_prod fresh_atm) -apply(auto)[1] -apply(drule_tac z="cb" in alpha_name) -apply(simp add: fresh_prod fresh_atm abs_fresh) -apply(simp) -apply(rule exI)+ -apply(rule conjI) -apply(rule_tac s="x" and t="[(x,cb)]\[(y,cb)]\y" in subst) -apply(simp add: calc_atm) -apply(rule refl) -apply(auto simp: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] -apply(perm_simp)+ -(* and2 case *) -apply(rule disjI2) -apply(rule disjI2) -apply(rule disjI2) -apply(rule disjI2) -apply(rule disjI1) -apply(simp add: trm.inject) -apply(erule conjE)+ -apply(generate_fresh "coname") -apply(simp add: abs_fresh fresh_prod fresh_atm) -apply(auto)[1] -apply(drule_tac c="c" in alpha_coname) -apply(simp add: fresh_prod fresh_atm abs_fresh) -apply(simp) -apply(rule exI)+ -apply(rule conjI) -apply(rule_tac s="a" and t="[(a,c)]\[(b,c)]\b" in subst) -apply(simp add: calc_atm) -apply(rule refl) -apply(generate_fresh "name") -apply(simp add: abs_fresh fresh_prod fresh_atm) -apply(auto)[1] -apply(drule_tac z="ca" in alpha_name) -apply(simp add: fresh_prod fresh_atm abs_fresh) -apply(simp) -apply(rule exI)+ -apply(rule conjI) -apply(rule_tac s="x" and t="[(x,ca)]\[(y,ca)]\y" in subst) -apply(simp add: calc_atm) -apply(rule refl) -apply(auto simp: fresh_left calc_atm abs_fresh split: if_splits)[1] -apply(generate_fresh "name") -apply(simp add: abs_fresh fresh_prod fresh_atm) -apply(auto)[1] -apply(drule_tac z="cb" in alpha_name) -apply(simp add: fresh_prod fresh_atm abs_fresh) -apply(simp) -apply(rule exI)+ -apply(rule conjI) -apply(rule_tac s="x" and t="[(x,cb)]\[(y,cb)]\y" in subst) -apply(simp add: calc_atm) -apply(rule refl) -apply(auto simp: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] -apply(perm_simp)+ -apply(generate_fresh "name") -apply(simp add: abs_fresh fresh_prod fresh_atm) -apply(auto)[1] -apply(drule_tac z="cb" in alpha_name) -apply(simp add: fresh_prod fresh_atm abs_fresh) -apply(simp) -apply(rule exI)+ -apply(rule conjI) -apply(rule_tac s="x" and t="[(x,cb)]\[(y,cb)]\y" in subst) -apply(simp add: calc_atm) -apply(rule refl) -apply(auto simp: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] -apply(perm_simp)+ -apply(generate_fresh "name") -apply(simp add: abs_fresh fresh_prod fresh_atm) -apply(auto)[1] -apply(drule_tac z="cb" in alpha_name) -apply(simp add: fresh_prod fresh_atm abs_fresh) -apply(simp) -apply(rule exI)+ -apply(rule conjI) -apply(rule_tac s="x" and t="[(x,cb)]\[(y,cb)]\y" in subst) -apply(simp add: calc_atm) -apply(rule refl) -apply(auto simp: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] -apply(perm_simp)+ -(* or1 case *) -apply(rule disjI2) -apply(rule disjI2) -apply(rule disjI2) -apply(rule disjI2) -apply(rule disjI2) -apply(rule disjI1) -apply(simp add: trm.inject) -apply(erule conjE)+ -apply(generate_fresh "coname") -apply(simp add: abs_fresh fresh_prod fresh_atm) -apply(auto)[1] -apply(drule_tac c="c" in alpha_coname) -apply(simp add: fresh_prod fresh_atm abs_fresh) -apply(simp) -apply(rule exI)+ -apply(rule conjI) -apply(rule_tac s="a" and t="[(a,c)]\[(b,c)]\b" in subst) -apply(simp add: calc_atm) -apply(rule refl) -apply(generate_fresh "name") -apply(simp add: abs_fresh fresh_prod fresh_atm) -apply(auto)[1] -apply(drule_tac z="ca" in alpha_name) -apply(simp add: fresh_prod fresh_atm abs_fresh) -apply(simp) -apply(rule exI)+ -apply(rule conjI) -apply(rule exI)+ -apply(rule_tac s="x" and t="[(x,ca)]\[(y,ca)]\y" in subst) -apply(simp add: calc_atm) -apply(rule refl) -apply(auto simp: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] -apply(perm_simp)+ -apply(generate_fresh "name") -apply(simp add: abs_fresh fresh_prod fresh_atm) -apply(auto)[1] -apply(drule_tac z="cb" in alpha_name) -apply(simp add: fresh_prod fresh_atm abs_fresh) -apply(simp) -apply(rule exI)+ -apply(rule conjI) -apply(rule exI)+ -apply(rule_tac s="x" and t="[(x,cb)]\[(y,cb)]\y" in subst) -apply(simp add: calc_atm) -apply(rule refl) -apply(auto simp: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] -apply(perm_simp)+ -(* or2 case *) -apply(rule disjI2) -apply(rule disjI2) -apply(rule disjI2) -apply(rule disjI2) -apply(rule disjI2) -apply(rule disjI2) -apply(rule disjI1) -apply(simp add: trm.inject) -apply(erule conjE)+ -apply(generate_fresh "coname") -apply(simp add: abs_fresh fresh_prod fresh_atm) -apply(auto)[1] -apply(drule_tac c="c" in alpha_coname) -apply(simp add: fresh_prod fresh_atm abs_fresh) -apply(simp) -apply(rule exI)+ -apply(rule conjI) -apply(rule_tac s="a" and t="[(a,c)]\[(b,c)]\b" in subst) -apply(simp add: calc_atm) -apply(rule refl) -apply(generate_fresh "name") -apply(simp add: abs_fresh fresh_prod fresh_atm) -apply(auto)[1] -apply(drule_tac z="ca" in alpha_name) -apply(simp add: fresh_prod fresh_atm abs_fresh) -apply(simp) -apply(rule exI)+ -apply(rule conjI) -apply(rule_tac s="x" and t="[(x,ca)]\[(y,ca)]\y" in subst) -apply(simp add: calc_atm) -apply(rule refl) -apply(auto simp: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] -apply(perm_simp)+ -apply(generate_fresh "name") -apply(simp add: abs_fresh fresh_prod fresh_atm) -apply(auto)[1] -apply(drule_tac z="cb" in alpha_name) -apply(simp add: fresh_prod fresh_atm abs_fresh) -apply(simp) -apply(rule exI)+ -apply(rule conjI) -apply(rule_tac s="x" and t="[(x,cb)]\[(y,cb)]\y" in subst) -apply(simp add: calc_atm) -apply(rule refl) -apply(auto simp: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] -apply(perm_simp)+ -(* imp-case *) -apply(rule disjI2) -apply(rule disjI2) -apply(rule disjI2) -apply(rule disjI2) -apply(rule disjI2) -apply(rule disjI2) -apply(rule disjI2) -apply(simp add: trm.inject) -apply(erule conjE)+ -apply(generate_fresh "coname") -apply(simp add: abs_fresh fresh_prod fresh_atm) -apply(auto)[1] -apply(drule_tac c="ca" in alpha_coname) -apply(simp add: fresh_prod fresh_atm abs_fresh) -apply(simp) -apply(rule exI)+ -apply(rule conjI) -apply(rule_tac s="a" and t="[(a,ca)]\[(b,ca)]\b" in subst) -apply(simp add: calc_atm) -apply(rule refl) -apply(generate_fresh "name") -apply(simp add: abs_fresh fresh_prod fresh_atm) -apply(auto)[1] -apply(drule_tac z="cb" in alpha_name) -apply(simp add: fresh_prod fresh_atm abs_fresh) -apply(simp) -apply(rule exI)+ -apply(rule conjI) -apply(rule_tac s="x" and t="[(x,cb)]\[(z,cb)]\z" in subst) -apply(simp add: calc_atm) -apply(rule refl) -apply(auto simp: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] -apply(perm_simp)+ -apply(generate_fresh "name") -apply(simp add: abs_fresh fresh_prod fresh_atm) -apply(auto)[1] -apply(drule_tac z="cc" in alpha_name) -apply(simp add: fresh_prod fresh_atm abs_fresh) -apply(simp) -apply(rule exI)+ -apply(rule conjI) -apply(rule_tac s="x" and t="[(x,cc)]\[(z,cc)]\z" in subst) -apply(simp add: calc_atm) -apply(rule refl) -apply(auto simp: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] -apply(perm_simp)+ -done +using assms +proof (cases rule: l_redu.cases) + case (LAxR x M a b) + then show ?thesis + apply(simp add: trm.inject) + by (metis alpha(2) crename_rename) +next + case (LAxL a M x y) + then show ?thesis + apply(simp add: trm.inject) + by (metis alpha(1) nrename_rename) +next + case (LNot y M N x a b) + then show ?thesis + apply(simp add: trm.inject) + apply(rule disjI2) + apply(rule disjI2) + apply(rule disjI1) + apply(erule conjE)+ + apply(generate_fresh "coname") + apply(simp add: abs_fresh fresh_prod fresh_atm) + apply(auto)[1] + apply(drule_tac c="c" in alpha_coname) + apply(simp add: fresh_prod fresh_atm abs_fresh) + apply(simp add: calc_atm) + apply(rule exI)+ + apply(rule conjI) + apply(rule refl) + apply(generate_fresh "name") + apply(simp add: calc_atm abs_fresh fresh_prod fresh_atm fresh_left) + apply(auto)[1] + apply(drule_tac z="ca" in alpha_name) + apply(simp add: fresh_prod fresh_atm abs_fresh) + apply(simp add: calc_atm) + apply(rule exI)+ + apply(rule conjI) + apply(rule refl) + apply(auto simp: calc_atm abs_fresh fresh_left)[1] + using nrename_fresh nrename_swap apply presburger + using crename_fresh crename_swap by presburger +next + case (LAnd1 b a1 M1 a2 M2 N y x) + then show ?thesis + apply - + apply(rule disjI2) + apply(rule disjI2) + apply(rule disjI2) + apply(rule disjI1) + apply(clarsimp simp add: trm.inject) + apply(generate_fresh "coname") + apply(clarsimp simp add: abs_fresh fresh_prod fresh_atm) + apply(drule_tac c="c" in alpha_coname) + apply(simp add: fresh_prod fresh_atm abs_fresh) + apply(simp) + apply(rule exI)+ + apply(rule conjI) + apply(rule exI)+ + apply(rule_tac s="a" and t="[(a,c)]\[(b,c)]\b" in subst) + apply(simp add: calc_atm) + apply(rule refl) + apply(generate_fresh "name") + apply(simp add: abs_fresh fresh_prod fresh_atm) + apply(auto)[1] + apply(drule_tac z="ca" in alpha_name) + apply(simp add: fresh_prod fresh_atm abs_fresh) + apply(simp) + apply (metis swap_simps(6)) + apply(generate_fresh "name") + apply(simp add: abs_fresh fresh_prod fresh_atm) + apply(auto)[1] + apply(drule_tac z="cb" in alpha_name) + apply(simp add: fresh_prod fresh_atm abs_fresh) + apply(simp) + apply(perm_simp) + apply (smt (verit, del_insts) abs_fresh(1,2) abs_perm(1,2) fic.intros(3) fin.intros(3) fresh_bij(1) perm_fresh_fresh(1,2) swap_simps(1,6)) + apply(perm_simp)+ + apply(generate_fresh "name") + apply(simp add: abs_fresh fresh_prod fresh_atm) + apply(auto)[1] + apply(drule_tac z="cb" in alpha_name) + apply(simp add: fresh_prod fresh_atm abs_fresh) + apply(perm_simp) + apply (smt (verit) abs_fresh(1,2) abs_perm(1,2) fic.intros(3) fin.intros(3) perm_fresh_fresh(1,2) perm_swap(3) swap_simps(1,6)) + apply(perm_simp)+ + apply(generate_fresh "name") + apply(simp add: abs_fresh fresh_prod fresh_atm) + apply(auto)[1] + apply(drule_tac z="cb" in alpha_name) + apply(simp add: fresh_prod fresh_atm abs_fresh cong: conj_cong) + apply(perm_simp) + by (smt (verit, best) abs_fresh(1,2) abs_perm(1) alpha(2) fic.intros(3) fin.intros(3) fresh_bij(1,2) perm_fresh_fresh(1,2) swap_simps(2,3)) +next + case (LAnd2 b a1 M1 a2 M2 N y x) + then show ?thesis + apply - + apply(rule disjI2) + apply(rule disjI2) + apply(rule disjI2) + apply(rule disjI2) + apply(rule disjI1) + apply(clarsimp simp add: trm.inject) + apply(generate_fresh "coname") + apply(clarsimp simp add: abs_fresh fresh_prod fresh_atm) + apply(drule_tac c="c" in alpha_coname) + apply(simp add: fresh_prod fresh_atm abs_fresh) + apply(simp) + apply(rule exI)+ + apply(rule conjI) + apply(rule_tac s="a" and t="[(a,c)]\[(b,c)]\b" in subst) + apply(simp add: calc_atm) + apply(rule refl) + apply(generate_fresh "name") + apply(auto simp add: abs_fresh fresh_prod fresh_atm) + apply(drule_tac z="ca" in alpha_name) + apply(simp add: fresh_prod fresh_atm abs_fresh) + apply (metis swap_simps(6)) + apply(generate_fresh "name") + apply(simp add: abs_fresh fresh_prod fresh_atm) + apply(auto)[1] + apply(drule_tac z="cb" in alpha_name) + apply(simp add: fresh_prod fresh_atm abs_fresh) + apply(perm_simp) + apply (smt (verit, ccfv_threshold) abs_fresh(1,2) abs_perm(1) fic.intros(3) fin.intros(4) fresh_bij(1,2) perm_fresh_fresh(1,2) swap_simps(1,4,6)) + apply(generate_fresh "name") + apply(simp add: abs_fresh fresh_prod fresh_atm) + apply(auto)[1] + apply(drule_tac z="cb" in alpha_name) + apply(simp add: fresh_prod fresh_atm abs_fresh) + apply(simp) + apply(perm_simp) + apply (smt (verit) abs_fresh(1,2) abs_perm(1,2) fic.intros(3) fin.intros(4) fresh_bij(1) perm_fresh_fresh(1,2) swap_simps(1,6)) + apply(generate_fresh "name") + apply(clarsimp simp add: abs_fresh fresh_prod fresh_atm) + apply(drule_tac z="cb" in alpha_name) + apply(simp add: fresh_prod fresh_atm abs_fresh) + apply(perm_simp) + by (smt (verit, ccfv_SIG) abs_fresh(1,2) abs_perm(1) alpha(2) fic.intros(3) fin.intros(4) fresh_bij(1,2) perm_fresh_fresh(1,2) swap_simps(2,3)) +next + case (LOr1 b a M N1 N2 y x1 x2) + then show ?thesis + apply - + apply(rule disjI2) + apply(rule disjI2) + apply(rule disjI2) + apply(rule disjI2) + apply(rule disjI2) + apply(rule disjI1) + apply(clarsimp simp add: trm.inject) + apply(generate_fresh "coname") + apply(simp add: abs_fresh fresh_prod fresh_atm) + apply(auto)[1] + apply(drule_tac c="c" in alpha_coname) + apply(simp add: fresh_prod fresh_atm abs_fresh) + apply(simp) + apply(perm_simp) + apply(rule exI)+ + apply(rule conjI) + apply(rule_tac s="a" and t="[(a,c)]\[(b,c)]\b" in subst) + apply(simp add: calc_atm) + apply(rule refl) + apply(generate_fresh "name") + apply(simp add: abs_fresh fresh_prod fresh_atm) + apply(auto)[1] + apply(drule_tac z="ca" in alpha_name) + apply(simp add: fresh_prod fresh_atm abs_fresh) + apply(simp) + apply(rule exI)+ + apply(rule conjI) + apply(rule exI)+ + apply(rule_tac s="x" and t="[(x,ca)]\[(y,ca)]\y" in subst) + apply(simp add: calc_atm) + apply(rule refl) + apply(auto simp: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] + apply(perm_simp)+ + apply(generate_fresh "name") + apply(simp add: abs_fresh fresh_prod fresh_atm) + apply(auto)[1] + apply(drule_tac z="cb" in alpha_name) + apply(simp add: fresh_prod fresh_atm abs_fresh) + apply(simp) + apply(rule exI)+ + apply(rule conjI) + apply(rule exI)+ + apply(rule_tac s="x" and t="[(x,cb)]\[(y,cb)]\y" in subst) + apply(simp add: calc_atm) + apply(rule refl) + apply(auto simp: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] + apply(perm_simp)+ + done +next + case (LOr2 b a M N1 N2 y x1 x2) + then show ?thesis + apply - + apply(rule disjI2) + apply(rule disjI2) + apply(rule disjI2) + apply(rule disjI2) + apply(rule disjI2) + apply(rule disjI2) + apply(rule disjI1) + apply(simp add: trm.inject) + apply(erule conjE)+ + apply(generate_fresh "coname") + apply(simp add: abs_fresh fresh_prod fresh_atm) + apply(auto)[1] + apply(drule_tac c="c" in alpha_coname) + apply(simp add: fresh_prod fresh_atm abs_fresh) + apply(simp) + apply(perm_simp) + apply(rule exI)+ + apply(rule conjI) + apply(rule_tac s="a" and t="[(a,c)]\[(b,c)]\b" in subst) + apply(simp add: calc_atm) + apply(rule refl) + apply(generate_fresh "name") + apply(simp add: abs_fresh fresh_prod fresh_atm) + apply(auto)[1] + apply(drule_tac z="ca" in alpha_name) + apply(simp add: fresh_prod fresh_atm abs_fresh) + apply(simp) + apply(rule exI)+ + apply(rule conjI) + apply(rule_tac s="x" and t="[(x,ca)]\[(y,ca)]\y" in subst) + apply(simp add: calc_atm) + apply(rule refl) + apply(auto simp: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] + apply(perm_simp)+ + apply(generate_fresh "name") + apply(simp add: abs_fresh fresh_prod fresh_atm) + apply(auto)[1] + apply(drule_tac z="cb" in alpha_name) + apply(simp add: fresh_prod fresh_atm abs_fresh) + apply(simp) + apply(rule exI)+ + apply(rule conjI) + apply(rule_tac s="x" and t="[(x,cb)]\[(y,cb)]\y" in subst) + apply(simp add: calc_atm) + apply(rule refl) + apply(auto simp: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] + apply(perm_simp)+ + done +next + case (LImp z N y P x M b a c) + then show ?thesis + apply(intro disjI2) + apply(clarsimp simp add: trm.inject) + apply(generate_fresh "coname") + apply(clarsimp simp add: abs_fresh fresh_prod fresh_atm) + apply(drule_tac c="ca" in alpha_coname) + apply(simp add: fresh_prod fresh_atm abs_fresh) + apply(simp) + apply(perm_simp) + apply(rule exI)+ + apply(rule conjI) + apply(rule_tac s="a" and t="[(a,ca)]\[(b,ca)]\b" in subst) + apply(simp add: calc_atm) + apply(rule refl) + apply(generate_fresh "name") + apply(simp add: abs_fresh fresh_prod fresh_atm) + apply(auto)[1] + apply(drule_tac z="caa" in alpha_name) + apply(simp add: fresh_prod fresh_atm abs_fresh) + apply(simp) + apply(perm_simp) + apply(rule exI)+ + apply(rule conjI) + apply(rule_tac s="x" and t="[(x,caa)]\[(z,caa)]\z" in subst) + apply(simp add: calc_atm) + apply(rule refl) + 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) + apply(auto)[1] + apply(drule_tac z="cb" in alpha_name) + apply(simp add: fresh_prod fresh_atm abs_fresh) + apply(simp) + apply(perm_simp) + apply(rule exI)+ + apply(rule conjI) + apply(rule_tac s="x" and t="[(x,cb)]\[(z,cb)]\z" in subst) + apply(simp add: calc_atm) + apply(rule refl) + apply(auto simp: fresh_left calc_atm abs_fresh alpha perm_fresh_fresh split: if_splits)[1] + apply(perm_simp)+ + done +qed + inductive c_redu :: "trm \ trm \ bool" ("_ \\<^sub>c _" [100,100] 100) @@ -3321,11 +3063,8 @@ have "Cut .M (x).N = Cut .([(a',a)]\M) (x').([(x',x)]\N)" 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) - apply(clarify) - apply(drule_tac a'="a" in fic_rename) - apply(simp add: perm_swap) + apply(intro left) + apply (metis fic_rename perm_swap(4)) apply(simp add: fresh_left calc_atm)+ done also have "\ = M{a:=(x).N}" using fs1 fs2 @@ -3342,12 +3081,9 @@ have "Cut .M (x).N = Cut .([(a',a)]\M) (x').([(x',x)]\N)" 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) - apply(clarify) - apply(drule_tac x'="x" in fin_rename) - apply(simp add: perm_swap) - apply(simp add: fresh_left calc_atm)+ + apply (intro right) + apply (metis fin_rename perm_swap(3)) + apply (simp add: fresh_left calc_atm)+ done also have "\ = N{x:=.M}" using fs1 fs2 by (simp add: subst_rename[symmetric] fresh_atm fresh_prod fresh_left calc_atm) @@ -3359,12 +3095,7 @@ and c::"coname" shows "M \\<^sub>c M' \ x\M \ x\M'" and "M \\<^sub>c M' \ c\M \ c\M'" -apply - -apply(induct rule: c_redu.induct) -apply(auto simp: abs_fresh rename_fresh subst_fresh) -apply(induct rule: c_redu.induct) -apply(auto simp: abs_fresh rename_fresh subst_fresh) -done + by (induct rule: c_redu.induct) (auto simp: abs_fresh rename_fresh subst_fresh)+ inductive a_redu :: "trm \ trm \ bool" ("_ \\<^sub>a _" [100,100] 100) @@ -3387,21 +3118,17 @@ | a_ImpL_r: "\a\N; x\(M,y); N\\<^sub>a N'\ \ ImpL .M (x).N y \\<^sub>a ImpL .M (x).N' y" | a_ImpR: "\a\b; M\\<^sub>a M'\ \ ImpR (x)..M b \\<^sub>a ImpR (x)..M' b" -lemma fresh_a_redu: +lemma fresh_a_redu1: fixes x::"name" - and c::"coname" shows "M \\<^sub>a M' \ x\M \ x\M'" - and "M \\<^sub>a M' \ c\M \ c\M'" -apply - -apply(induct rule: a_redu.induct) -apply(simp add: fresh_l_redu) -apply(simp add: fresh_c_redu) -apply(auto simp: 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: abs_fresh abs_supp fin_supp) -done + by (induct rule: a_redu.induct) (auto simp: fresh_l_redu fresh_c_redu abs_fresh abs_supp fin_supp) + +lemma fresh_a_redu2: + fixes c::"coname" + shows "M \\<^sub>a M' \ c\M \ c\M'" + by (induct rule: a_redu.induct) (auto simp: fresh_l_redu fresh_c_redu abs_fresh abs_supp fin_supp) + +lemmas fresh_a_redu = fresh_a_redu1 fresh_a_redu2 equivariance a_redu @@ -3611,405 +3338,94 @@ lemma ax_do_not_a_reduce: shows "Ax x a \\<^sub>a M \ False" apply(erule_tac a_redu.cases) -apply(auto simp: trm.inject) -apply(drule ax_do_not_l_reduce) -apply(simp) -apply(drule ax_do_not_c_reduce) -apply(simp) -done +apply(simp_all add: trm.inject) + using ax_do_not_l_reduce ax_do_not_c_reduce by blast+ lemma a_redu_NotL_elim: - assumes a: "NotL .M x \\<^sub>a R" + assumes "NotL .M x \\<^sub>a R" shows "\M'. R = NotL .M' x \ M\\<^sub>aM'" -using a [[simproc del: defined_all]] -apply(erule_tac a_redu.cases, simp_all add: trm.inject) -apply(erule_tac l_redu.cases, simp_all add: trm.inject) -apply(erule_tac c_redu.cases, simp_all add: trm.inject) -apply(auto) -apply(rotate_tac 2) -apply(erule_tac a_redu.cases, simp_all add: trm.inject) -apply(erule_tac l_redu.cases, simp_all add: trm.inject) -apply(erule_tac c_redu.cases, simp_all add: trm.inject) -apply(auto simp: alpha a_redu.eqvt) -apply(rule_tac x="([(a,aa)]\M'a)" in exI) -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: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu) -apply(simp add: perm_swap) -done + using assms + 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 (smt (verit, best) a_redu.eqvt(2) alpha(2) fresh_a_redu2)+ + done lemma a_redu_NotR_elim: - assumes a: "NotR (x).M a \\<^sub>a R" + assumes "NotR (x).M a \\<^sub>a R" shows "\M'. R = NotR (x).M' a \ M\\<^sub>aM'" -using a [[simproc del: defined_all]] -apply(erule_tac a_redu.cases, simp_all add: trm.inject) -apply(erule_tac l_redu.cases, simp_all add: trm.inject) -apply(erule_tac c_redu.cases, simp_all add: trm.inject) -apply(auto) -apply(rotate_tac 2) -apply(erule_tac a_redu.cases, simp_all add: trm.inject) -apply(erule_tac l_redu.cases, simp_all add: trm.inject) -apply(erule_tac c_redu.cases, simp_all add: trm.inject) -apply(auto simp: alpha a_redu.eqvt) -apply(rule_tac x="([(x,xa)]\M'a)" in exI) -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: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu) -apply(simp add: perm_swap) -done + using assms + 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 (smt (verit, best) a_redu.eqvt(1) alpha(1) fresh_a_redu1)+ + done lemma a_redu_AndR_elim: - assumes a: "AndR .M .N c\\<^sub>a R" + assumes "AndR .M .N c\\<^sub>a R" shows "(\M'. R = AndR .M' .N c \ M\\<^sub>aM') \ (\N'. R = AndR .M .N' c \ N\\<^sub>aN')" -using a [[simproc del: defined_all]] -apply(erule_tac a_redu.cases, simp_all add: trm.inject) -apply(erule_tac l_redu.cases, simp_all add: trm.inject) -apply(erule_tac c_redu.cases, simp_all add: trm.inject) -apply(rotate_tac 6) -apply(erule_tac a_redu.cases, simp_all add: trm.inject) -apply(erule_tac l_redu.cases, simp_all add: trm.inject) -apply(erule_tac c_redu.cases, simp_all add: trm.inject) -apply(rule disjI1) -apply(auto simp: alpha a_redu.eqvt)[1] -apply(rule_tac x="([(a,aa)]\M'a)" in exI) -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: 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: 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: 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: 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: 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: 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: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] -apply(rule disjI2) -apply(auto simp: alpha a_redu.eqvt)[1] -apply(rule_tac x="([(b,ba)]\N')" in exI) -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: 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: 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: 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: 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: 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: 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: 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: alpha a_redu.eqvt)[1] -apply(rule_tac x="([(a,aa)]\M')" in exI) -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: 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: 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: 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: 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: 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: 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: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] -apply(rule disjI2) -apply(auto simp: alpha a_redu.eqvt)[1] -apply(rule_tac x="([(b,ba)]\N'a)" in exI) -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: 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: 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: 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: 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: 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: 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: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] -done + using assms + 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 (smt (verit) a_redu.eqvt(2) alpha'(2) fresh_a_redu2) + by (metis a_NotL a_redu_NotL_elim trm.inject(4)) lemma a_redu_AndL1_elim: - assumes a: "AndL1 (x).M y \\<^sub>a R" + assumes "AndL1 (x).M y \\<^sub>a R" shows "\M'. R = AndL1 (x).M' y \ M\\<^sub>aM'" -using a [[simproc del: defined_all]] -apply(erule_tac a_redu.cases, simp_all add: trm.inject) -apply(erule_tac l_redu.cases, simp_all add: trm.inject) -apply(erule_tac c_redu.cases, simp_all add: trm.inject) -apply(auto) -apply(rotate_tac 3) -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: alpha a_redu.eqvt) -apply(rule_tac x="([(x,xa)]\M'a)" in exI) -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: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu) -apply(simp add: perm_swap) -done + using assms + 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) + by (metis a_NotR a_redu_NotR_elim trm.inject(3)) lemma a_redu_AndL2_elim: assumes a: "AndL2 (x).M y \\<^sub>a R" shows "\M'. R = AndL2 (x).M' y \ M\\<^sub>aM'" -using a [[simproc del: defined_all]] -apply(erule_tac a_redu.cases, simp_all add: trm.inject) -apply(erule_tac l_redu.cases, simp_all add: trm.inject) -apply(erule_tac c_redu.cases, simp_all add: trm.inject) -apply(auto) -apply(rotate_tac 3) -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: alpha a_redu.eqvt) -apply(rule_tac x="([(x,xa)]\M'a)" in exI) -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: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu) -apply(simp add: perm_swap) -done + using a + apply(erule_tac a_redu.cases, simp_all add: trm.inject) + apply(erule_tac l_redu.cases, simp_all add: trm.inject) + apply(erule_tac c_redu.cases, simp_all add: trm.inject) + by (metis a_NotR a_redu_NotR_elim trm.inject(3)) lemma a_redu_OrL_elim: - assumes a: "OrL (x).M (y).N z\\<^sub>a R" + assumes "OrL (x).M (y).N z\\<^sub>a R" shows "(\M'. R = OrL (x).M' (y).N z \ M\\<^sub>aM') \ (\N'. R = OrL (x).M (y).N' z \ N\\<^sub>aN')" -using a [[simproc del: defined_all]] -apply(erule_tac a_redu.cases, simp_all add: trm.inject) -apply(erule_tac l_redu.cases, simp_all add: trm.inject) -apply(erule_tac c_redu.cases, simp_all add: trm.inject) -apply(rotate_tac 6) -apply(erule_tac a_redu.cases, simp_all add: trm.inject) -apply(erule_tac l_redu.cases, simp_all add: trm.inject) -apply(erule_tac c_redu.cases, simp_all add: trm.inject) -apply(rule disjI1) -apply(auto simp: alpha a_redu.eqvt)[1] -apply(rule_tac x="([(x,xa)]\M'a)" in exI) -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: 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: 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: 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: 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: 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: 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: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] -apply(rule disjI2) -apply(auto simp: alpha a_redu.eqvt)[1] -apply(rule_tac x="([(y,ya)]\N')" in exI) -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: 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: 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: 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: 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: 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: 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: 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: alpha a_redu.eqvt)[1] -apply(rule_tac x="([(x,xa)]\M')" in exI) -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: 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: 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: 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: 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: 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: 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: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] -apply(rule disjI2) -apply(auto simp: alpha a_redu.eqvt)[1] -apply(rule_tac x="([(y,ya)]\N'a)" in exI) -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: 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: 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: 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: 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: 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: 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: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] + using assms + 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 (metis a_NotR a_redu_NotR_elim trm.inject(3))+ done lemma a_redu_OrR1_elim: - assumes a: "OrR1 .M b \\<^sub>a R" + assumes "OrR1 .M b \\<^sub>a R" shows "\M'. R = OrR1 .M' b \ M\\<^sub>aM'" -using a [[simproc del: defined_all]] -apply(erule_tac a_redu.cases, simp_all add: trm.inject) -apply(erule_tac l_redu.cases, simp_all add: trm.inject) -apply(erule_tac c_redu.cases, simp_all add: trm.inject) -apply(auto) -apply(rotate_tac 3) -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: alpha a_redu.eqvt) -apply(rule_tac x="([(a,aa)]\M'a)" in exI) -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: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu) -apply(simp add: perm_swap) -done + using assms + 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) + by (metis a_NotL a_redu_NotL_elim trm.inject(4)) lemma a_redu_OrR2_elim: assumes a: "OrR2 .M b \\<^sub>a R" shows "\M'. R = OrR2 .M' b \ M\\<^sub>aM'" -using a [[simproc del: defined_all]] -apply(erule_tac a_redu.cases, simp_all add: trm.inject) -apply(erule_tac l_redu.cases, simp_all add: trm.inject) -apply(erule_tac c_redu.cases, simp_all add: trm.inject) -apply(auto) -apply(rotate_tac 3) -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: alpha a_redu.eqvt) -apply(rule_tac x="([(a,aa)]\M'a)" in exI) -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: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu) -apply(simp add: perm_swap) -done + using assms + 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) + by (metis a_redu_OrR1_elim better_OrR1_intro trm.inject(8)) lemma a_redu_ImpL_elim: assumes a: "ImpL .M (y).N z\\<^sub>a R" shows "(\M'. R = ImpL .M' (y).N z \ M\\<^sub>aM') \ (\N'. R = ImpL .M (y).N' z \ N\\<^sub>aN')" -using a [[simproc del: defined_all]] -apply(erule_tac a_redu.cases, simp_all add: trm.inject) -apply(erule_tac l_redu.cases, simp_all add: trm.inject) -apply(erule_tac c_redu.cases, simp_all add: trm.inject) -apply(rotate_tac 5) -apply(erule_tac a_redu.cases, simp_all add: trm.inject) -apply(erule_tac l_redu.cases, simp_all add: trm.inject) -apply(erule_tac c_redu.cases, simp_all add: trm.inject) -apply(rule disjI1) -apply(auto simp: alpha a_redu.eqvt)[1] -apply(rule_tac x="([(a,aa)]\M'a)" in exI) -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: 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: 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: 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: 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: 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: 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: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] -apply(rule disjI2) -apply(auto simp: alpha a_redu.eqvt)[1] -apply(rule_tac x="([(y,xa)]\N')" in exI) -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: 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: 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: 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: 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: 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: 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: 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: alpha a_redu.eqvt)[1] -apply(rule_tac x="([(a,aa)]\M')" in exI) -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: 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: 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: 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: 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: 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: 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: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] -apply(rule disjI2) -apply(auto simp: alpha a_redu.eqvt)[1] -apply(rule_tac x="([(y,xa)]\N'a)" in exI) -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: 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: 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: 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: 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: 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: 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: perm_swap fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu)[1] -done + using assms + 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 (smt (verit) a_redu.eqvt(2) alpha'(2) fresh_a_redu2) + by (metis a_NotR a_redu_NotR_elim trm.inject(3)) lemma a_redu_ImpR_elim: assumes a: "ImpR (x)..M b \\<^sub>a R" @@ -4022,7 +3438,7 @@ apply(rotate_tac 3) 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(erule_tac c_redu.cases, simp_all add: trm.inject) apply(auto simp: alpha a_redu.eqvt abs_perm abs_fresh) apply(rule_tac x="([(a,aa)]\M'a)" in exI) apply(auto simp: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap) @@ -4054,56 +3470,35 @@ 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: 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: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap) -apply(rule sym) -apply(rule trans) -apply(rule perm_compose) -apply(simp add: calc_atm perm_swap) -done + apply (simp add: cp_coname_name1 perm_dj(4) perm_swap(3,4)) + apply(rule_tac x="([(a,aaa)]\[(x,xaa)]\M'a)" in exI) +by(simp add: fresh_left alpha a_redu.eqvt calc_atm fresh_a_redu perm_swap perm_compose(4) perm_dj(4)) text \Transitive Closure\ abbreviation - a_star_redu :: "trm \ trm \ bool" ("_ \\<^sub>a* _" [100,100] 100) + a_star_redu :: "trm \ trm \ bool" ("_ \\<^sub>a* _" [80,80] 80) where "M \\<^sub>a* M' \ (a_redu)\<^sup>*\<^sup>* M M'" -lemma a_starI: - assumes a: "M \\<^sub>a M'" - shows "M \\<^sub>a* M'" -using a by blast +lemmas a_starI = r_into_rtranclp lemma a_starE: assumes a: "M \\<^sub>a* M'" shows "M = M' \ (\N. M \\<^sub>a N \ N \\<^sub>a* M')" -using a -by (induct) (auto) + using a by (induct) (auto) lemma a_star_refl: shows "M \\<^sub>a* M" by blast -lemma a_star_trans[trans]: - assumes a1: "M1\\<^sub>a* M2" - and a2: "M2\\<^sub>a* M3" - shows "M1 \\<^sub>a* M3" -using a2 a1 -by (induct) (auto) +declare rtranclp_trans [trans] text \congruence rules for \\\<^sub>a*\\ lemma ax_do_not_a_star_reduce: shows "Ax x a \\<^sub>a* M \ M = Ax x a" -apply(induct set: rtranclp) -apply(auto) -apply(drule ax_do_not_a_reduce) -apply(simp) -done + using a_starE ax_do_not_a_reduce by blast lemma a_star_CutL: @@ -4186,500 +3581,199 @@ a_star_OrL a_star_OrR1 a_star_OrR2 a_star_ImpL a_star_ImpR lemma a_star_redu_NotL_elim: - assumes a: "NotL .M x \\<^sub>a* R" + assumes "NotL .M x \\<^sub>a* R" shows "\M'. R = NotL .M' x \ M \\<^sub>a* M'" -using a -apply(induct set: rtranclp) -apply(auto) -apply(drule a_redu_NotL_elim) -apply(auto) -done + using assms by (induct set: rtranclp) (use a_redu_NotL_elim in force)+ lemma a_star_redu_NotR_elim: - assumes a: "NotR (x).M a \\<^sub>a* R" + assumes "NotR (x).M a \\<^sub>a* R" shows "\M'. R = NotR (x).M' a \ M \\<^sub>a* M'" -using a -apply(induct set: rtranclp) -apply(auto) -apply(drule a_redu_NotR_elim) -apply(auto) -done + using assms by (induct set: rtranclp) (use a_redu_NotR_elim in force)+ lemma a_star_redu_AndR_elim: - assumes a: "AndR .M .N c\\<^sub>a* R" + assumes "AndR .M .N c\\<^sub>a* R" shows "(\M' N'. R = AndR .M' .N' c \ M \\<^sub>a* M' \ N \\<^sub>a* N')" -using a -apply(induct set: rtranclp) -apply(auto) -apply(drule a_redu_AndR_elim) -apply(auto simp: alpha trm.inject) -done + using assms by (induct set: rtranclp) (use a_redu_AndR_elim in force)+ lemma a_star_redu_AndL1_elim: - assumes a: "AndL1 (x).M y \\<^sub>a* R" + assumes "AndL1 (x).M y \\<^sub>a* R" shows "\M'. R = AndL1 (x).M' y \ M \\<^sub>a* M'" -using a -apply(induct set: rtranclp) -apply(auto) -apply(drule a_redu_AndL1_elim) -apply(auto simp: alpha trm.inject) -done + using assms by (induct set: rtranclp) (use a_redu_AndL1_elim in force)+ lemma a_star_redu_AndL2_elim: - assumes a: "AndL2 (x).M y \\<^sub>a* R" + assumes "AndL2 (x).M y \\<^sub>a* R" shows "\M'. R = AndL2 (x).M' y \ M \\<^sub>a* M'" -using a -apply(induct set: rtranclp) -apply(auto) -apply(drule a_redu_AndL2_elim) -apply(auto simp: alpha trm.inject) -done + using assms by (induct set: rtranclp) (use a_redu_AndL2_elim in force)+ lemma a_star_redu_OrL_elim: - assumes a: "OrL (x).M (y).N z \\<^sub>a* R" + assumes "OrL (x).M (y).N z \\<^sub>a* R" shows "(\M' N'. R = OrL (x).M' (y).N' z \ M \\<^sub>a* M' \ N \\<^sub>a* N')" -using a -apply(induct set: rtranclp) -apply(auto) -apply(drule a_redu_OrL_elim) -apply(auto simp: alpha trm.inject) -done + using assms by (induct set: rtranclp) (use a_redu_OrL_elim in force)+ lemma a_star_redu_OrR1_elim: - assumes a: "OrR1 .M y \\<^sub>a* R" + assumes "OrR1 .M y \\<^sub>a* R" shows "\M'. R = OrR1 .M' y \ M \\<^sub>a* M'" -using a -apply(induct set: rtranclp) -apply(auto) -apply(drule a_redu_OrR1_elim) -apply(auto simp: alpha trm.inject) -done + using assms by (induct set: rtranclp) (use a_redu_OrR1_elim in force)+ lemma a_star_redu_OrR2_elim: - assumes a: "OrR2 .M y \\<^sub>a* R" + assumes "OrR2 .M y \\<^sub>a* R" shows "\M'. R = OrR2 .M' y \ M \\<^sub>a* M'" -using a -apply(induct set: rtranclp) -apply(auto) -apply(drule a_redu_OrR2_elim) -apply(auto simp: alpha trm.inject) -done + using assms by (induct set: rtranclp) (use a_redu_OrR2_elim in force)+ lemma a_star_redu_ImpR_elim: - assumes a: "ImpR (x)..M y \\<^sub>a* R" + assumes "ImpR (x)..M y \\<^sub>a* R" shows "\M'. R = ImpR (x)..M' y \ M \\<^sub>a* M'" -using a -apply(induct set: rtranclp) -apply(auto) -apply(drule a_redu_ImpR_elim) -apply(auto simp: alpha trm.inject) -done + using assms by (induct set: rtranclp) (use a_redu_ImpR_elim in force)+ lemma a_star_redu_ImpL_elim: - assumes a: "ImpL .M (y).N z \\<^sub>a* R" + assumes "ImpL .M (y).N z \\<^sub>a* R" shows "(\M' N'. R = ImpL .M' (y).N' z \ M \\<^sub>a* M' \ N \\<^sub>a* N')" -using a -apply(induct set: rtranclp) -apply(auto) -apply(drule a_redu_ImpL_elim) -apply(auto simp: alpha trm.inject) -done + using assms by (induct set: rtranclp) (use a_redu_ImpL_elim in force)+ text \Substitution\ lemma subst_not_fin1: shows "\fin(M{x:=.P}) x" -apply(nominal_induct M avoiding: x c P rule: trm.strong_induct) -apply(auto) -apply(drule fin_elims, simp) -apply(drule fin_elims, simp) -apply(erule fin.cases, simp_all add: trm.inject) -apply(erule fin.cases, simp_all add: trm.inject) -apply(erule fin.cases, simp_all add: trm.inject) -apply(subgoal_tac "\x'::name. x'\(trm{x:=.P},P)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_NotL) -apply(erule fin.cases, simp_all add: trm.inject) -apply(rule exists_fresh'(1)[OF fs_name1]) -apply(erule fin.cases, simp_all add: trm.inject) -apply(erule fin.cases, simp_all add: trm.inject) -apply(subgoal_tac "\x'::name. x'\(trm{x:=.P},P,name1)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_AndL1) -apply(erule fin.cases, simp_all add: trm.inject) -apply(rule exists_fresh'(1)[OF fs_name1]) -apply(erule fin.cases, simp_all add: trm.inject) -apply(subgoal_tac "\x'::name. x'\(trm{x:=.P},P,name1)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_AndL2) -apply(erule fin.cases, simp_all add: trm.inject) -apply(rule exists_fresh'(1)[OF fs_name1]) -apply(erule fin.cases, simp_all add: trm.inject) -apply(erule fin.cases, simp_all add: trm.inject) -apply(erule fin.cases, simp_all add: trm.inject) -apply(subgoal_tac "\x'::name. x'\(trm1{x:=.P},P,name1,trm2{x:=.P},name2)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_OrL) -apply(erule fin.cases, simp_all add: trm.inject) -apply(rule exists_fresh'(1)[OF fs_name1]) -apply(erule fin.cases, simp_all add: trm.inject) -apply(erule fin.cases, simp_all add: trm.inject) -apply(subgoal_tac "\x'::name. x'\(trm1{name2:=.P},P,name1,trm2{name2:=.P})") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_ImpL) -apply(erule fin.cases, simp_all add: trm.inject) -apply(rule exists_fresh'(1)[OF fs_name1]) -apply(erule fin.cases, simp_all add: trm.inject) -done - -lemma subst_not_fin2: - assumes a: "\fin M y" - shows "\fin(M{c:=(x).P}) y" -using a -apply(nominal_induct M avoiding: x c P y rule: trm.strong_induct) -apply(auto) -apply(drule fin_elims, simp) -apply(drule fin_elims, simp) -apply(drule fin_elims, simp) -apply(subgoal_tac "\c'::coname. c'\(trm{coname:=(x).P},P)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_NotR) -apply(drule fin_elims, simp) -apply(rule exists_fresh'(2)[OF fs_coname1]) -apply(drule fin_elims, simp) -apply(drule fin_elims, simp) -apply(drule freshn_after_substc) -apply(simp add: fin.intros) -apply(subgoal_tac "\c'::coname. c'\(trm1{coname3:=(x).P},P,coname1,trm2{coname3:=(x).P},coname2)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_AndR) -apply(drule fin_elims, simp) -apply(rule exists_fresh'(2)[OF fs_coname1]) -apply(drule fin_elims, simp) -apply(drule fin_elims, simp) -apply(simp add: abs_fresh fresh_atm) -apply(drule freshn_after_substc) -apply(simp add: fin.intros abs_fresh) -apply(drule fin_elims, simp) -apply(simp add: abs_fresh fresh_atm) -apply(drule freshn_after_substc) -apply(simp add: fin.intros abs_fresh) -apply(subgoal_tac "\c'::coname. c'\(trm{coname2:=(x).P},P,coname1)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_OrR1) -apply(drule fin_elims, simp) -apply(rule exists_fresh'(2)[OF fs_coname1]) -apply(drule fin_elims, simp) -apply(subgoal_tac "\c'::coname. c'\(trm{coname2:=(x).P},P,coname1)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_OrR2) -apply(drule fin_elims, simp) -apply(rule exists_fresh'(2)[OF fs_coname1]) -apply(drule fin_elims, simp) -apply(drule fin_elims, simp) -apply(auto)[1] -apply(simp add: abs_fresh fresh_atm) -apply(drule freshn_after_substc) -apply(drule freshn_after_substc) -apply(simp add: fin.intros abs_fresh) -apply(subgoal_tac "\c'::coname. c'\(trm{coname2:=(x).P},P,coname1,coname2)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_ImpR) -apply(drule fin_elims, simp) -apply(rule exists_fresh'(2)[OF fs_coname1]) -apply(drule fin_elims, simp) -apply(drule fin_elims, simp) -apply(auto)[1] -apply(simp add: abs_fresh fresh_atm) -apply(drule freshn_after_substc) -apply(drule freshn_after_substc) -apply(simp add: fin.intros abs_fresh) -done - -lemma subst_not_fic1: - shows "\fic (M{a:=(x).P}) a" -apply(nominal_induct M avoiding: a x P rule: trm.strong_induct) -apply(auto) -apply(erule fic.cases, simp_all add: trm.inject) -apply(erule fic.cases, simp_all add: trm.inject) -apply(erule fic.cases, simp_all add: trm.inject) -apply(erule fic.cases, simp_all add: trm.inject) -apply(subgoal_tac "\a'::coname. a'\(trm{coname:=(x).P},P)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_NotR) -apply(erule fic.cases, simp_all add: trm.inject) -apply(rule exists_fresh'(2)[OF fs_coname1]) -apply(erule fic.cases, simp_all add: trm.inject) -apply(erule fic.cases, simp_all add: trm.inject) -apply(subgoal_tac "\a'::coname. a'\(trm1{coname3:=(x).P},P,trm2{coname3:=(x).P},coname1,coname2)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_AndR) -apply(erule fic.cases, simp_all add: trm.inject) -apply(rule exists_fresh'(2)[OF fs_coname1]) -apply(erule fic.cases, simp_all add: trm.inject) -apply(erule fic.cases, simp_all add: trm.inject) -apply(erule fic.cases, simp_all add: trm.inject) -apply(subgoal_tac "\a'::coname. a'\(trm{coname2:=(x).P},P,coname1)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_OrR1) -apply(erule fic.cases, simp_all add: trm.inject) -apply(rule exists_fresh'(2)[OF fs_coname1]) -apply(erule fic.cases, simp_all add: trm.inject) -apply(subgoal_tac "\a'::coname. a'\(trm{coname2:=(x).P},P,coname1)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_OrR2) -apply(erule fic.cases, simp_all add: trm.inject) -apply(rule exists_fresh'(2)[OF fs_coname1]) -apply(erule fic.cases, simp_all add: trm.inject) -apply(erule fic.cases, simp_all add: trm.inject) -apply(subgoal_tac "\a'::coname. a'\(trm{coname2:=(x).P},P,coname1)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_ImpR) -apply(erule fic.cases, simp_all add: trm.inject) -apply(rule exists_fresh'(2)[OF fs_coname1]) -apply(erule fic.cases, simp_all add: trm.inject) -apply(erule fic.cases, simp_all add: trm.inject) -done - -lemma subst_not_fic2: - assumes a: "\fic M a" - shows "\fic(M{x:=.P}) a" -using a -apply(nominal_induct M avoiding: x a P b rule: trm.strong_induct) -apply(auto) -apply(drule fic_elims, simp) -apply(drule fic_elims, simp) -apply(drule fic_elims, simp) -apply(drule fic_elims, simp) -apply(drule freshc_after_substn) -apply(simp add: fic.intros) -apply(subgoal_tac "\x'::name. x'\(trm{x:=.P},P)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_NotL) -apply(drule fic_elims, simp) -apply(rule exists_fresh'(1)[OF fs_name1]) -apply(drule fic_elims, simp) -apply(drule fic_elims, simp) -apply(auto)[1] -apply(simp add: abs_fresh fresh_atm) -apply(drule freshc_after_substn) -apply(drule freshc_after_substn) -apply(simp add: fic.intros abs_fresh) -apply(subgoal_tac "\x'::name. x'\(trm{x:=.P},P,name1)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_AndL1) -apply(drule fic_elims, simp) -apply(rule exists_fresh'(1)[OF fs_name1]) -apply(drule fic_elims, simp) -apply(subgoal_tac "\x'::name. x'\(trm{x:=.P},P,name1)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_AndL2) -apply(drule fic_elims, simp) -apply(rule exists_fresh'(1)[OF fs_name1]) -apply(drule fic_elims, simp) -apply(drule fic_elims, simp) -apply(simp add: abs_fresh fresh_atm) -apply(drule freshc_after_substn) -apply(simp add: fic.intros abs_fresh) -apply(drule fic_elims, simp) -apply(simp add: abs_fresh fresh_atm) -apply(drule freshc_after_substn) -apply(simp add: fic.intros abs_fresh) -apply(subgoal_tac "\x'::name. x'\(trm1{x:=.P},P,name1,trm2{x:=.P},name2)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_OrL) -apply(drule fic_elims, simp) -apply(rule exists_fresh'(1)[OF fs_name1]) -apply(drule fic_elims, simp) -apply(drule fic_elims, simp) -apply(simp add: abs_fresh fresh_atm) -apply(drule freshc_after_substn) -apply(simp add: fic.intros abs_fresh) -apply(subgoal_tac "\x'::name. x'\(trm1{name2:=.P},trm2{name2:=.P},P,name1)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_ImpL) -apply(drule fic_elims, simp) -apply(rule exists_fresh'(1)[OF fs_name1]) -apply(drule fic_elims, simp) -done +proof (nominal_induct M avoiding: x c P rule: trm.strong_induct) + case (Ax name coname) + with fin_rest_elims show ?case + by (auto simp: fin_Ax_elim) +next + case (NotL coname trm name) + obtain x'::name where "x'\(trm{x:=.P},P)" + by (meson exists_fresh(1) fs_name1) + with NotL fin_NotL_elim fresh_fun_simp_NotL show ?case + by simp (metis fin_rest_elims(1) fresh_prod) +next + case (AndL1 name1 trm name2) + obtain x'::name where "x' \ (trm{x:=.P},P,name1)" + by (meson exists_fresh(1) fs_name1) + with AndL1 fin_AndL1_elim fresh_fun_simp_AndL1 show ?case + by simp (metis fin_rest_elims(1) fresh_prod) +next + case (AndL2 name2 trm name2) + obtain x'::name where "x' \ (trm{x:=.P},P,name2)" + by (meson exists_fresh(1) fs_name1) + with AndL2 fin_AndL2_elim fresh_fun_simp_AndL2 better_AndL2_substn show ?case + by (metis abs_fresh(1) fresh_atm(1) not_fin_subst2) +next + case (OrL name1 trm1 name2 trm2 name3) + obtain x'::name where "x' \ (trm1{x:=.P},P,name1,trm2{x:=.P},name2)" + by (meson exists_fresh(1) fs_name1) + with OrL fin_OrL_elim fresh_fun_simp_OrL show ?case + by simp (metis fin_rest_elims(1) fresh_prod) +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 fin_ImpL_elim fresh_fun_simp_ImpL show ?case + by simp (metis fin_rest_elims(1) fresh_prod) +qed (use fin_rest_elims not_fin_subst2 in auto) + +lemmas subst_not_fin2 = not_fin_subst1 text \Reductions\ lemma fin_l_reduce: - assumes a: "fin M x" - and b: "M \\<^sub>l M'" + assumes "fin M x" and "M \\<^sub>l M'" shows "fin M' x" -using b a -apply(induct) -apply(erule fin.cases) -apply(simp_all add: trm.inject) -apply(rotate_tac 3) -apply(erule fin.cases) -apply(simp_all add: trm.inject) -apply(erule fin.cases, simp_all add: trm.inject)+ -done + using assms fin_rest_elims(1) l_redu.simps by fastforce lemma fin_c_reduce: - assumes a: "fin M x" - and b: "M \\<^sub>c M'" + assumes "fin M x" and "M \\<^sub>c M'" shows "fin M' x" -using b a -apply(induct) -apply(erule fin.cases, simp_all add: trm.inject)+ -done + using assms c_redu.simps fin_rest_elims(1) by fastforce lemma fin_a_reduce: assumes a: "fin M x" and b: "M \\<^sub>a M'" shows "fin M' x" -using a b -apply(induct) -apply(drule ax_do_not_a_reduce) -apply(simp) -apply(drule a_redu_NotL_elim) -apply(auto) -apply(rule fin.intros) -apply(simp add: fresh_a_redu) -apply(drule a_redu_AndL1_elim) -apply(auto) -apply(rule fin.intros) -apply(force simp add: abs_fresh fresh_a_redu) -apply(drule a_redu_AndL2_elim) -apply(auto) -apply(rule fin.intros) -apply(force simp add: abs_fresh fresh_a_redu) -apply(drule a_redu_OrL_elim) -apply(auto) -apply(rule fin.intros) -apply(force simp add: abs_fresh fresh_a_redu) -apply(force simp add: abs_fresh fresh_a_redu) -apply(rule fin.intros) -apply(force simp add: abs_fresh fresh_a_redu) -apply(force simp add: abs_fresh fresh_a_redu) -apply(drule a_redu_ImpL_elim) -apply(auto) -apply(rule fin.intros) -apply(force simp add: abs_fresh fresh_a_redu) -apply(force simp add: abs_fresh fresh_a_redu) -apply(rule fin.intros) -apply(force simp add: abs_fresh fresh_a_redu) -apply(force simp add: abs_fresh fresh_a_redu) -done +using assms +proof (induct) + case (1 x a) + then show ?case + using ax_do_not_a_reduce by auto +next + case (2 x M a) + then show ?case + using a_redu_NotL_elim fresh_a_redu1 by blast +next + case (3 y x M) + then show ?case + by (metis a_redu_AndL1_elim abs_fresh(1) fin.intros(3) fresh_a_redu1) +next + case (4 y x M) + then show ?case + by (metis a_redu_AndL2_elim abs_fresh(1) fin.intros(4) fresh_a_redu1) +next + case (5 z x M y N) + then show ?case + by (smt (verit) a_redu_OrL_elim abs_fresh(1) fin.intros(5) fresh_a_redu1) +next + case (6 y M x N a) + then show ?case + by (smt (verit, best) a_redu_ImpL_elim abs_fresh(1) fin.simps fresh_a_redu1) +qed + lemma fin_a_star_reduce: assumes a: "fin M x" and b: "M \\<^sub>a* M'" shows "fin M' x" -using b a -apply(induct set: rtranclp) -apply(auto simp: fin_a_reduce) -done +using b a by (induct set: rtranclp)(auto simp: fin_a_reduce) lemma fic_l_reduce: assumes a: "fic M x" and b: "M \\<^sub>l M'" shows "fic M' x" -using b a -apply(induct) -apply(erule fic.cases) -apply(simp_all add: trm.inject) -apply(rotate_tac 3) -apply(erule fic.cases) -apply(simp_all add: trm.inject) -apply(erule fic.cases, simp_all add: trm.inject)+ -done + using b a + by induction (use fic_rest_elims in force)+ lemma fic_c_reduce: assumes a: "fic M x" and b: "M \\<^sub>c M'" shows "fic M' x" using b a -apply(induct) -apply(erule fic.cases, simp_all add: trm.inject)+ -done + by induction (use fic_rest_elims in force)+ lemma fic_a_reduce: assumes a: "fic M x" and b: "M \\<^sub>a M'" - shows "fic M' x" -using a b -apply(induct) -apply(drule ax_do_not_a_reduce) -apply(simp) -apply(drule a_redu_NotR_elim) -apply(auto) -apply(rule fic.intros) -apply(simp add: fresh_a_redu) -apply(drule a_redu_AndR_elim) -apply(auto) -apply(rule fic.intros) -apply(force simp add: abs_fresh fresh_a_redu) -apply(force simp add: abs_fresh fresh_a_redu) -apply(rule fic.intros) -apply(force simp add: abs_fresh fresh_a_redu) -apply(force simp add: abs_fresh fresh_a_redu) -apply(drule a_redu_OrR1_elim) -apply(auto) -apply(rule fic.intros) -apply(force simp add: abs_fresh fresh_a_redu) -apply(drule a_redu_OrR2_elim) -apply(auto) -apply(rule fic.intros) -apply(force simp add: abs_fresh fresh_a_redu) -apply(drule a_redu_ImpR_elim) -apply(auto) -apply(rule fic.intros) -apply(force simp add: abs_fresh fresh_a_redu) -done +shows "fic M' x" + using assms +proof induction + case (1 x a) + then show ?case + using ax_do_not_a_reduce by fastforce +next + case (2 a M x) + then show ?case + using a_redu_NotR_elim fresh_a_redu2 by blast +next + case (3 c a M b N) + then show ?case + by (smt (verit) a_redu_AndR_elim abs_fresh(2) fic.intros(3) fresh_a_redu2) +next + case (4 b a M) + then show ?case + by (metis a_redu_OrR1_elim abs_fresh(2) fic.intros(4) fresh_a_redu2) +next + case (5 b a M) + then show ?case + by (metis a_redu_OrR2_elim abs_fresh(2) fic.simps fresh_a_redu2) +next + case (6 b a M x) + then show ?case + by (metis a_redu_ImpR_elim abs_fresh(2) fic.simps fresh_a_redu2) +qed + lemma fic_a_star_reduce: assumes a: "fic M x" and b: "M \\<^sub>a* M'" shows "fic M' x" using b a -apply(induct set: rtranclp) -apply(auto simp: fic_a_reduce) -done + by (induct set: rtranclp) (auto simp: fic_a_reduce) text \substitution properties\ @@ -4744,11 +3838,7 @@ using new by (simp add: fresh_fun_simp_NotL fresh_prod) also have "\ \\<^sub>a* (NotL .(M{x:=.Ax y a}) x')[x'\n>y]" using new - apply(rule_tac a_starI) - apply(rule al_redu) - apply(rule better_LAxL_intro) - apply(auto) - done + by (intro a_starI al_redu better_LAxL_intro) auto also have "\ = NotL .(M{x:=.Ax y a}) y" using new by (simp add: nrename_fresh) also have "\ \\<^sub>a* NotL .(M[x\n>y]) y" using ih by (auto intro: a_star_congs) also have "\ = (NotL .M z)[x\n>y]" using eq by simp @@ -4786,12 +3876,7 @@ using new by (simp add: fresh_fun_simp_AndL1 fresh_prod) also have "\ \\<^sub>a* (AndL1 (u).(M{x:=.Ax y a}) v')[v'\n>y]" using new - apply(rule_tac a_starI) - apply(rule a_redu.intros) - apply(rule better_LAxL_intro) - apply(rule fin.intros) - apply(simp add: abs_fresh) - done + by (intro a_starI a_redu.intros better_LAxL_intro fin.intros) (simp add: abs_fresh) also have "\ = AndL1 (u).(M{x:=.Ax y a}) y" using fs new 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) @@ -4820,12 +3905,7 @@ using new by (simp add: fresh_fun_simp_AndL2 fresh_prod) also have "\ \\<^sub>a* (AndL2 (u).(M{x:=.Ax y a}) v')[v'\n>y]" using new - apply(rule_tac a_starI) - apply(rule a_redu.intros) - apply(rule better_LAxL_intro) - apply(rule fin.intros) - apply(simp add: abs_fresh) - done + by (intro a_starI a_redu.intros better_LAxL_intro fin.intros) (simp add: abs_fresh) also have "\ = AndL2 (u).(M{x:=.Ax y a}) y" using fs new 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) @@ -4870,12 +3950,7 @@ using new by (simp add: fresh_fun_simp_OrL) also have "\ \\<^sub>a* (OrL (u).(M{x:=.Ax y a}) (v).(N{x:=.Ax y a}) z')[z'\n>y]" using new - apply(rule_tac a_starI) - apply(rule a_redu.intros) - apply(rule better_LAxL_intro) - apply(rule fin.intros) - apply(simp_all add: abs_fresh) - done + by (intro a_starI a_redu.intros better_LAxL_intro fin.intros) (simp_all add: abs_fresh) also have "\ = OrL (u).(M{x:=.Ax y a}) (v).(N{x:=.Ax y a}) y" using fs new 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" @@ -4916,12 +3991,7 @@ using new by (simp add: fresh_fun_simp_ImpL) also have "\ \\<^sub>a* (ImpL .(M{x:=.Ax y a}) (u).(N{x:=.Ax y a}) v')[v'\n>y]" using new - apply(rule_tac a_starI) - apply(rule a_redu.intros) - apply(rule better_LAxL_intro) - apply(rule fin.intros) - apply(simp_all add: abs_fresh) - done + by (intro a_starI a_redu.intros better_LAxL_intro fin.intros) (simp_all add: abs_fresh) also have "\ = ImpL .(M{x:=.Ax y a}) (u).(N{x:=.Ax y a}) y" using fs new 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" @@ -4994,12 +4064,7 @@ using new by (simp add: fresh_fun_simp_NotR fresh_prod) also have "\ \\<^sub>a* (NotR (z).(M{b:=(x).Ax x a}) a')[a'\c>a]" using new - apply(rule_tac a_starI) - apply(rule a_redu.intros) - apply(rule better_LAxR_intro) - apply(rule fic.intros) - apply(simp) - done + by (intro a_starI a_redu.intros better_LAxR_intro fic.intros) auto also have "\ = NotR (z).(M{b:=(x).Ax x a}) a" using new by (simp add: crename_fresh) also have "\ \\<^sub>a* NotR (z).(M[b\c>a]) a" using ih by (auto intro: a_star_congs) also have "\ = (NotR (z).M c)[b\c>a]" using eq by simp @@ -5036,12 +4101,7 @@ using new by (simp add: fresh_fun_simp_AndR fresh_prod) also have "\ \\<^sub>a* (AndR .(M{b:=(x).Ax x a}) .(N{b:=(x).Ax x a}) e')[e'\c>a]" using new - apply(rule_tac a_starI) - apply(rule a_redu.intros) - apply(rule better_LAxR_intro) - apply(rule fic.intros) - apply(simp_all add: abs_fresh) - done + by (intro a_starI a_redu.intros better_LAxR_intro fic.intros) (simp_all add: abs_fresh) also have "\ = AndR .(M{b:=(x).Ax x a}) .(N{b:=(x).Ax x a}) a" using fs new 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) @@ -5086,12 +4146,7 @@ using new by (simp add: fresh_fun_simp_OrR1) also have "\ \\<^sub>a* (OrR1 .M{b:=(x).Ax x a} a')[a'\c>a]" using new - apply(rule_tac a_starI) - apply(rule a_redu.intros) - apply(rule better_LAxR_intro) - apply(rule fic.intros) - apply(simp_all add: abs_fresh) - done + by (intro a_starI a_redu.intros better_LAxR_intro fic.intros) (simp_all add: abs_fresh) also have "\ = OrR1 .M{b:=(x).Ax x a} a" using fs new 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) @@ -5120,12 +4175,7 @@ using new by (simp add: fresh_fun_simp_OrR2) also have "\ \\<^sub>a* (OrR2 .M{b:=(x).Ax x a} a')[a'\c>a]" using new - apply(rule_tac a_starI) - apply(rule a_redu.intros) - apply(rule better_LAxR_intro) - apply(rule fic.intros) - apply(simp_all add: abs_fresh) - done + by (intro a_starI a_redu.intros better_LAxR_intro fic.intros) (simp_all add: abs_fresh) also have "\ = OrR2 .M{b:=(x).Ax x a} a" using fs new 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) @@ -5163,12 +4213,7 @@ using new by (simp add: fresh_fun_simp_ImpR) also have "\ \\<^sub>a* (ImpR z..M{b:=(x).Ax x a} a')[a'\c>a]" using new - apply(rule_tac a_starI) - apply(rule a_redu.intros) - apply(rule better_LAxR_intro) - apply(rule fic.intros) - apply(simp_all add: abs_fresh) - done + by (intro a_starI a_redu.intros better_LAxR_intro fic.intros) (simp_all add: abs_fresh) also have "\ = ImpR z..M{b:=(x).Ax x a} a" using fs new 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) @@ -5196,468 +4241,141 @@ text \substitution lemmas\ -lemma not_Ax1: - shows "\(b\M) \ M{b:=(y).Q} \ Ax x a" -apply(nominal_induct M avoiding: b y Q x a rule: trm.strong_induct) -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) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_NotR abs_fresh fresh_atm) -apply(rule exists_fresh'(2)[OF fs_coname1]) -apply(subgoal_tac "\x'::coname. x'\(trm{coname:=(y).Q},Q)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_NotR abs_fresh fresh_atm) -apply(rule exists_fresh'(2)[OF fs_coname1]) -apply(subgoal_tac "\x'::coname. x'\(trm1{coname3:=(y).Q},Q,trm2{coname3:=(y).Q},coname1,coname2)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_AndR abs_fresh fresh_atm) -apply(rule exists_fresh'(2)[OF fs_coname1]) -apply(subgoal_tac "\x'::coname. x'\(trm1{coname3:=(y).Q},Q,trm2{coname3:=(y).Q},coname1,coname2)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_AndR abs_fresh fresh_atm) -apply(rule exists_fresh'(2)[OF fs_coname1]) -apply(subgoal_tac "\x'::coname. x'\(trm1{coname3:=(y).Q},Q,trm2{coname3:=(y).Q},coname1,coname2)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_AndR abs_fresh fresh_atm) -apply(rule exists_fresh'(2)[OF fs_coname1]) -apply(subgoal_tac "\x'::coname. x'\(trm{coname2:=(y).Q},Q,coname1)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_OrR1 abs_fresh fresh_atm) -apply(rule exists_fresh'(2)[OF fs_coname1]) -apply(subgoal_tac "\x'::coname. x'\(trm{coname2:=(y).Q},Q,coname1)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_OrR1 abs_fresh fresh_atm) -apply(rule exists_fresh'(2)[OF fs_coname1]) -apply(subgoal_tac "\x'::coname. x'\(trm{coname2:=(y).Q},Q,coname1)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_OrR2 abs_fresh fresh_atm) -apply(rule exists_fresh'(2)[OF fs_coname1]) -apply(subgoal_tac "\x'::coname. x'\(trm{coname2:=(y).Q},Q,coname1)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_OrR2 abs_fresh fresh_atm) -apply(rule exists_fresh'(2)[OF fs_coname1]) -apply(subgoal_tac "\x'::coname. x'\(trm{coname2:=(y).Q},Q,coname1)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_ImpR abs_fresh abs_supp fin_supp fresh_atm) -apply(rule exists_fresh'(2)[OF fs_coname1]) -apply(subgoal_tac "\x'::coname. x'\(trm{coname2:=(y).Q},Q,coname1)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_ImpR abs_fresh abs_supp fin_supp fresh_atm) -apply(rule exists_fresh'(2)[OF fs_coname1]) -done - -lemma not_Ax2: - shows "\(x\M) \ M{x:=.Q} \ Ax y a" -apply(nominal_induct M avoiding: b y Q x a rule: trm.strong_induct) -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) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_NotL abs_fresh fresh_atm) -apply(rule exists_fresh'(1)[OF fs_name1]) -apply(subgoal_tac "\x'::name. x'\(trm{x:=.Q},Q)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_NotL abs_fresh fresh_atm) -apply(rule exists_fresh'(1)[OF fs_name1]) -apply(subgoal_tac "\x'::name. x'\(trm{x:=.Q},Q,name1)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_AndL1 abs_fresh fresh_atm) -apply(rule exists_fresh'(1)[OF fs_name1]) -apply(subgoal_tac "\x'::name. x'\(trm{x:=.Q},Q,name1)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_AndL1 abs_fresh fresh_atm) -apply(rule exists_fresh'(1)[OF fs_name1]) -apply(subgoal_tac "\x'::name. x'\(trm{x:=.Q},Q,name1)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_AndL2 abs_fresh fresh_atm) -apply(rule exists_fresh'(1)[OF fs_name1]) -apply(subgoal_tac "\x'::name. x'\(trm{x:=.Q},Q,name1)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_AndL2 abs_fresh fresh_atm) -apply(rule exists_fresh'(1)[OF fs_name1]) -apply(subgoal_tac "\x'::name. x'\(trm1{x:=.Q},Q,trm2{x:=.Q},name1,name2)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_OrL abs_fresh fresh_atm) -apply(rule exists_fresh'(1)[OF fs_name1]) -apply(subgoal_tac "\x'::name. x'\(trm1{x:=.Q},Q,trm2{x:=.Q},name1,name2)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_OrL abs_fresh fresh_atm) -apply(rule exists_fresh'(1)[OF fs_name1]) -apply(subgoal_tac "\x'::name. x'\(trm1{x:=.Q},Q,trm2{x:=.Q},name1,name2)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_OrL abs_fresh fresh_atm) -apply(rule exists_fresh'(1)[OF fs_name1]) -apply(subgoal_tac "\x'::name. x'\(trm1{name2:=.Q},Q,trm2{name2:=.Q},name1)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_ImpL abs_fresh fresh_atm) -apply(rule exists_fresh'(1)[OF fs_name1]) -apply(subgoal_tac "\x'::name. x'\(trm1{name2:=.Q},Q,trm2{name2:=.Q},name1)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_ImpL abs_fresh fresh_atm) -apply(rule exists_fresh'(1)[OF fs_name1]) -apply(subgoal_tac "\x'::name. x'\(trm1{name2:=.Q},Q,trm2{name2:=.Q},name1)") -apply(erule exE) -apply(simp add: fresh_prod) -apply(erule conjE)+ -apply(simp add: fresh_fun_simp_ImpL abs_fresh fresh_atm) -apply(rule exists_fresh'(1)[OF fs_name1]) -done +lemma not_Ax1: "\(b\M) \ M{b:=(y).Q} \ Ax x a" +proof (nominal_induct M avoiding: b y Q x a rule: trm.strong_induct) + case (NotR name trm coname) + then show ?case + by (metis fin.intros(1) fin_rest_elims(2) subst_not_fin2) +next + case (AndR coname1 trm1 coname2 trm2 coname3) + then show ?case + by (metis fin.intros(1) fin_rest_elims(3) subst_not_fin2) +next + case (OrR1 coname1 trm coname2) + then show ?case + by (metis fin.intros(1) fin_rest_elims(4) subst_not_fin2) +next + case (OrR2 coname1 trm coname2) + then show ?case + by (metis fin.intros(1) fin_rest_elims(5) subst_not_fin2) +next + case (ImpR name coname1 trm coname2) + then show ?case + by (metis fin.intros(1) fin_rest_elims(6) subst_not_fin2) +qed (auto simp: fresh_atm abs_fresh abs_supp fin_supp) + + + +lemma not_Ax2: "\(x\M) \ M{x:=.Q} \ Ax y a" +proof (nominal_induct M avoiding: b y Q x a rule: trm.strong_induct) + case (NotL coname trm name) + then show ?case + by (metis fic.intros(1) fic_rest_elims(2) not_fic_subst1) +next + case (AndL1 name1 trm name2) + then show ?case + by (metis fic.intros(1) fic_rest_elims(4) not_fic_subst1) +next + case (AndL2 name1 trm name2) + then show ?case + by (metis fic.intros(1) fic_rest_elims(5) not_fic_subst1) +next + case (OrL name1 trm1 name2 trm2 name3) + then show ?case + by (metis fic.intros(1) fic_rest_elims(3) not_fic_subst1) +next + case (ImpL coname trm1 name1 trm2 name2) + then show ?case + by (metis fic.intros(1) fic_rest_elims(6) not_fic_subst1) +qed (auto simp: fresh_atm abs_fresh abs_supp fin_supp) + lemma interesting_subst1: assumes a: "x\y" "x\P" "y\P" shows "N{y:=.P}{x:=.P} = N{x:=.Ax y c}{y:=.P}" -using a + using a proof(nominal_induct N avoiding: x y c P rule: trm.strong_induct) - case Ax - then show ?case - 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) - apply(rule trans) - apply(rule better_Cut_substn) - apply(simp add: abs_fresh) - apply(simp) - apply(simp) - apply(auto) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substn) - apply(simp add: abs_fresh) - apply(simp) - apply(simp) - apply(simp add: trm.inject alpha forget) - apply(rule trans) - apply(rule better_Cut_substn) - apply(simp add: abs_fresh) - apply(simp) - apply(simp) - apply(auto) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substn) - apply(simp add: abs_fresh) - apply(simp) - apply(simp) - apply(rule impI) - apply(simp add: trm.inject alpha forget) - apply(simp add: trm.inject alpha forget) - apply(rule trans) - apply(rule better_Cut_substn) - apply(simp add: abs_fresh) - apply(simp) - apply(simp) - apply(auto) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substn) - apply(simp add: abs_fresh) - apply(simp) - apply(simp) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substn) - apply(simp add: abs_fresh) - apply(simp) - apply(simp) - apply(simp add: trm.inject alpha forget) - apply(rule trans) - apply(rule better_Cut_substn) - apply(simp add: abs_fresh) - apply(simp) - apply(simp) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substn) - apply(simp add: abs_fresh) - apply(simp) - apply(simp) - apply(rule trans) - apply(rule better_Cut_substn) - apply(simp add: abs_fresh) - apply(simp) - apply(simp) - apply(auto) - apply(case_tac "y'\M") - apply(simp add: forget) - apply(simp add: not_Ax2) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substn) - apply(simp add: abs_fresh) - apply(simp) - apply(simp) - apply(auto) - apply(case_tac "x'\M") - apply(simp add: forget) - apply(simp add: not_Ax2) - done -next - case NotR - then show ?case - by (auto simp: abs_fresh fresh_atm forget) + apply (simp add: abs_fresh) + by (smt (verit) forget(1) not_Ax2) next case (NotL d M u) + obtain x'::name and x''::name + where "x' \ (P,M{y:=.P},M{x:=.Ax y c}{y:=.P},y,x)" + and "x''\ (P,M{x:=.Ax y c},M{x:=.Ax y c}{y:=.P},Ax y c,y,x)" + by (meson exists_fresh(1) fs_name1) then show ?case - 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)+ - apply(simp add: fresh_fun_simp_NotL) - apply(rule trans) - apply(rule better_Cut_substn) - apply(simp add: abs_fresh) - apply(simp) - apply(simp) - apply(auto simp: fresh_atm) - apply(simp add: trm.inject alpha forget) - apply(rule exists_fresh'(1)[OF fs_name1]) - apply(subgoal_tac "\x'::name. x'\(P,M{x:=.Ax y c},M{x:=.Ax y c}{y:=.P},Ax y c,y,x)") - apply(erule exE, simp only: fresh_prod) - apply(erule conjE)+ - apply(simp only: fresh_fun_simp_NotL) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substn) - apply(simp add: abs_fresh) - apply(simp) - apply(simp) - apply(simp add: trm.inject alpha forget subst_fresh) - apply(rule trans) - apply(rule substn.simps) - apply(simp add: abs_fresh fresh_prod fresh_atm) - apply(simp add: fresh_atm) - apply(rule exists_fresh'(1)[OF fs_name1]) - done -next - case (AndR d1 M d2 M' d3) - then show ?case - by (auto simp: abs_fresh fresh_atm forget trm.inject subst_fresh) + using NotL + by (auto simp: perm_fresh_fresh(1) swap_simps(3,7) trm.inject alpha forget + fresh_atm abs_fresh fresh_fun_simp_NotL fresh_prod) next case (AndL1 u M d) + obtain x'::name and x''::name + where "x' \ (P,M{y:=.P},M{x:=.Ax y c}{y:=.P},u,y,x)" + and "x''\ (P,Ax y c,M{x:=.Ax y c},M{x:=.Ax y c}{y:=.P},u,y,x)" + by (meson exists_fresh(1) fs_name1) then show ?case - 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)+ - apply(simp add: fresh_fun_simp_AndL1) - apply(rule trans) - apply(rule better_Cut_substn) - apply(simp add: abs_fresh) - apply(simp) - apply(simp) - apply(auto simp: fresh_atm) - apply(simp add: trm.inject alpha forget) - apply(rule exists_fresh'(1)[OF fs_name1]) - apply(subgoal_tac "\x'::name. x'\(P,Ax y c,M{x:=.Ax y c},M{x:=.Ax y c}{y:=.P},u,y,x)") - apply(erule exE, simp only: fresh_prod) - apply(erule conjE)+ - apply(simp only: fresh_fun_simp_AndL1) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substn) - apply(simp add: abs_fresh) - apply(simp) - apply(simp) - apply(auto simp: fresh_atm) - apply(rule exists_fresh'(1)[OF fs_name1]) - done + using AndL1 + by (auto simp: perm_fresh_fresh(1) swap_simps(3,7) trm.inject alpha forget + fresh_atm abs_fresh fresh_fun_simp_AndL1 fresh_prod) next case (AndL2 u M d) + obtain x'::name and x''::name + where "x' \ (P,M{y:=.P},M{x:=.Ax y c}{y:=.P},u,y,x)" + and "x''\ (P,Ax y c,M{x:=.Ax y c},M{x:=.Ax y c}{y:=.P},u,y,x)" + by (meson exists_fresh(1) fs_name1) then show ?case - 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)+ - apply(simp add: fresh_fun_simp_AndL2) - apply(rule trans) - apply(rule better_Cut_substn) - apply(simp add: abs_fresh) - apply(simp) - apply(simp) - apply(auto simp: fresh_atm) - apply(simp add: trm.inject alpha forget) - apply(rule exists_fresh'(1)[OF fs_name1]) - apply(subgoal_tac "\x'::name. x'\(P,Ax y c,M{x:=.Ax y c},M{x:=.Ax y c}{y:=.P},u,y,x)") - apply(erule exE, simp only: fresh_prod) - apply(erule conjE)+ - apply(simp only: fresh_fun_simp_AndL2) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substn) - apply(simp add: abs_fresh) - apply(simp) - apply(simp) - apply(auto simp: fresh_atm) - apply(rule exists_fresh'(1)[OF fs_name1]) - done -next - case OrR1 - then show ?case - by (auto simp: abs_fresh fresh_atm forget trm.inject subst_fresh) -next - case OrR2 - then show ?case - by (auto simp: abs_fresh fresh_atm forget trm.inject subst_fresh) + using AndL2 + by (auto simp: perm_fresh_fresh(1) swap_simps(3,7) trm.inject alpha forget + fresh_atm abs_fresh fresh_fun_simp_AndL2 fresh_prod) next case (OrL x1 M x2 M' x3) + obtain x'::name and x''::name + where "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)" + and "x''\ (P,Ax y c,M{x:=.Ax y c},M{x:=.Ax y c}{y:=.P}, + M'{x:=.Ax y c},M'{x:=.Ax y c}{y:=.P},x1,x2,x3,y,x)" + by (meson exists_fresh(1) fs_name1) then show ?case - 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) - apply(erule conjE)+ - apply(simp add: fresh_fun_simp_OrL) - apply(rule trans) - apply(rule better_Cut_substn) - apply(simp add: abs_fresh) - apply(simp) - apply(simp) - apply(auto simp: fresh_atm) - apply(simp add: trm.inject alpha forget) - apply(rule trans) - apply(rule substn.simps) - apply(simp add: abs_fresh subst_fresh fresh_atm) - apply(simp add: abs_fresh subst_fresh fresh_atm) - apply(force) - apply(simp) - apply(rule exists_fresh'(1)[OF fs_name1]) - apply(subgoal_tac "\x'::name. x'\(P,Ax y c,M{x:=.Ax y c},M{x:=.Ax y c}{y:=.P}, - M'{x:=.Ax y c},M'{x:=.Ax y c}{y:=.P},x1,x2,x3,y,x)") - apply(erule exE, simp only: fresh_prod) - apply(erule conjE)+ - apply(simp only: fresh_fun_simp_OrL) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substn) - apply(simp add: abs_fresh) - apply(simp) - apply(simp) - apply(simp add: trm.inject alpha) - apply(rule trans) - apply(rule substn.simps) - apply(simp add: abs_fresh subst_fresh fresh_atm) - apply(simp add: abs_fresh subst_fresh fresh_atm) - apply(force) - apply(simp) - apply(auto simp: fresh_atm) - apply(rule exists_fresh'(1)[OF fs_name1]) - done -next - case ImpR - then show ?case - by (auto simp: abs_fresh fresh_atm forget trm.inject subst_fresh) + using OrL + by (auto simp: perm_fresh_fresh(1) swap_simps(3,7) trm.inject alpha forget + fresh_atm abs_fresh fresh_fun_simp_OrL fresh_prod subst_fresh) next case (ImpL a M x1 M' x2) + obtain x'::name and x''::name + where "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)" + and "x''\ (P,Ax y c,M{x2:=.Ax y c},M{x2:=.Ax y c}{y:=.P}, + M'{x2:=.Ax y c},M'{x2:=.Ax y c}{y:=.P},x1,x2,y,x)" + by (meson exists_fresh(1) fs_name1) then show ?case - 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) - apply(erule conjE)+ - apply(simp add: fresh_fun_simp_ImpL) - apply(rule trans) - apply(rule better_Cut_substn) - apply(simp add: abs_fresh) - apply(simp) - apply(simp) - apply(auto simp: fresh_atm) - apply(simp add: trm.inject alpha forget) - apply(rule trans) - apply(rule substn.simps) - apply(simp add: abs_fresh subst_fresh fresh_atm) - apply(simp add: abs_fresh subst_fresh fresh_atm) - apply(force) - apply(rule exists_fresh'(1)[OF fs_name1]) - apply(subgoal_tac "\x'::name. x'\(P,Ax y c,M{x2:=.Ax y c},M{x2:=.Ax y c}{y:=.P}, - M'{x2:=.Ax y c},M'{x2:=.Ax y c}{y:=.P},x1,x2,y,x)") - apply(erule exE, simp only: fresh_prod) - apply(erule conjE)+ - apply(simp only: fresh_fun_simp_ImpL) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substn) - apply(simp add: abs_fresh) - apply(simp) - apply(simp) - apply(simp add: trm.inject alpha) - apply(rule trans) - apply(rule substn.simps) - apply(simp add: abs_fresh subst_fresh fresh_atm) - apply(simp add: abs_fresh subst_fresh fresh_atm) - apply(simp) - apply(auto simp: fresh_atm) - apply(rule exists_fresh'(1)[OF fs_name1]) - done -qed + using ImpL + by (auto simp: perm_fresh_fresh(1) swap_simps(3,7) trm.inject alpha forget + fresh_atm abs_fresh fresh_fun_simp_ImpL fresh_prod subst_fresh) +qed (auto simp: abs_fresh fresh_atm forget trm.inject subst_fresh) lemma interesting_subst1': - assumes a: "x\y" "x\P" "y\P" + assumes "x\y" "x\P" "y\P" shows "N{y:=.P}{x:=.P} = N{x:=.Ax y a}{y:=.P}" proof - show ?thesis proof (cases "c=a") - case True then show ?thesis using a by (simp add: interesting_subst1) + case True with assms show ?thesis + by (simp add: interesting_subst1) next - case False then show ?thesis using a - apply - - apply(subgoal_tac "N{x:=.Ax y a} = N{x:=.([(c,a)]\Ax y a)}") - apply(simp add: interesting_subst1 calc_atm) - apply(rule subst_rename) - apply(simp add: fresh_prod fresh_atm) - done + case False + then have "N{x:=.Ax y a} = N{x:=.([(c,a)]\Ax y a)}" + by (simp add: fresh_atm(2,4) fresh_prod substn_rename4) + with assms show ?thesis + by (simp add: interesting_subst1 swap_simps) qed qed lemma interesting_subst2: assumes a: "a\b" "a\P" "b\P" shows "N{a:=(y).P}{b:=(y).P} = N{b:=(y).Ax y a}{a:=(y).P}" -using a + using a proof(nominal_induct N avoiding: a b y P rule: trm.strong_induct) case Ax then show ?case @@ -5665,144 +4383,34 @@ next case (Cut d M u M' x' y' c P) with assms show ?case - apply(simp) apply(auto simp: trm.inject) - apply(rule trans) - apply(rule better_Cut_substc) - apply(simp) - apply(simp add: abs_fresh) - apply(simp add: forget) - apply(auto) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substc) - apply(simp add: abs_fresh) - apply(simp add: abs_fresh) - apply(simp) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substc) - apply(simp add: abs_fresh) - apply(simp add: abs_fresh) - apply(simp) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substc) - apply(simp add: abs_fresh) - apply(simp add: abs_fresh) - apply(simp) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substc) - apply(simp add: abs_fresh) - apply(simp add: abs_fresh) - apply(simp) - apply(rule trans) - apply(rule better_Cut_substc) - apply(simp add: abs_fresh) - apply(simp add: abs_fresh) - apply(auto)[1] - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substc) - apply(simp add: abs_fresh) - apply(simp add: abs_fresh) - apply(simp) - apply(rule impI) - apply(simp add: fresh_atm trm.inject alpha forget) - apply(case_tac "x'\M'") - apply(simp add: forget) - apply(simp add: not_Ax1) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substc) - apply(simp add: abs_fresh) - apply(simp add: abs_fresh) - apply(auto) - apply(case_tac "y'\M'") - apply(simp add: forget) - apply(simp add: not_Ax1) - done + apply (force simp add: abs_fresh forget) + apply (simp add: abs_fresh forget) + by (smt (verit, ccfv_threshold) abs_fresh(1) better_Cut_substc forget(2) fresh_prod not_Ax1) next case NotL then show ?case by (auto simp: abs_fresh fresh_atm forget) next case (NotR u M d) - then show ?case - 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)+ - apply(simp add: fresh_fun_simp_NotR) - apply(rule trans) - apply(rule better_Cut_substc) - apply(simp add: abs_fresh) - apply(simp add: abs_fresh) - apply(simp) - apply(auto simp: fresh_atm) - apply(simp add: trm.inject alpha forget) - apply(rule exists_fresh'(2)[OF fs_coname1]) - apply(subgoal_tac "\a'::coname. a'\(P,M{d:=(y).Ax y a},M{d:=(y).Ax y a}{a:=(y).P},Ax y a,y,d)") - apply(erule exE, simp only: fresh_prod) - apply(erule conjE)+ - apply(simp only: fresh_fun_simp_NotR) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substc) - apply(simp add: abs_fresh) - apply(simp add: abs_fresh) - apply(simp) - apply(simp add: trm.inject alpha forget subst_fresh) - apply(rule trans) - apply(rule substc.simps) - apply(simp add: abs_fresh fresh_prod fresh_atm) - apply(simp add: fresh_atm) - apply(rule exists_fresh'(2)[OF fs_coname1]) - done + obtain a' a''::coname + where "a' \ (b,P,M{d:=(y).P},M{b:=(y).Ax y d}{d:=(y).P},u,y)" + and "a'' \ (P,M{d:=(y).Ax y a},M{d:=(y).Ax y a}{a:=(y).P},Ax y a,y,d)" + by (meson exists_fresh'(2) fs_coname1) + with NotR show ?case + by (auto simp: abs_fresh fresh_atm forget fresh_prod fresh_fun_simp_NotR) next case (AndR d1 M d2 M' d3) - then show ?case + obtain a' a''::coname + where "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)" + and "a'' \ (P,Ax y a,M{d3:=(y).Ax y a},M{d3:=(y).Ax y a}{a:=(y).P}, + M'{d3:=(y).Ax y a},M'{d3:=(y).Ax y a}{a:=(y).P},d1,d2,d3,y,b)" + by (meson exists_fresh'(2) fs_coname1) + with AndR show ?case 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) - apply(erule conjE)+ - apply(simp add: fresh_fun_simp_AndR) - apply(rule trans) - apply(rule better_Cut_substc) - apply(simp add: abs_fresh fresh_atm) - apply(simp add: abs_fresh fresh_atm) - apply(simp) - apply(auto simp: fresh_atm) - apply(simp add: trm.inject alpha forget) - apply(rule trans) - apply(rule substc.simps) - apply(simp add: abs_fresh subst_fresh fresh_atm) - apply(simp add: abs_fresh subst_fresh fresh_atm) - apply(force) - apply(simp) - apply(rule exists_fresh'(2)[OF fs_coname1]) - apply(subgoal_tac "\a'::coname. a'\(P,Ax y a,M{d3:=(y).Ax y a},M{d3:=(y).Ax y a}{a:=(y).P}, - M'{d3:=(y).Ax y a},M'{d3:=(y).Ax y a}{a:=(y).P},d1,d2,d3,y,b)") - apply(erule exE, simp only: fresh_prod) - apply(erule conjE)+ - apply(simp only: fresh_fun_simp_AndR) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substc) - apply(simp add: abs_fresh) - apply(simp add: abs_fresh) - apply(simp) - apply(simp add: trm.inject alpha) - apply(rule trans) - apply(rule substc.simps) - apply(simp add: abs_fresh subst_fresh fresh_atm) - apply(simp add: abs_fresh subst_fresh fresh_atm) - apply(force) - apply(simp) - apply(auto simp: fresh_atm) - apply(rule exists_fresh'(2)[OF fs_coname1]) + apply(auto simp only: fresh_prod fresh_fun_simp_AndR) + apply (force simp: forget abs_fresh subst_fresh fresh_atm)+ done next case (AndL1 u M d) @@ -5814,70 +4422,20 @@ by (auto simp: abs_fresh fresh_atm forget trm.inject subst_fresh) next case (OrR1 d M e) - then show ?case - 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)+ - apply(simp add: fresh_fun_simp_OrR1) - apply(rule trans) - apply(rule better_Cut_substc) - apply(simp add: abs_fresh) - apply(simp add: abs_fresh) - apply(simp) - apply(auto simp: fresh_atm) - apply(simp add: trm.inject alpha forget) - apply(rule exists_fresh'(2)[OF fs_coname1]) - apply(subgoal_tac "\a'::coname. a'\(b,P,Ax y a,M{e:=(y).Ax y a},M{e:=(y).Ax y a}{a:=(y).P},d,e)") - apply(erule exE, simp only: fresh_prod) - apply(erule conjE)+ - apply(simp only: fresh_fun_simp_OrR1) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substc) - apply(simp add: abs_fresh) - apply(simp add: abs_fresh) - apply(simp) - apply(simp add: trm.inject alpha forget subst_fresh) - apply(rule trans) - apply(rule substc.simps) - apply(simp add: abs_fresh fresh_prod fresh_atm) - apply(simp add: fresh_atm) - apply(rule exists_fresh'(2)[OF fs_coname1]) - done + obtain a' a''::coname + where "a' \ (b,P,M{e:=(y).P},M{b:=(y).Ax y e}{e:=(y).P},d,e)" + and "a'' \ (b,P,Ax y a,M{e:=(y).Ax y a},M{e:=(y).Ax y a}{a:=(y).P},d,e)" + by (meson exists_fresh'(2) fs_coname1) + with OrR1 show ?case + by (auto simp: abs_fresh fresh_atm forget fresh_prod fresh_fun_simp_OrR1) next case (OrR2 d M e) - then show ?case - 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)+ - apply(simp add: fresh_fun_simp_OrR2) - apply(rule trans) - apply(rule better_Cut_substc) - apply(simp add: abs_fresh) - apply(simp add: abs_fresh) - apply(simp) - apply(auto simp: fresh_atm) - apply(simp add: trm.inject alpha forget) - apply(rule exists_fresh'(2)[OF fs_coname1]) - apply(subgoal_tac "\a'::coname. a'\(b,P,Ax y a,M{e:=(y).Ax y a},M{e:=(y).Ax y a}{a:=(y).P},d,e)") - apply(erule exE, simp only: fresh_prod) - apply(erule conjE)+ - apply(simp only: fresh_fun_simp_OrR2) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substc) - apply(simp add: abs_fresh) - apply(simp add: abs_fresh) - apply(simp) - apply(simp add: trm.inject alpha forget subst_fresh) - apply(rule trans) - apply(rule substc.simps) - apply(simp add: abs_fresh fresh_prod fresh_atm) - apply(simp add: fresh_atm) - apply(rule exists_fresh'(2)[OF fs_coname1]) - done + obtain a' a''::coname + where "a' \ (b,P,M{e:=(y).P},M{b:=(y).Ax y e}{e:=(y).P},d,e)" + and "a'' \ (b,P,Ax y a,M{e:=(y).Ax y a},M{e:=(y).Ax y a}{a:=(y).P},d,e)" + by (meson exists_fresh'(2) fs_coname1) + with OrR2 show ?case + by (auto simp: abs_fresh fresh_atm forget fresh_prod fresh_fun_simp_OrR2) next case (OrL x1 M x2 M' x3) then show ?case @@ -5888,352 +4446,108 @@ 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: abs_fresh fresh_atm forget trm.inject subst_fresh) - apply(subgoal_tac "\a'::coname. a'\(b,e,d,P,M{d:=(y).P},M{b:=(y).Ax y d}{d:=(y).P})") - apply(erule exE, simp add: fresh_prod) - apply(erule conjE)+ - apply(simp add: fresh_fun_simp_ImpR) - apply(rule trans) - apply(rule better_Cut_substc) - apply(simp add: abs_fresh) - apply(simp add: abs_fresh) - apply(simp) - apply(auto simp: fresh_atm) - apply(simp add: trm.inject alpha forget) - apply(rule exists_fresh'(2)[OF fs_coname1]) - apply(subgoal_tac "\a'::coname. a'\(e,d,P,Ax y a,M{d:=(y).Ax y a},M{d:=(y).Ax y a}{a:=(y).P})") - apply(erule exE, simp only: fresh_prod) - apply(erule conjE)+ - apply(simp only: fresh_fun_simp_ImpR) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substc) - apply(simp add: abs_fresh) - apply(simp add: abs_fresh) - apply(simp) - apply(simp add: trm.inject alpha) - apply(rule trans) - apply(rule substc.simps) - apply(simp add: abs_fresh subst_fresh fresh_atm) - apply(simp add: abs_fresh subst_fresh fresh_atm) - apply(simp) - apply(auto simp: fresh_atm) - apply(rule exists_fresh'(2)[OF fs_coname1]) - done + obtain a' a''::coname + where "a' \ (b,e,d,P,M{d:=(y).P},M{b:=(y).Ax y d}{d:=(y).P})" + and "a'' \ (e,d,P,Ax y a,M{d:=(y).Ax y a},M{d:=(y).Ax y a}{a:=(y).P})" + by (meson exists_fresh'(2) fs_coname1) + with ImpR show ?case + by (auto simp: abs_fresh fresh_atm forget fresh_prod fresh_fun_simp_ImpR) qed lemma interesting_subst2': - assumes a: "a\b" "a\P" "b\P" + assumes "a\b" "a\P" "b\P" shows "N{a:=(y).P}{b:=(y).P} = N{b:=(z).Ax z a}{a:=(y).P}" proof - show ?thesis proof (cases "z=y") - case True then show ?thesis using a by (simp add: interesting_subst2) + case True then show ?thesis using assms by (simp add: interesting_subst2) next - case False then show ?thesis using a - apply - - apply(subgoal_tac "N{b:=(z).Ax z a} = N{b:=(y).([(y,z)]\Ax z a)}") - apply(simp add: interesting_subst2 calc_atm) - apply(rule subst_rename) - apply(simp add: fresh_prod fresh_atm) - done + case False + then have "N{b:=(z).Ax z a} = N{b:=(y).([(y,z)]\Ax z a)}" + by (metis fresh_atm(1,3) fresh_prod substc_rename2 trm.fresh(1)) + with False assms show ?thesis + by (simp add: interesting_subst2 calc_atm) qed qed + lemma subst_subst1: assumes a: "a\(Q,b)" "x\(y,P,Q)" "b\Q" "y\P" shows "M{x:=.P}{b:=(y).Q} = M{b:=(y).Q}{x:=.(P{b:=(y).Q})}" -using a + using a proof(nominal_induct M avoiding: x a P b y Q rule: trm.strong_induct) case (Ax z c) - have fs: "a\(Q,b)" "x\(y,P,Q)" "b\Q" "y\P" by fact+ - { assume asm: "z=x \ c=b" - have "(Ax x b){x:=.P}{b:=(y).Q} = (Cut .P (x).Ax x b){b:=(y).Q}" using fs by simp - also have "\ = Cut .(P{b:=(y).Q}) (y).Q" - using fs by (simp_all add: fresh_prod fresh_atm) - also have "\ = Cut .(P{b:=(y).Q}) (y).(Q{x:=.(P{b:=(y).Q})})" using fs by (simp add: forget) - also have "\ = (Cut .Ax x b (y).Q){x:=.(P{b:=(y).Q})}" - using fs asm by (auto simp: fresh_prod fresh_atm subst_fresh) - also have "\ = (Ax x b){b:=(y).Q}{x:=.(P{b:=(y).Q})}" using fs by simp - finally have "(Ax z c){x:=.P}{b:=(y).Q} = (Ax z c){b:=(y).Q}{x:=.(P{b:=(y).Q})}" - using asm by simp - } - moreover - { assume asm: "z\x \ c=b" - have "(Ax z c){x:=.P}{b:=(y).Q} = (Ax z c){b:=(y).Q}" using asm by simp - also have "\ = Cut .Ax z c (y).Q" using fs asm by simp - also have "\ = Cut .(Ax z c{x:=.(P{b:=(y).Q})}) (y).(Q{x:=.(P{b:=(y).Q})})" - using fs asm by (simp add: forget) - also have "\ = (Cut .Ax z c (y).Q){x:=.(P{b:=(y).Q})}" using asm fs - by (auto simp: 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: trm.inject abs_fresh) - also have "\ = (Ax z c){x:=.(P{b:=(y).Q})}" using fs asm by simp - also have "\ = (Ax z c){b:=(y).Q}{x:=.(P{b:=(y).Q})}" using asm by auto - finally have "(Ax z c){x:=.P}{b:=(y).Q} = (Ax z c){b:=(y).Q}{x:=.(P{b:=(y).Q})}" by simp - } - moreover - { assume asm: "z\x \ c\b" - have "(Ax z c){x:=.P}{b:=(y).Q} = (Ax z c){b:=(y).Q}{x:=.(P{b:=(y).Q})}" using asm by auto - } - ultimately show ?case by blast + then show ?case + by (auto simp add: abs_fresh fresh_prod fresh_atm subst_fresh trm.inject forget) next case (Cut c M z N) - { assume asm: "M = Ax x c \ N = Ax z b" - have "(Cut .M (z).N){x:=.P}{b:=(y).Q} = (Cut .P (z).(N{x:=.P})){b:=(y).Q}" - using 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: 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: 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 - have "(Cut .M (z).N){x:=.P}{b:=(y).Q} = (Cut .M (z).N){b:=(y).Q}{x:=.(P{b:=(y).Q})}" - using eq1 eq2 by simp - } - moreover - { assume asm: "M \ Ax x c \ N = Ax z b" - have neq: "M{b:=(y).Q} \ Ax x c" - proof (cases "b\M") - case True then show ?thesis using asm by (simp add: forget) - next - case False then show ?thesis by (simp add: not_Ax1) - qed - have "(Cut .M (z).N){x:=.P}{b:=(y).Q} = (Cut .(M{x:=.P}) (z).(N{x:=.P})){b:=(y).Q}" - using Cut asm by simp - also have "\ = (Cut .(M{x:=.P}) (z).N){b:=(y).Q}" using Cut asm by (simp add: fresh_atm) - also have "\ = Cut .(M{x:=.P}{b:=(y).Q}) (y).Q" using Cut asm by (simp add: abs_fresh) - also have "\ = Cut .(M{b:=(y).Q}{x:=.P{b:=(y).Q}}) (y).Q" using Cut asm by simp - finally - have eq1: "(Cut .M (z).N){x:=.P}{b:=(y).Q} = Cut .(M{b:=(y).Q}{x:=.P{b:=(y).Q}}) (y).Q" - by simp - have "(Cut .M (z).N){b:=(y).Q}{x:=.(P{b:=(y).Q})} = - (Cut .(M{b:=(y).Q}) (y).Q){x:=.(P{b:=(y).Q})}" using 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: 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 - have "(Cut .M (z).N){x:=.P}{b:=(y).Q} = (Cut .M (z).N){b:=(y).Q}{x:=.(P{b:=(y).Q})}" - using eq1 eq2 by simp - } - moreover - { assume asm: "M = Ax x c \ N \ Ax z b" - have neq: "N{x:=.P} \ Ax z b" - proof (cases "x\N") - case True then show ?thesis using asm by (simp add: forget) - next - case False then show ?thesis by (simp add: not_Ax2) - qed - have "(Cut .M (z).N){x:=.P}{b:=(y).Q} = (Cut .P (z).(N{x:=.P})){b:=(y).Q}" - using Cut asm by simp - also have "\ = Cut .(P{b:=(y).Q}) (z).(N{x:=.P}{b:=(y).Q})" using Cut asm neq - by (simp add: abs_fresh) - also have "\ = Cut .(P{b:=(y).Q}) (z).(N{b:=(y).Q}{x:=.(P{b:=(y).Q})})" using Cut asm by simp - finally have eq1: "(Cut .M (z).N){x:=.P}{b:=(y).Q} - = Cut .(P{b:=(y).Q}) (z).(N{b:=(y).Q}{x:=.(P{b:=(y).Q})})" by simp - have "(Cut .M (z).N){b:=(y).Q}{x:=.(P{b:=(y).Q})} - = (Cut .(M{b:=(y).Q}) (z).(N{b:=(y).Q})){x:=.(P{b:=(y).Q})}" - using Cut asm by auto - also have "\ = (Cut .M (z).(N{b:=(y).Q})){x:=.(P{b:=(y).Q})}" - 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 - have eq2: "(Cut .M (z).N){b:=(y).Q}{x:=.(P{b:=(y).Q})} - = Cut .(P{b:=(y).Q}) (z).(N{b:=(y).Q}{x:=.(P{b:=(y).Q})})" by simp - have "(Cut .M (z).N){x:=.P}{b:=(y).Q} = (Cut .M (z).N){b:=(y).Q}{x:=.(P{b:=(y).Q})}" - using eq1 eq2 by simp - } - moreover - { assume asm: "M \ Ax x c \ N \ Ax z b" - have neq1: "N{x:=.P} \ Ax z b" - proof (cases "x\N") - case True then show ?thesis using asm by (simp add: forget) - next - case False then show ?thesis by (simp add: not_Ax2) - qed - have neq2: "M{b:=(y).Q} \ Ax x c" - proof (cases "b\M") - case True then show ?thesis using asm by (simp add: forget) - next - case False then show ?thesis by (simp add: not_Ax1) - qed - have "(Cut .M (z).N){x:=.P}{b:=(y).Q} = (Cut .(M{x:=.P}) (z).(N{x:=.P})){b:=(y).Q}" - using Cut asm by simp - also have "\ = Cut .(M{x:=.P}{b:=(y).Q}) (z).(N{x:=.P}{b:=(y).Q})" using Cut asm neq1 - by (simp add: abs_fresh) - also have "\ = Cut .(M{b:=(y).Q}{x:=.(P{b:=(y).Q})}) (z).(N{b:=(y).Q}{x:=.(P{b:=(y).Q})})" - using Cut asm by simp - finally have eq1: "(Cut .M (z).N){x:=.P}{b:=(y).Q} - = Cut .(M{b:=(y).Q}{x:=.(P{b:=(y).Q})}) (z).(N{b:=(y).Q}{x:=.(P{b:=(y).Q})})" by simp - have "(Cut .M (z).N){b:=(y).Q}{x:=.(P{b:=(y).Q})} = - (Cut .(M{b:=(y).Q}) (z).(N{b:=(y).Q})){x:=.(P{b:=(y).Q})}" using Cut asm neq1 by simp - also have "\ = Cut .(M{b:=(y).Q}{x:=.(P{b:=(y).Q})}) (z).(N{b:=(y).Q}{x:=.(P{b:=(y).Q})})" - using Cut asm neq2 by (simp add: fresh_prod fresh_atm subst_fresh) - finally have eq2: "(Cut .M (z).N){b:=(y).Q}{x:=.(P{b:=(y).Q})} = - Cut .(M{b:=(y).Q}{x:=.(P{b:=(y).Q})}) (z).(N{b:=(y).Q}{x:=.(P{b:=(y).Q})})" by simp - have "(Cut .M (z).N){x:=.P}{b:=(y).Q} = (Cut .M (z).N){b:=(y).Q}{x:=.(P{b:=(y).Q})}" - using eq1 eq2 by simp - } - ultimately show ?case by blast + then show ?case + apply (clarsimp simp add: abs_fresh fresh_prod fresh_atm subst_fresh) + apply (metis forget not_Ax1 not_Ax2) + done next case (NotR z M c) then show ?case - 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) - apply(erule conjE)+ - apply(simp add: fresh_fun_simp_NotR abs_fresh fresh_atm) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substn) - apply(simp add: fresh_prod fresh_atm subst_fresh abs_fresh) - apply(simp add: fresh_prod fresh_atm subst_fresh abs_fresh) - apply(simp add: forget) - apply(simp add: fresh_prod fresh_atm subst_fresh) - apply(rule exists_fresh'(2)[OF fs_coname1]) - done + by (auto simp: forget fresh_prod fresh_atm subst_fresh abs_fresh fresh_fun_simp_NotR) next case (NotL c M z) - then show ?case - 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) - apply(erule conjE)+ - apply(simp add: fresh_fun_simp_NotL abs_fresh fresh_atm) - apply(rule exists_fresh'(1)[OF fs_name1]) - done + obtain x'::name where "x' \ (P,M{x:=.P},P{b:=(y).Q},M{b:=(y).Q}{x:=.P{b:=(y).Q}},y,Q)" + by (meson exists_fresh(1) fs_name1) + with NotL show ?case + apply(auto simp: fresh_prod fresh_atm abs_fresh subst_fresh subst_fresh fresh_fun_simp_NotL) + by (metis fresh_fun_simp_NotL fresh_prod subst_fresh(5)) next case (AndR c1 M c2 N c3) then show ?case - 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) - apply(simp add: fresh_prod) - apply(erule conjE)+ - apply(simp add: fresh_fun_simp_AndR abs_fresh fresh_atm) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substn) - apply(simp_all add: fresh_atm abs_fresh subst_fresh) - apply(simp add: forget) - apply(rule exists_fresh'(2)[OF fs_coname1]) - done + by (auto simp: forget fresh_prod fresh_atm subst_fresh abs_fresh fresh_fun_simp_AndR) next case (AndL1 z1 M z2) - then show ?case - 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) - apply(erule conjE)+ - apply(simp add: fresh_fun_simp_AndL1 abs_fresh fresh_atm) - apply(rule exists_fresh'(1)[OF fs_name1]) - done + obtain x'::name where "x' \ (P,M{x:=.P},P{b:=(y).Q},z1,y,Q,M{b:=(y).Q}{x:=.P{b:=(y).Q}})" + by (meson exists_fresh(1) fs_name1) + with AndL1 show ?case + apply(auto simp: fresh_prod fresh_atm abs_fresh subst_fresh fresh_fun_simp_AndL1) + by (metis fresh_atm(1) fresh_fun_simp_AndL1 fresh_prod subst_fresh(5)) next case (AndL2 z1 M z2) - then show ?case - 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) - apply(erule conjE)+ - apply(simp add: fresh_fun_simp_AndL2 abs_fresh fresh_atm) - apply(rule exists_fresh'(1)[OF fs_name1]) - done + obtain x'::name where "x' \ (P,M{x:=.P},P{b:=(y).Q},z1,y,Q,M{b:=(y).Q}{x:=.P{b:=(y).Q}})" + by (meson exists_fresh(1) fs_name1) + with AndL2 show ?case + apply(auto simp: fresh_prod fresh_atm abs_fresh subst_fresh fresh_fun_simp_AndL2) + by (metis fresh_atm(1) fresh_fun_simp_AndL2 fresh_prod subst_fresh(5)) next case (OrL z1 M z2 N z3) - then show ?case - 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) - apply(simp add: fresh_prod) - apply(erule conjE)+ - apply(simp add: fresh_fun_simp_OrL abs_fresh fresh_atm) - apply(simp add: trm.inject alpha) - apply(rule trans) - apply(rule substc.simps) - apply(simp_all add: fresh_atm subst_fresh) - apply(rule exists_fresh'(1)[OF fs_name1]) - done + obtain x'::name where "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)" + by (meson exists_fresh(1) fs_name1) + with OrL show ?case + apply(auto simp: fresh_prod fresh_atm abs_fresh subst_fresh fresh_fun_simp_OrL) + by (metis (full_types) fresh_atm(1) fresh_fun_simp_OrL fresh_prod subst_fresh(5)) next case (OrR1 c1 M c2) then show ?case - 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) - apply(simp add: fresh_prod) - apply(erule conjE)+ - apply(simp add: fresh_fun_simp_OrR1 abs_fresh fresh_atm) - apply(simp_all add: fresh_atm subst_fresh abs_fresh) - apply(simp add: forget) - apply(rule exists_fresh'(2)[OF fs_coname1]) - done + by (auto simp: forget fresh_prod fresh_atm subst_fresh abs_fresh fresh_fun_simp_OrR1) next case (OrR2 c1 M c2) then show ?case - 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) - apply(simp add: fresh_prod) - apply(erule conjE)+ - apply(simp add: fresh_fun_simp_OrR2 abs_fresh fresh_atm) - apply(simp_all add: fresh_atm subst_fresh abs_fresh) - apply(simp add: forget) - apply(rule exists_fresh'(2)[OF fs_coname1]) - done + by (auto simp: forget fresh_prod fresh_atm subst_fresh abs_fresh fresh_fun_simp_OrR2) next case (ImpR z c M d) then show ?case - 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) - apply(simp add: fresh_prod) - apply(erule conjE)+ - apply(simp add: fresh_fun_simp_ImpR abs_fresh fresh_atm) - apply(simp_all add: fresh_atm subst_fresh forget abs_fresh) - apply(rule exists_fresh'(2)[OF fs_coname1]) - done + by (auto simp: forget fresh_prod fresh_atm subst_fresh abs_fresh fresh_fun_simp_ImpR) next case (ImpL c M z N u) - then show ?case - 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) - apply(simp add: fresh_prod) - apply(erule conjE)+ - apply(simp add: fresh_fun_simp_ImpL abs_fresh fresh_atm) - apply(simp add: trm.inject alpha) - apply(rule trans) - apply(rule substc.simps) - apply(simp_all add: fresh_atm subst_fresh forget) - apply(rule exists_fresh'(1)[OF fs_name1]) - done + obtain z'::name where "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)" + by (meson exists_fresh(1) fs_name1) + with ImpL show ?case + apply(auto simp: fresh_prod fresh_atm abs_fresh subst_fresh fresh_fun_simp_ImpL) + by (metis (full_types) fresh_atm(1) fresh_prod subst_fresh(5) fresh_fun_simp_ImpL) qed lemma subst_subst2: assumes a: "a\(b,P,N)" "x\(y,P,M)" "b\(M,N)" "y\P" shows "M{a:=(x).N}{y:=.P} = M{y:=.P}{a:=(x).N{y:=.P}}" -using a + using a proof(nominal_induct M avoiding: a x N y b P rule: trm.strong_induct) case (Ax z c) then show ?case @@ -6241,142 +4555,54 @@ next case (Cut d M' u M'') then show ?case - apply(simp add: fresh_atm fresh_prod trm.inject abs_fresh) - apply(auto) - apply(simp add: fresh_atm) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substc) - apply(simp add: abs_fresh subst_fresh fresh_prod fresh_atm) - apply(simp add: fresh_prod subst_fresh fresh_atm abs_fresh) - apply(simp) - apply(simp add: forget) - apply(simp add: fresh_atm) - apply(case_tac "a\M'") - apply(simp add: forget) - apply(simp add: not_Ax1) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substc) - apply(simp add: abs_fresh subst_fresh fresh_prod fresh_atm) - apply(simp add: fresh_prod subst_fresh fresh_atm abs_fresh) - apply(auto)[1] - apply(case_tac "y\M''") - apply(simp add: forget) - apply(simp add: not_Ax2) - apply(simp add: forget) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substc) - apply(simp add: subst_fresh fresh_atm) - apply(simp add: abs_fresh subst_fresh) - apply(auto)[1] - apply(case_tac "y\M''") - apply(simp add: forget) - apply(simp add: not_Ax2) - apply(case_tac "a\M'") - apply(simp add: forget) - apply(simp add: not_Ax1) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substc) - apply(simp add: subst_fresh) - apply(simp add: subst_fresh abs_fresh) - apply(simp) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substc) - apply(simp add: subst_fresh fresh_atm) - apply(simp add: subst_fresh abs_fresh) - apply(auto)[1] - apply(case_tac "y\M''") - apply(simp add: forget) - apply(simp add: not_Ax2) + apply (clarsimp simp add: abs_fresh fresh_prod fresh_atm subst_fresh) + apply (metis forget not_Ax1 not_Ax2) done next case (NotR z M' d) - then show ?case - 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)+ - apply(simp add: fresh_fun_simp_NotR) - apply(rule trans) - apply(rule better_Cut_substn) - apply(simp add: abs_fresh subst_fresh) - apply(simp add: fresh_prod subst_fresh fresh_atm) - apply(simp) - apply(simp add: trm.inject alpha) - apply(rule trans) - apply(rule substn.simps) - apply(simp add: fresh_prod fresh_atm) - apply(simp add: fresh_atm) - apply(rule exists_fresh'(2)[OF fs_coname1]) + obtain a'::coname where "a' \ (y,P,N,N{y:=.P},M'{d:=(x).N},M'{y:=.P}{d:=(x).N{y:=.P}})" + by (meson exists_fresh'(2) fs_coname1) + with NotR show ?case + apply(simp add: fresh_atm subst_fresh) + apply(auto simp only: fresh_prod fresh_fun_simp_NotR; simp add: abs_fresh NotR) done next case (NotL d M' z) - then show ?case - 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)+ - apply(simp add: fresh_fun_simp_NotL) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substc) - apply(simp add: abs_fresh subst_fresh fresh_atm) - apply(simp add: fresh_prod subst_fresh fresh_atm abs_fresh) - apply(simp) - apply(simp add: trm.inject alpha forget) - apply(rule trans) - apply(rule substc.simps) - apply(simp add: fresh_prod fresh_atm) - apply(simp add: fresh_atm subst_fresh) - apply(simp) - apply(rule exists_fresh'(1)[OF fs_name1]) - done + obtain x'::name where "x' \ (z,y,P,N,N{y:=.P},M'{y:=.P},M'{y:=.P}{a:=(x).N{y:=.P}})" + by (meson exists_fresh(1) fs_name1) + with NotL show ?case + apply(simp add: fresh_prod fresh_atm abs_fresh) + apply(auto simp only: fresh_fun_simp_NotL) + by (auto simp: subst_fresh abs_fresh fresh_atm forget) next case (AndR d M' e M'' f) - then show ?case - apply(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget trm.inject) + obtain a'::coname where "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}})" + by (meson exists_fresh'(2) fs_coname1) + with AndR show ?case + apply(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod) apply(subgoal_tac "\a'::coname. a'\(P,b,d,e,N,N{y:=.P},M'{f:=(x).N},M''{f:=(x).N}, M'{y:=.P}{f:=(x).N{y:=.P}},M''{y:=.P}{f:=(x).N{y:=.P}})") - apply(erule exE, simp add: fresh_prod) - apply(erule conjE)+ - apply(simp add: fresh_fun_simp_AndR) - apply(rule trans) - apply(rule better_Cut_substn) - apply(simp add: abs_fresh subst_fresh fresh_atm) - apply(simp add: fresh_prod subst_fresh fresh_atm abs_fresh) - apply(simp) - apply(simp add: trm.inject alpha forget) - apply(rule trans) - apply(rule substn.simps) - apply(simp add: fresh_prod fresh_atm subst_fresh) - apply(simp add: fresh_atm subst_fresh) - apply(simp add: fresh_atm) - apply(simp) + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_AndR) + apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp) apply(rule exists_fresh'(2)[OF fs_coname1]) done next case (AndL1 z M' u) - then show ?case - apply(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget trm.inject) + obtain x'::name where "x' \ (P,b,z,u,x,N,M'{y:=.P},M'{y:=.P}{a:=(x).N{y:=.P}})" + by (meson exists_fresh(1) fs_name1) + with AndL1 show ?case + apply(simp add: fresh_prod fresh_atm abs_fresh) + apply(auto simp only: fresh_fun_simp_AndL1) + apply (auto simp: subst_fresh abs_fresh fresh_atm forget) + apply(subgoal_tac "\x'::name. x'\(P,b,z,u,x,N,M'{y:=.P},M'{y:=.P}{a:=(x).N{y:=.P}})") - apply(erule exE, simp add: fresh_prod) - apply(erule conjE)+ - apply(simp add: fresh_fun_simp_AndL1) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substc) - apply(simp add: abs_fresh subst_fresh fresh_atm) - apply(simp add: fresh_prod subst_fresh fresh_atm abs_fresh) - apply(simp) - apply(simp add: trm.inject alpha forget) - apply(rule trans) - apply(rule substc.simps) - apply(simp add: fresh_prod fresh_atm subst_fresh) - apply(simp) + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_AndL1) + apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp) apply(rule exists_fresh'(1)[OF fs_name1]) done next @@ -6384,20 +4610,10 @@ then show ?case 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)+ - apply(simp add: fresh_fun_simp_AndL2) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substc) - apply(simp add: abs_fresh subst_fresh fresh_atm) - apply(simp add: fresh_prod subst_fresh fresh_atm abs_fresh) - apply(simp) - apply(simp add: trm.inject alpha forget) - apply(rule trans) - apply(rule substc.simps) - apply(simp add: fresh_prod fresh_atm subst_fresh) - apply(simp) + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_AndL2) + apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp) apply(rule exists_fresh'(1)[OF fs_name1]) done next @@ -6406,22 +4622,10 @@ 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) - apply(erule conjE)+ - apply(simp add: fresh_fun_simp_OrL) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substc) - apply(simp add: abs_fresh subst_fresh fresh_atm) - apply(simp add: fresh_prod subst_fresh fresh_atm abs_fresh) - apply(simp) - apply(simp add: trm.inject alpha forget) - apply(rule trans) - apply(rule substc.simps) - apply(simp add: fresh_prod fresh_atm subst_fresh) - apply(simp add: fresh_atm subst_fresh) - apply(simp add: fresh_atm) - apply(simp) + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_OrL) + apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp) apply(rule exists_fresh'(1)[OF fs_name1]) done next @@ -6430,19 +4634,10 @@ 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) - apply(erule conjE)+ - apply(simp add: fresh_fun_simp_OrR1) - apply(rule trans) - apply(rule better_Cut_substn) - apply(simp add: abs_fresh subst_fresh fresh_atm) - apply(simp add: fresh_prod subst_fresh fresh_atm abs_fresh) - apply(simp) - apply(simp add: trm.inject alpha forget) - apply(rule trans) - apply(rule substn.simps) - apply(simp add: fresh_prod fresh_atm subst_fresh) - apply(simp) + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_OrR1) + apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp) apply(rule exists_fresh'(2)[OF fs_coname1]) done next @@ -6451,42 +4646,24 @@ 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) - apply(erule conjE)+ - apply(simp add: fresh_fun_simp_OrR2) - apply(rule trans) - apply(rule better_Cut_substn) - apply(simp add: abs_fresh subst_fresh fresh_atm) - apply(simp add: fresh_prod subst_fresh fresh_atm abs_fresh) - apply(simp) - apply(simp add: trm.inject alpha forget) - apply(rule trans) - apply(rule substn.simps) - apply(simp add: fresh_prod fresh_atm subst_fresh) - apply(simp) + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_OrR2) + apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp) apply(rule exists_fresh'(2)[OF fs_coname1]) done next case (ImpR x e M' f) then show ?case 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}, + apply(subgoal_tac "\c'::coname. c'\(P,b,e,f,x,N,N{y:=.P}, M'{f:=(x).N},M'{y:=.P}{f:=(x).N{y:=.P}})") - apply(erule exE, simp add: fresh_prod) - apply(erule conjE)+ - apply(simp add: fresh_fun_simp_ImpR) - apply(rule trans) - apply(rule better_Cut_substn) - apply(simp add: abs_fresh subst_fresh fresh_atm) - apply(simp add: fresh_prod subst_fresh fresh_atm abs_fresh) - apply(simp) - apply(simp add: trm.inject alpha forget) - apply(rule trans) - apply(rule substn.simps) - apply(simp add: fresh_prod fresh_atm subst_fresh) - apply(simp add: fresh_atm) - apply(simp add: fresh_atm trm.inject alpha abs_fresh fin_supp abs_supp) - apply(rule exists_fresh'(2)[OF fs_coname1]) + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_ImpR) + apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp) + apply(simp add: fresh_atm trm.inject alpha abs_fresh fin_supp abs_supp) + apply(rule exists_fresh'(2)[OF fs_coname1]) apply(simp add: fresh_atm trm.inject alpha abs_fresh fin_supp abs_supp) done next @@ -6495,21 +4672,10 @@ 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) - apply(erule conjE)+ - apply(simp add: fresh_fun_simp_ImpL) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substc) - apply(simp add: abs_fresh subst_fresh fresh_atm) - apply(simp add: fresh_prod subst_fresh fresh_atm abs_fresh) - apply(simp) - apply(simp add: trm.inject alpha forget) - apply(rule trans) - apply(rule substc.simps) - apply(simp add: fresh_prod fresh_atm subst_fresh) - apply(simp add: fresh_atm subst_fresh) - apply(simp add: fresh_atm) + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_ImpL) + apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp) apply(rule exists_fresh'(1)[OF fs_name1]) done qed @@ -6517,7 +4683,7 @@ lemma subst_subst3: assumes a: "a\(P,N,c)" "c\(M,N)" "x\(y,P,M)" "y\(P,x)" "M\Ax y a" shows "N{x:=.M}{y:=.P} = N{y:=.P}{x:=.(M{y:=.P})}" -using a + using a proof(nominal_induct N avoiding: x y a c M P rule: trm.strong_induct) case (Ax z c) then show ?case @@ -6527,37 +4693,17 @@ then show ?case apply(simp add: fresh_atm fresh_prod trm.inject abs_fresh) apply(auto) - apply(simp add: fresh_atm) - apply(simp add: trm.inject) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substn) - apply(simp add: abs_fresh subst_fresh) - apply(simp add: fresh_prod subst_fresh fresh_atm) - apply(subgoal_tac "P \ Ax x c") - apply(simp) - apply(simp add: forget) - apply(clarify) - apply(simp add: fresh_atm) - apply(case_tac "x\M'") - apply(simp add: forget) - apply(simp add: not_Ax2) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substn) - apply(simp add: abs_fresh subst_fresh) - apply(simp add: fresh_prod subst_fresh fresh_atm) - apply(simp) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substn) - apply(simp add: abs_fresh subst_fresh) - apply(simp add: fresh_prod subst_fresh fresh_atm) - apply(auto) - apply(case_tac "y\M'") - apply(simp add: forget) - apply(simp add: not_Ax2) - done + apply(auto simp add: fresh_atm trm.inject forget) + apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp) + apply (metis forget(1) not_Ax2) + apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp) + apply(rule sym) + apply(rule trans) + apply(rule better_Cut_substn) + apply(simp add: abs_fresh subst_fresh) + apply(simp add: fresh_prod subst_fresh fresh_atm) + apply(simp) + by (metis forget(1) not_Ax2) next case NotR then show ?case @@ -6566,39 +4712,17 @@ case (NotL d M' u) then show ?case 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)+ - apply(simp add: fresh_fun_simp_NotL) - apply(rule trans) - apply(rule better_Cut_substn) - apply(simp add: abs_fresh subst_fresh) - apply(simp add: fresh_prod subst_fresh fresh_atm) - apply(simp) - apply(simp add: trm.inject alpha) - apply(rule trans) - apply(rule substn.simps) - apply(simp add: fresh_prod fresh_atm) - apply(simp add: fresh_atm) - apply(rule exists_fresh'(1)[OF fs_name1]) + apply(subgoal_tac "\x'::name. x'\(y,P,M,M{y:=.P},M'{x:=.M},M'{y:=.P}{x:=.M{y:=.P}})") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_NotL) + apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp) + apply(rule exists_fresh'(1)[OF fs_name1]) apply(subgoal_tac "\x'::name. x'\(x,y,P,M,M'{y:=.P},M'{y:=.P}{x:=.M{y:=.P}})") - apply(erule exE, simp add: fresh_prod) - apply(erule conjE)+ - apply(simp add: fresh_fun_simp_NotL) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substn) - apply(simp add: abs_fresh subst_fresh) - apply(simp add: fresh_atm subst_fresh fresh_prod) - apply(subgoal_tac "P \ Ax x c") - apply(simp) - apply(simp add: forget trm.inject alpha) - apply(rule trans) - apply(rule substn.simps) - apply(simp add: fresh_atm subst_fresh) - apply(simp add: fresh_atm) - apply(clarify) - apply(simp add: fresh_atm) + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_NotL) + apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp) apply(rule exists_fresh'(1)[OF fs_name1]) done next @@ -6609,78 +4733,34 @@ case (AndL1 u M' v) then show ?case 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)+ - apply(simp add: fresh_fun_simp_AndL1) - apply(rule trans) - apply(rule better_Cut_substn) - apply(simp add: abs_fresh subst_fresh) - apply(simp add: fresh_prod subst_fresh fresh_atm) - apply(simp) - apply(simp add: trm.inject alpha) - apply(rule trans) - apply(rule substn.simps) - apply(simp add: fresh_prod fresh_atm) - apply(simp add: fresh_atm) - apply(rule exists_fresh'(1)[OF fs_name1]) + apply(subgoal_tac "\x'::name. x'\(u,y,v,P,M,M{y:=.P},M'{x:=.M},M'{y:=.P}{x:=.M{y:=.P}})") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_AndL1) + apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp) + apply(rule exists_fresh'(1)[OF fs_name1]) apply(subgoal_tac "\x'::name. x'\(x,y,u,v,P,M,M'{y:=.P},M'{y:=.P}{x:=.M{y:=.P}})") - apply(erule exE, simp add: fresh_prod) - apply(erule conjE)+ - apply(simp add: fresh_fun_simp_AndL1) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substn) - apply(simp add: abs_fresh subst_fresh) - apply(simp add: fresh_atm subst_fresh fresh_prod) - apply(subgoal_tac "P \ Ax x c") - apply(simp) - apply(simp add: forget trm.inject alpha) - apply(rule trans) - apply(rule substn.simps) - apply(simp add: fresh_atm subst_fresh) - apply(simp add: fresh_atm) - apply(clarify) - apply(simp add: fresh_atm) + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_AndL1) + apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp) apply(rule exists_fresh'(1)[OF fs_name1]) done next case (AndL2 u M' v) then show ?case 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)+ - apply(simp add: fresh_fun_simp_AndL2) - apply(rule trans) - apply(rule better_Cut_substn) - apply(simp add: abs_fresh subst_fresh) - apply(simp add: fresh_prod subst_fresh fresh_atm) - apply(simp) - apply(simp add: trm.inject alpha) - apply(rule trans) - apply(rule substn.simps) - apply(simp add: fresh_prod fresh_atm) - apply(simp add: fresh_atm) - apply(rule exists_fresh'(1)[OF fs_name1]) + apply(subgoal_tac "\x'::name. x'\(u,y,v,P,M,M{y:=.P},M'{x:=.M},M'{y:=.P}{x:=.M{y:=.P}})") + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_AndL2) + apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp) + apply(rule exists_fresh'(1)[OF fs_name1]) apply(subgoal_tac "\x'::name. x'\(x,y,u,v,P,M,M'{y:=.P},M'{y:=.P}{x:=.M{y:=.P}})") - apply(erule exE, simp add: fresh_prod) - apply(erule conjE)+ - apply(simp add: fresh_fun_simp_AndL2) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substn) - apply(simp add: abs_fresh subst_fresh) - apply(simp add: fresh_atm subst_fresh fresh_prod) - apply(subgoal_tac "P \ Ax x c") - apply(simp) - apply(simp add: forget trm.inject alpha) - apply(rule trans) - apply(rule substn.simps) - apply(simp add: fresh_atm subst_fresh) - apply(simp add: fresh_atm) - apply(clarify) - apply(simp add: fresh_atm) + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_AndL2) + apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp) apply(rule exists_fresh'(1)[OF fs_name1]) done next @@ -6695,44 +4775,19 @@ case (OrL x1 M' x2 M'' x3) then show ?case 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(subgoal_tac "\x'::name. x'\(y,P,M,M{y:=.P},M'{x:=.M},M'{y:=.P}{x:=.M{y:=.P}}, x1,x2,x3,M''{x:=.M},M''{y:=.P}{x:=.M{y:=.P}})") - apply(erule exE, simp add: fresh_prod) - apply(erule conjE)+ - apply(simp add: fresh_fun_simp_OrL) - apply(rule trans) - apply(rule better_Cut_substn) - apply(simp add: abs_fresh subst_fresh) - apply(simp add: fresh_prod subst_fresh fresh_atm) - apply(simp) - apply(simp add: trm.inject alpha) - apply(rule trans) - apply(rule substn.simps) - apply(simp add: fresh_prod fresh_atm subst_fresh) - apply(simp add: fresh_prod fresh_atm subst_fresh) - apply(simp add: fresh_atm) - apply(simp add: fresh_atm) - apply(rule exists_fresh'(1)[OF fs_name1]) + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_OrL) + apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp) + apply(rule exists_fresh'(1)[OF fs_name1]) apply(subgoal_tac "\x'::name. x'\(x,y,P,M,M'{y:=.P},M'{y:=.P}{x:=.M{y:=.P}}, x1,x2,x3,M''{y:=.P},M''{y:=.P}{x:=.M{y:=.P}})") - apply(erule exE, simp add: fresh_prod) - apply(erule conjE)+ - apply(simp add: fresh_fun_simp_OrL) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substn) - apply(simp add: abs_fresh subst_fresh) - apply(simp add: fresh_atm subst_fresh fresh_prod) - apply(simp add: fresh_prod fresh_atm) - apply(auto) - apply(simp add: fresh_atm) - apply(simp add: forget trm.inject alpha) - apply(rule trans) - apply(rule substn.simps) - apply(simp add: fresh_atm subst_fresh) - apply(simp add: fresh_atm subst_fresh) - apply(simp add: fresh_atm) - apply(simp add: fresh_atm) + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_OrL) + apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp) apply(rule exists_fresh'(1)[OF fs_name1]) done next @@ -6743,42 +4798,19 @@ case (ImpL d M' x1 M'' x2) then show ?case 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}}, + apply(subgoal_tac "\x'::name. x'\(y,P,M,M{y:=.P},M'{x2:=.M},M'{y:=.P}{x2:=.M{y:=.P}}, x1,x2,M''{x2:=.M},M''{y:=.P}{x2:=.M{y:=.P}})") - apply(erule exE, simp add: fresh_prod) - apply(erule conjE)+ - apply(simp add: fresh_fun_simp_ImpL) - apply(rule trans) - apply(rule better_Cut_substn) - apply(simp add: abs_fresh subst_fresh) - apply(simp add: fresh_prod subst_fresh fresh_atm) - apply(simp) - apply(simp add: trm.inject alpha) - apply(rule trans) - apply(rule substn.simps) - apply(simp add: fresh_prod fresh_atm subst_fresh) - apply(simp add: fresh_prod fresh_atm subst_fresh) - apply(simp add: fresh_atm) - apply(rule exists_fresh'(1)[OF fs_name1]) + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_ImpL) + apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp) + apply(rule exists_fresh'(1)[OF fs_name1]) apply(subgoal_tac "\x'::name. x'\(x,y,P,M,M'{x2:=.P},M'{x2:=.P}{x:=.M{x2:=.P}}, x1,x2,M''{x2:=.P},M''{x2:=.P}{x:=.M{x2:=.P}})") - apply(erule exE, simp add: fresh_prod) - apply(erule conjE)+ - apply(simp add: fresh_fun_simp_ImpL) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substn) - apply(simp add: abs_fresh subst_fresh) - apply(simp add: fresh_atm subst_fresh fresh_prod) - apply(simp add: fresh_prod fresh_atm) - apply(auto) - apply(simp add: fresh_atm) - apply(simp add: forget trm.inject alpha) - apply(rule trans) - apply(rule substn.simps) - apply(simp add: fresh_atm subst_fresh) - apply(simp add: fresh_atm subst_fresh) - apply(simp add: fresh_atm) + apply(erule exE, simp add: fresh_prod) + apply(erule conjE)+ + apply(simp add: fresh_fun_simp_ImpL) + apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp) apply(rule exists_fresh'(1)[OF fs_name1]) done qed @@ -6786,239 +4818,96 @@ lemma subst_subst4: assumes a: "x\(P,N,y)" "y\(M,N)" "a\(c,P,M)" "c\(P,a)" "M\Ax x c" shows "N{a:=(x).M}{c:=(y).P} = N{c:=(y).P}{a:=(x).(M{c:=(y).P})}" -using a + using a 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: subst_fresh abs_fresh fresh_atm fresh_prod forget) -next case (Cut d M' u M'') then show ?case apply(simp add: fresh_atm fresh_prod trm.inject abs_fresh) apply(auto) - apply(simp add: fresh_atm) - apply(simp add: trm.inject) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substc) - apply(simp add: abs_fresh subst_fresh fresh_atm) - apply(simp add: fresh_prod subst_fresh abs_fresh fresh_atm) - apply(subgoal_tac "P \ Ax y a") - apply(simp) - apply(simp add: forget) - apply(clarify) - apply(simp add: fresh_atm) - apply(case_tac "a\M''") - apply(simp add: forget) - apply(simp add: not_Ax1) + apply(simp add: fresh_atm) + apply(simp add: trm.inject) + apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp) + apply (metis forget(2) not_Ax1) + apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp) apply(rule sym) apply(rule trans) - apply(rule better_Cut_substc) - apply(simp add: fresh_prod subst_fresh fresh_atm) - apply(simp add: abs_fresh subst_fresh) - apply(simp) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substc) - apply(simp add: fresh_prod subst_fresh fresh_atm) - apply(simp add: abs_fresh subst_fresh) + apply(rule better_Cut_substc) + apply(simp add: fresh_prod subst_fresh fresh_atm) + apply(simp add: abs_fresh subst_fresh) apply(auto) - apply(case_tac "c\M''") - apply(simp add: forget) - apply(simp add: not_Ax1) - done -next - case NotL - then show ?case - by(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget) + by (metis forget(2) not_Ax1) next case (NotR u M' d) then show ?case apply(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget) - apply(generate_fresh "coname") - apply(fresh_fun_simp) - apply(fresh_fun_simp) - apply(simp add: abs_fresh subst_fresh) - apply(rule trans) - apply(rule better_Cut_substc) - apply(simp) - apply(simp add: abs_fresh) - apply(simp) - apply(simp add: trm.inject alpha) - apply(rule trans) - apply(rule substc.simps) - apply(simp add: fresh_prod fresh_atm) - apply(auto simp: fresh_atm fresh_prod)[1] + apply(generate_fresh "coname") + apply(fresh_fun_simp) + apply(fresh_fun_simp) + apply(simp add: abs_fresh subst_fresh) + apply(simp add: abs_fresh subst_fresh) + apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp) apply(generate_fresh "coname") apply(fresh_fun_simp) apply(fresh_fun_simp) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substc) - apply(simp add: fresh_prod fresh_atm subst_fresh) - apply(simp add: abs_fresh subst_fresh) - apply(auto simp: 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: fresh_prod fresh_atm) + apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp) done next - case AndL1 - then show ?case - by(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget) -next - case AndL2 - then show ?case - 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: subst_fresh abs_fresh fresh_atm fresh_prod forget) - apply(generate_fresh "coname") - apply(fresh_fun_simp) - apply(fresh_fun_simp) - apply(simp add: abs_fresh subst_fresh) - apply(rule trans) - apply(rule better_Cut_substc) - apply(simp) - apply(simp add: abs_fresh) - apply(simp) - apply(simp add: trm.inject alpha) - apply(rule trans) - apply(rule substc.simps) - apply(auto simp: fresh_prod fresh_atm subst_fresh)[1] - apply(auto simp: fresh_prod fresh_atm subst_fresh)[1] - apply(simp) - apply(auto simp: fresh_atm fresh_prod)[1] + apply(generate_fresh "coname") + apply(fresh_fun_simp) + apply(fresh_fun_simp) + apply(simp add: abs_fresh subst_fresh) + apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp) apply(generate_fresh "coname") apply(fresh_fun_simp) apply(fresh_fun_simp) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substc) - apply(simp add: subst_fresh fresh_atm fresh_prod) - apply(simp add: abs_fresh subst_fresh) - apply(auto simp: fresh_atm)[1] - apply(simp add: trm.inject alpha forget) - apply(rule trans) - apply(rule substc.simps) - 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: fresh_atm fresh_prod)[1] + apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp) done next - case OrL - then show ?case - by(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget) -next case (OrR1 d M' e) then show ?case apply(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget) + apply(generate_fresh "coname") + apply(fresh_fun_simp) + apply(fresh_fun_simp) + apply(simp add: abs_fresh subst_fresh) + apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp) apply(generate_fresh "coname") apply(fresh_fun_simp) apply(fresh_fun_simp) - apply(simp add: abs_fresh subst_fresh) - apply(rule trans) - apply(rule better_Cut_substc) - apply(simp) - apply(simp add: abs_fresh) - apply(simp) - apply(simp add: trm.inject alpha) - apply(rule trans) - apply(rule substc.simps) - apply(auto simp: 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) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substc) - apply(simp add: subst_fresh fresh_atm fresh_prod) - apply(simp add: abs_fresh subst_fresh) - apply(auto simp: fresh_atm)[1] - apply(simp add: trm.inject alpha forget) - apply(rule trans) - apply(rule substc.simps) - apply(auto simp: fresh_prod fresh_atm subst_fresh)[1] - apply(auto simp: fresh_prod fresh_atm subst_fresh)[1] + apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp) done next case (OrR2 d M' e) then show ?case apply(auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget) - apply(generate_fresh "coname") - apply(fresh_fun_simp) - apply(fresh_fun_simp) - apply(simp add: abs_fresh subst_fresh) - apply(rule trans) - apply(rule better_Cut_substc) - apply(simp) - apply(simp add: abs_fresh) - apply(simp) - apply(simp add: trm.inject alpha) - apply(rule trans) - apply(rule substc.simps) - apply(auto simp: 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) + apply(simp add: abs_fresh subst_fresh) + apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp) apply(generate_fresh "coname") apply(fresh_fun_simp) apply(fresh_fun_simp) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substc) - apply(simp add: subst_fresh fresh_atm fresh_prod) - apply(simp add: abs_fresh subst_fresh) - apply(auto simp: fresh_atm)[1] - apply(simp add: trm.inject alpha forget) - apply(rule trans) - apply(rule substc.simps) - apply(auto simp: fresh_prod fresh_atm subst_fresh)[1] - apply(auto simp: fresh_prod fresh_atm subst_fresh)[1] + apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp) done next - case ImpL - then show ?case - 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: subst_fresh abs_fresh fresh_atm fresh_prod forget) - apply(generate_fresh "coname") - apply(fresh_fun_simp) - apply(fresh_fun_simp) - apply(simp add: abs_fresh subst_fresh) - apply(rule trans) - apply(rule better_Cut_substc) - apply(simp) - apply(simp add: abs_fresh) - apply(simp) - apply(simp add: trm.inject alpha) - apply(rule trans) - apply(rule substc.simps) - apply(auto simp: 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) + apply(simp add: abs_fresh subst_fresh) + apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp) apply(generate_fresh "coname") apply(fresh_fun_simp) apply(fresh_fun_simp) - apply(rule sym) - apply(rule trans) - apply(rule better_Cut_substc) - apply(simp add: subst_fresh fresh_atm fresh_prod) - apply(simp add: abs_fresh subst_fresh) - apply(auto simp: fresh_atm)[1] - apply(simp add: trm.inject alpha forget) - apply(rule trans) - apply(rule substc.simps) - 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] + apply (force simp add: trm.inject alpha forget fresh_prod fresh_atm subst_fresh abs_fresh abs_supp fin_supp)+ done -qed +qed (auto simp: subst_fresh abs_fresh fresh_atm fresh_prod forget) + end diff -r 9f8034d29365 -r fbb38db0435d src/HOL/Nominal/Examples/Class2.thy --- a/src/HOL/Nominal/Examples/Class2.thy Mon Jul 29 15:26:03 2024 +0200 +++ b/src/HOL/Nominal/Examples/Class2.thy Mon Jul 29 15:26:56 2024 +0200 @@ -257,7 +257,7 @@ apply(auto) apply(rule aux4) apply(simp add: trm.inject alpha calc_atm fresh_atm) - apply(rule a_star_trans) + apply(rule rtranclp_trans) apply(rule a_starI) apply(rule al_redu) apply(rule l_redu.intros) @@ -339,7 +339,7 @@ then show ?thesis apply - apply(rule a_star_CutL) - apply(rule a_star_trans) + apply(rule rtranclp_trans) apply(rule a_starI) apply(rule ac_redu) apply(rule better_left) @@ -414,7 +414,7 @@ then show ?thesis apply - apply(rule a_star_CutL) - apply(rule a_star_trans) + apply(rule rtranclp_trans) apply(rule a_starI) apply(rule ac_redu) apply(rule better_left) @@ -490,7 +490,7 @@ then show ?thesis apply - apply(rule a_star_CutL) - apply(rule a_star_trans) + apply(rule rtranclp_trans) apply(rule a_starI) apply(rule ac_redu) apply(rule better_left) @@ -566,7 +566,7 @@ then show ?thesis apply - apply(rule a_star_CutL) - apply(rule a_star_trans) + apply(rule rtranclp_trans) apply(rule a_starI) apply(rule ac_redu) apply(rule better_left) @@ -642,7 +642,7 @@ then show ?thesis apply - apply(rule a_star_CutL) - apply(rule a_star_trans) + apply(rule rtranclp_trans) apply(rule a_starI) apply(rule ac_redu) apply(rule better_left) @@ -723,7 +723,7 @@ apply - apply(rule a_star_CutL) apply(rule a_star_CutL) - apply(rule a_star_trans) + apply(rule rtranclp_trans) apply(rule a_starI) apply(rule ac_redu) apply(rule better_left) @@ -792,7 +792,7 @@ apply(auto) apply(rule aux4) apply(simp add: trm.inject alpha calc_atm fresh_atm) - apply(rule a_star_trans) + apply(rule rtranclp_trans) apply(rule a_starI) apply(rule al_redu) apply(rule l_redu.intros) @@ -849,7 +849,7 @@ then show ?thesis apply - apply(rule a_star_CutR) - apply(rule a_star_trans) + apply(rule rtranclp_trans) apply(rule a_starI) apply(rule ac_redu) apply(rule better_right) @@ -924,7 +924,7 @@ then show ?thesis apply - apply(rule a_star_CutR) - apply(rule a_star_trans) + apply(rule rtranclp_trans) apply(rule a_starI) apply(rule ac_redu) apply(rule better_right) @@ -1000,7 +1000,7 @@ then show ?thesis apply - apply(rule a_star_CutR) - apply(rule a_star_trans) + apply(rule rtranclp_trans) apply(rule a_starI) apply(rule ac_redu) apply(rule better_right) @@ -1076,7 +1076,7 @@ then show ?thesis apply - apply(rule a_star_CutR) - apply(rule a_star_trans) + apply(rule rtranclp_trans) apply(rule a_starI) apply(rule ac_redu) apply(rule better_right) @@ -1152,7 +1152,7 @@ then show ?thesis apply - apply(rule a_star_CutR) - apply(rule a_star_trans) + apply(rule rtranclp_trans) apply(rule a_starI) apply(rule ac_redu) apply(rule better_right) @@ -1219,7 +1219,7 @@ apply - apply(rule a_star_CutL) apply(rule a_star_CutR) - apply(rule a_star_trans) + apply(rule rtranclp_trans) apply(rule a_starI) apply(rule al_redu) apply(rule better_LAxL_intro) @@ -1233,7 +1233,7 @@ apply - apply(rule a_star_CutL) apply(rule a_star_CutR) - apply(rule a_star_trans) + apply(rule rtranclp_trans) apply(rule a_starI) apply(rule ac_redu) apply(rule better_right) @@ -1287,7 +1287,7 @@ then show ?thesis using LImp apply - apply(rule a_star_CutR) - apply(rule a_star_trans) + apply(rule rtranclp_trans) apply(rule a_starI) apply(rule ac_redu) apply(rule better_right) @@ -1339,7 +1339,7 @@ then show ?thesis using LImp apply - apply(rule a_star_CutR) - apply(rule a_star_trans) + apply(rule rtranclp_trans) apply(rule a_starI) apply(rule ac_redu) apply(rule better_right) @@ -1367,7 +1367,7 @@ apply - apply(rule a_star_CutL) apply(rule a_star_CutR) - apply(rule a_star_trans) + apply(rule rtranclp_trans) apply(rule a_starI) apply(rule ac_redu) apply(rule better_right) @@ -1426,17 +1426,17 @@ apply(drule_tac x="c" in meta_spec) apply(drule_tac x="P" in meta_spec) apply(simp) - apply(rule a_star_trans) + apply(rule rtranclp_trans) apply(rule a_star_CutL) apply(assumption) - apply(rule a_star_trans) + apply(rule rtranclp_trans) apply(rule_tac M'="P[c\c>a]" in a_star_CutL) apply(case_tac "fic P c") apply(rule a_starI) apply(rule al_redu) apply(rule better_LAxR_intro) apply(simp) - apply(rule a_star_trans) + apply(rule rtranclp_trans) apply(rule a_starI) apply(rule ac_redu) apply(rule better_left) @@ -1633,17 +1633,17 @@ apply(drule_tac x="y" in meta_spec) apply(drule_tac x="P" in meta_spec) apply(simp) - apply(rule a_star_trans) + apply(rule rtranclp_trans) apply(rule a_star_CutR) apply(assumption) - apply(rule a_star_trans) + apply(rule rtranclp_trans) apply(rule_tac N'="P[y\n>x]" in a_star_CutR) apply(case_tac "fin P y") apply(rule a_starI) apply(rule al_redu) apply(rule better_LAxL_intro) apply(simp) - apply(rule a_star_trans) + apply(rule rtranclp_trans) apply(rule a_starI) apply(rule ac_redu) apply(rule better_right) diff -r 9f8034d29365 -r fbb38db0435d src/HOL/Nominal/Examples/Class3.thy --- a/src/HOL/Nominal/Examples/Class3.thy Mon Jul 29 15:26:03 2024 +0200 +++ b/src/HOL/Nominal/Examples/Class3.thy Mon Jul 29 15:26:56 2024 +0200 @@ -6115,7 +6115,7 @@ apply(case_tac "coname\(idc \ a)") apply(simp add: lookup3) apply(simp add: lookup4) - apply(rule a_star_trans) + apply(rule rtranclp_trans) apply(rule a_starI) apply(rule al_redu) apply(rule better_LAxR_intro) @@ -6124,14 +6124,14 @@ apply(simp add: lookup2) apply(case_tac "coname\(idc \ a)") apply(simp add: lookup5) - apply(rule a_star_trans) + apply(rule rtranclp_trans) apply(rule a_starI) apply(rule al_redu) apply(rule better_LAxR_intro) apply(rule fic.intros) apply(simp) apply(simp add: lookup6) - apply(rule a_star_trans) + apply(rule rtranclp_trans) apply(rule a_starI) apply(rule al_redu) apply(rule better_LAxR_intro) @@ -6156,7 +6156,7 @@ apply(drule idc_cmaps) apply(simp) apply(subgoal_tac "c\idn \ x,idc \ a") - apply(rule a_star_trans) + apply(rule rtranclp_trans) apply(rule a_starI) apply(rule al_redu) apply(rule better_LAxR_intro) @@ -6181,7 +6181,7 @@ apply(drule idn_nmaps) apply(simp) apply(subgoal_tac "c\idn \ x,idc \ a") - apply(rule a_star_trans) + apply(rule rtranclp_trans) apply(rule a_starI) apply(rule al_redu) apply(rule better_LAxL_intro) @@ -6207,7 +6207,7 @@ apply(simp) apply(subgoal_tac "c\idn \ x,idc \ a") apply(subgoal_tac "c\idn \ x,idc \ a") - apply(rule a_star_trans) + apply(rule rtranclp_trans) apply(rule a_starI) apply(rule al_redu) apply(rule better_LAxR_intro) @@ -6258,7 +6258,7 @@ apply(drule idn_nmaps) apply(simp) apply(subgoal_tac "c\idn \ x,idc \ a") - apply(rule a_star_trans) + apply(rule rtranclp_trans) apply(rule a_starI) apply(rule al_redu) apply(rule better_LAxL_intro) @@ -6287,7 +6287,7 @@ apply(drule idn_nmaps) apply(simp) apply(subgoal_tac "c\idn \ x,idc \ a") - apply(rule a_star_trans) + apply(rule rtranclp_trans) apply(rule a_starI) apply(rule al_redu) apply(rule better_LAxL_intro) @@ -6316,7 +6316,7 @@ apply(drule idc_cmaps) apply(simp) apply(subgoal_tac "c\idn \ x,idc \ a") - apply(rule a_star_trans) + apply(rule rtranclp_trans) apply(rule a_starI) apply(rule al_redu) apply(rule better_LAxR_intro) @@ -6345,7 +6345,7 @@ apply(drule idc_cmaps) apply(simp) apply(subgoal_tac "c\idn \ x,idc \ a") - apply(rule a_star_trans) + apply(rule rtranclp_trans) apply(rule a_starI) apply(rule al_redu) apply(rule better_LAxR_intro) @@ -6375,7 +6375,7 @@ apply(simp) apply(subgoal_tac "c\idn \ x,idc \ a") apply(subgoal_tac "c\idn \ x,idc \ a") - apply(rule a_star_trans) + apply(rule rtranclp_trans) apply(rule a_starI) apply(rule al_redu) apply(rule better_LAxL_intro) @@ -6425,7 +6425,7 @@ apply(drule idc_cmaps) apply(simp) apply(subgoal_tac "c\idn \ x,idc \ a") - apply(rule a_star_trans) + apply(rule rtranclp_trans) apply(rule a_starI) apply(rule al_redu) apply(rule better_LAxR_intro) @@ -6455,7 +6455,7 @@ apply(simp) apply(subgoal_tac "c\idn \ x,idc \ a") apply(subgoal_tac "c\idn \ x,idc \ a") - apply(rule a_star_trans) + apply(rule rtranclp_trans) apply(rule a_starI) apply(rule al_redu) apply(rule better_LAxL_intro) diff -r 9f8034d29365 -r fbb38db0435d src/HOL/Num.thy --- a/src/HOL/Num.thy Mon Jul 29 15:26:03 2024 +0200 +++ b/src/HOL/Num.thy Mon Jul 29 15:26:56 2024 +0200 @@ -677,6 +677,33 @@ lemma not_numeral_less_zero: "\ numeral n < 0" by (simp add: not_less zero_le_numeral) +lemma one_of_nat_le_iff [simp]: "1 \ of_nat k \ 1 \ k" + using of_nat_le_iff [of 1] by simp + +lemma numeral_nat_le_iff [simp]: "numeral n \ of_nat k \ numeral n \ k" + using of_nat_le_iff [of "numeral n"] by simp + +lemma of_nat_le_1_iff [simp]: "of_nat k \ 1 \ k \ 1" + using of_nat_le_iff [of _ 1] by simp + +lemma of_nat_le_numeral_iff [simp]: "of_nat k \ numeral n \ k \ numeral n" + using of_nat_le_iff [of _ "numeral n"] by simp + +lemma one_of_nat_less_iff [simp]: "1 < of_nat k \ 1 < k" + using of_nat_less_iff [of 1] by simp + +lemma numeral_nat_less_iff [simp]: "numeral n < of_nat k \ numeral n < k" + using of_nat_less_iff [of "numeral n"] by simp + +lemma of_nat_less_1_iff [simp]: "of_nat k < 1 \ k < 1" + using of_nat_less_iff [of _ 1] by simp + +lemma of_nat_less_numeral_iff [simp]: "of_nat k < numeral n \ k < numeral n" + using of_nat_less_iff [of _ "numeral n"] by simp + +lemma of_nat_eq_numeral_iff [simp]: "of_nat k = numeral n \ k = numeral n" + using of_nat_eq_iff [of _ "numeral n"] by simp + lemmas le_numeral_extra = zero_le_one not_one_le_zero order_refl [of 0] order_refl [of 1] diff -r 9f8034d29365 -r fbb38db0435d src/HOL/Number_Theory/Residues.thy --- a/src/HOL/Number_Theory/Residues.thy Mon Jul 29 15:26:03 2024 +0200 +++ b/src/HOL/Number_Theory/Residues.thy Mon Jul 29 15:26:56 2024 +0200 @@ -283,8 +283,10 @@ defines "R \ residue_ring (int p)" sublocale residues_prime < residues p - unfolding R_def residues_def - by (auto simp: p_prime prime_gt_1_int) +proof + show "1 < int p" + using prime_gt_1_nat by auto +qed context residues_prime begin diff -r 9f8034d29365 -r fbb38db0435d src/HOL/Proofs/ex/XML_Data.thy --- a/src/HOL/Proofs/ex/XML_Data.thy Mon Jul 29 15:26:03 2024 +0200 +++ b/src/HOL/Proofs/ex/XML_Data.thy Mon Jul 29 15:26:56 2024 +0200 @@ -14,11 +14,11 @@ ML \ fun export_proof thy thm = thm |> Proof_Syntax.standard_proof_of {full = true, expand_name = Thm.expand_name thm} - |> Proofterm.encode (Sign.consts_of thy); + |> Proofterm.encode thy; fun import_proof thy xml = let - val prf = Proofterm.decode (Sign.consts_of thy) xml; + val prf = Proofterm.decode thy xml; val (prf', _) = Proofterm.freeze_thaw_prf prf; in Drule.export_without_context (Proof_Checker.thm_of_proof thy prf') end; \ diff -r 9f8034d29365 -r fbb38db0435d src/HOL/Real.thy --- a/src/HOL/Real.thy Mon Jul 29 15:26:03 2024 +0200 +++ b/src/HOL/Real.thy Mon Jul 29 15:26:56 2024 +0200 @@ -750,16 +750,10 @@ lemma less_RealD: assumes "cauchy Y" shows "x < Real Y \ \n. x < of_rat (Y n)" - apply (erule contrapos_pp) - apply (simp add: not_less) - apply (erule Real_leI [OF assms]) - done + by (meson Real_leI assms leD leI) lemma of_nat_less_two_power [simp]: "of_nat n < (2::'a::linordered_idom) ^ n" - apply (induct n) - apply simp - apply (metis add_le_less_mono mult_2 of_nat_Suc one_le_numeral one_le_power power_Suc) - done + by auto lemma complete_real: fixes S :: "real set" @@ -1309,6 +1303,8 @@ for x y :: real by auto +lemma mult_ge1_I: "\x\1; y\1\ \ x*y \ (1::real)" + using mult_mono by fastforce subsection \Lemmas about powers\ @@ -1620,6 +1616,18 @@ lemma Rats_no_bot_less: "\q \ \. q < x" for x :: real by (auto intro!: bexI[of _ "of_int (\x\ - 1)"]) linarith +lemma floor_ceiling_diff_le: "0 \ r \ nat\real k - r\ \ k - nat\r\" + by linarith + +lemma floor_ceiling_diff_le': "nat\r - real k\ \ nat\r\ - k" + by linarith + +lemma ceiling_floor_diff_ge: "nat\r - real k\ \ nat\r\ - k" + by linarith + +lemma ceiling_floor_diff_ge': "r \ k \ nat\r - real k\ \ k - nat\r\" + by linarith + subsection \Exponentiation with floor\ diff -r 9f8034d29365 -r fbb38db0435d src/HOL/Set_Interval.thy --- a/src/HOL/Set_Interval.thy Mon Jul 29 15:26:03 2024 +0200 +++ b/src/HOL/Set_Interval.thy Mon Jul 29 15:26:56 2024 +0200 @@ -316,6 +316,71 @@ with * show "a = b \ b = c" by auto qed simp +text \The following results generalise their namesakes in @{theory HOL.Nat} to intervals\ + +lemma lift_Suc_mono_le_ivl: + assumes mono: "\n. n\N \ f n \ f (Suc n)" + and "n \ n'" and subN: "{n.. N" + shows "f n \ f n'" +proof (cases "n < n'") + case True + then show ?thesis + using subN + proof (induction n n' rule: less_Suc_induct) + case (1 i) + then show ?case + by (simp add: mono subsetD) + next + case (2 i j k) + have "f i \ f j" "f j \ f k" + using 2 by force+ + then show ?case by auto + qed +next + case False + with \n \ n'\ show ?thesis by auto +qed + +lemma lift_Suc_antimono_le_ivl: + assumes mono: "\n. n\N \ f n \ f (Suc n)" + and "n \ n'" and subN: "{n.. N" + shows "f n \ f n'" +proof (cases "n < n'") + case True + then show ?thesis + using subN + proof (induction n n' rule: less_Suc_induct) + case (1 i) + then show ?case + by (simp add: mono subsetD) + next + case (2 i j k) + have "f i \ f j" "f j \ f k" + using 2 by force+ + then show ?case by auto + qed +next + case False + with \n \ n'\ show ?thesis by auto +qed + +lemma lift_Suc_mono_less_ivl: + assumes mono: "\n. n\N \ f n < f (Suc n)" + and "n < n'" and subN: "{n.. N" + shows "f n < f n'" + using \n < n'\ + using subN +proof (induction n n' rule: less_Suc_induct) + case (1 i) + then show ?case + by (simp add: mono subsetD) +next + case (2 i j k) + have "f i < f j" "f j < f k" + using 2 by force+ + then show ?case by auto +qed + end context no_top @@ -2247,6 +2312,30 @@ finally show ?thesis by metis qed +lemma prod_divide_nat_ivl: + fixes f :: "nat \ 'a::idom_divide" + shows "\ m \ n; n \ p; prod f {m.. 0\ \ prod f {m.. syntax not available?*) + fixes f:: "nat \ 'a::idom_divide" + assumes "m \ n" "prod f {.. 0" + shows "(prod f {..n}) div (prod f {..i. f(n - i)) {..n - m}" +proof - + have "\i. i \ n-m \ \k\m. k \ n \ i = n-k" + by (metis Nat.le_diff_conv2 add.commute \m\n\ diff_diff_cancel diff_le_self order.trans) + then have eq: "{..n-m} = (-)n ` {m..n}" + by force + have inj: "inj_on ((-)n) {m..n}" + by (auto simp: inj_on_def) + have "prod (\i. f(n - i)) {..n - m} = prod f {m..n}" + by (simp add: eq prod.reindex_cong [OF inj]) + also have "\ = prod f {..n} div prod f {..Telescoping sums\ diff -r 9f8034d29365 -r fbb38db0435d src/HOL/Tools/inductive_realizer.ML --- a/src/HOL/Tools/inductive_realizer.ML Mon Jul 29 15:26:03 2024 +0200 +++ b/src/HOL/Tools/inductive_realizer.ML Mon Jul 29 15:26:56 2024 +0200 @@ -16,7 +16,7 @@ fun thm_name_of thm = (case Proofterm.fold_proof_atoms false (fn PThm ({thm_name, ...}, _) => cons thm_name | _ => I) [Thm.proof_of thm] [] of - [thm_name] => thm_name + [(thm_name, _)] => thm_name | _ => raise THM ("thm_name_of: bad proof of theorem", 0, [thm])); val short_name_of = Thm_Name.short o thm_name_of; diff -r 9f8034d29365 -r fbb38db0435d src/HOL/Tools/rewrite_hol_proof.ML --- a/src/HOL/Tools/rewrite_hol_proof.ML Mon Jul 29 15:26:03 2024 +0200 +++ b/src/HOL/Tools/rewrite_hol_proof.ML Mon Jul 29 15:26:56 2024 +0200 @@ -305,9 +305,9 @@ (** Replace congruence rules by substitution rules **) -fun strip_cong ps (PThm ({thm_name = ("HOL.cong", 0), ...}, _) % _ % _ % SOME x % SOME y %% +fun strip_cong ps (PThm ({thm_name = (("HOL.cong", 0), _), ...}, _) % _ % _ % SOME x % SOME y %% prfa %% prfT %% prf1 %% prf2) = strip_cong (((x, y), (prf2, prfa)) :: ps) prf1 - | strip_cong ps (PThm ({thm_name = ("HOL.refl", 0), ...}, _) % SOME f %% _) = SOME (f, ps) + | strip_cong ps (PThm ({thm_name = (("HOL.refl", 0), _), ...}, _) % SOME f %% _) = SOME (f, ps) | strip_cong _ _ = NONE; val subst_prf = Proofterm.any_head_of (Thm.proof_of @{thm subst}); @@ -330,15 +330,15 @@ fun mk_AbsP P t = AbsP ("H", Option.map HOLogic.mk_Trueprop P, t); -fun elim_cong_aux Ts (PThm ({thm_name = ("HOL.iffD1", 0), ...}, _) % _ % _ %% prf1 %% prf2) = +fun elim_cong_aux Ts (PThm ({thm_name = (("HOL.iffD1", 0), _), ...}, _) % _ % _ %% prf1 %% prf2) = Option.map (make_subst Ts prf2 []) (strip_cong [] prf1) - | elim_cong_aux Ts (PThm ({thm_name = ("HOL.iffD1", 0), ...}, _) % P % _ %% prf) = + | elim_cong_aux Ts (PThm ({thm_name = (("HOL.iffD1", 0), _), ...}, _) % P % _ %% prf) = Option.map (mk_AbsP P o make_subst Ts (PBound 0) []) (strip_cong [] (Proofterm.incr_boundvars 1 0 prf)) - | elim_cong_aux Ts (PThm ({thm_name = ("HOL.iffD2", 0), ...}, _) % _ % _ %% prf1 %% prf2) = + | elim_cong_aux Ts (PThm ({thm_name = (("HOL.iffD2", 0), _), ...}, _) % _ % _ %% prf1 %% prf2) = Option.map (make_subst Ts prf2 [] o apsnd (map (make_sym Ts))) (strip_cong [] prf1) - | elim_cong_aux Ts (PThm ({thm_name = ("HOL.iffD2", 0), ...}, _) % _ % P %% prf) = + | elim_cong_aux Ts (PThm ({thm_name = (("HOL.iffD2", 0), _), ...}, _) % _ % P %% prf) = Option.map (mk_AbsP P o make_subst Ts (PBound 0) [] o apsnd (map (make_sym Ts))) (strip_cong [] (Proofterm.incr_boundvars 1 0 prf)) | elim_cong_aux _ _ = NONE; diff -r 9f8034d29365 -r fbb38db0435d src/HOL/Topological_Spaces.thy --- a/src/HOL/Topological_Spaces.thy Mon Jul 29 15:26:03 2024 +0200 +++ b/src/HOL/Topological_Spaces.thy Mon Jul 29 15:26:56 2024 +0200 @@ -893,6 +893,15 @@ simp: eventually_at_filter less_le elim: eventually_elim2) +lemma (in order_topology) + shows at_within_Ici_at_right: "at a within {a..} = at_right a" + and at_within_Iic_at_left: "at a within {..a} = at_left a" + using order_tendstoD(2)[OF tendsto_ident_at [where s = "{a<..}"]] + using order_tendstoD(1)[OF tendsto_ident_at[where s = "{.. x < b \ at x within {a..b} = at x" by (rule at_within_open_subset[where S="{a<..n. X n \ X (Suc n)) \ incseq X" - using lift_Suc_mono_le[of X] by (auto simp: incseq_def) + by (simp add: mono_iff_le_Suc) lemma incseqD: "incseq f \ i \ j \ f i \ f j" by (auto simp: incseq_def) @@ -1252,7 +1261,7 @@ unfolding incseq_def by auto lemma decseq_SucI: "(\n. X (Suc n) \ X n) \ decseq X" - using order.lift_Suc_mono_le[OF dual_order, of X] by (auto simp: decseq_def) + by (simp add: antimono_iff_le_Suc) lemma decseqD: "decseq f \ i \ j \ f j \ f i" by (auto simp: decseq_def) diff -r 9f8034d29365 -r fbb38db0435d src/HOL/Transcendental.thy --- a/src/HOL/Transcendental.thy Mon Jul 29 15:26:03 2024 +0200 +++ b/src/HOL/Transcendental.thy Mon Jul 29 15:26:56 2024 +0200 @@ -1494,6 +1494,16 @@ for x y :: real by (auto simp: linorder_not_less [symmetric]) +lemma exp_mono: + fixes x y :: real + assumes "x \ y" + shows "exp x \ exp y" + using assms exp_le_cancel_iff by fastforce + +lemma exp_minus': "exp (-x) = 1 / (exp x)" + for x :: "'a::{real_normed_field,banach}" + by (simp add: exp_minus inverse_eq_divide) + lemma exp_inj_iff [iff]: "exp x = exp y \ x = y" for x y :: real by (simp add: order_eq_iff) @@ -1668,8 +1678,11 @@ for x :: real by (simp add: linorder_not_less [symmetric]) -lemma ln_mono: "\x::real. \x \ y; 0 < x; 0 < y\ \ ln x \ ln y" - using ln_le_cancel_iff by presburger +lemma ln_mono: "\x::real. \x \ y; 0 < x\ \ ln x \ ln y" + by simp + +lemma ln_strict_mono: "\x::real. \x < y; 0 < x\ \ ln x < ln y" + by simp lemma ln_inj_iff [simp]: "0 < x \ 0 < y \ ln x = ln y \ x = y" for x :: real @@ -2327,6 +2340,9 @@ \ \logarithm of \<^term>\x\ to base \<^term>\a\\ where "log a x = ln x / ln a" +lemma log_exp [simp]: "log b (exp x) = x / ln b" + by (simp add: log_def) + lemma tendsto_log [tendsto_intros]: "(f \ a) F \ (g \ b) F \ 0 < a \ a \ 1 \ b\0 \ ((\x. log (f x) (g x)) \ log a b) F" @@ -6988,7 +7004,7 @@ end lemma sinh_real_strict_mono: "strict_mono (sinh :: real \ real)" - by (rule pos_deriv_imp_strict_mono derivative_intros)+ auto + by (force intro: strict_monoI DERIV_pos_imp_increasing [where f=sinh] derivative_intros) lemma cosh_real_strict_mono: assumes "0 \ x" and "x < (y::real)" @@ -7004,10 +7020,10 @@ lemma tanh_real_strict_mono: "strict_mono (tanh :: real \ real)" proof - - have *: "tanh x ^ 2 < 1" for x :: real + have "tanh x ^ 2 < 1" for x :: real using tanh_real_bounds[of x] by (simp add: abs_square_less_1 abs_if) - show ?thesis - by (rule pos_deriv_imp_strict_mono) (insert *, auto intro!: derivative_intros) + then show ?thesis + by (force intro!: strict_monoI DERIV_pos_imp_increasing [where f=tanh] derivative_intros) qed lemma sinh_real_abs [simp]: "sinh (abs x :: real) = abs (sinh x)" diff -r 9f8034d29365 -r fbb38db0435d src/Pure/Admin/isabelle_cronjob.scala --- a/src/Pure/Admin/isabelle_cronjob.scala Mon Jul 29 15:26:03 2024 +0200 +++ b/src/Pure/Admin/isabelle_cronjob.scala Mon Jul 29 15:26:56 2024 +0200 @@ -356,7 +356,7 @@ " -e ISABELLE_SWIPL=/usr/local/bin/swipl", args = "-a -d '~~/src/Benchmarks'")), List( - Remote_Build("AFP (macOS 14 Sonoma, Apple Silicon)", "studio1-sonoma", history = 120, + Remote_Build("AFP macOS (macOS 14 Sonoma, Apple Silicon)", "studio1-sonoma", history = 120, history_base = "build_history_base_arm", java_heap = "8g", options = "-m32 -M1x5 -p pide_session=false -t AFP" + @@ -387,9 +387,9 @@ args = "-a -d '~~/src/Benchmarks'", count = () => 3)), List( - Remote_Build("Windows/AFP", "windows2", + Remote_Build("AFP Windows", "windows2", java_heap = "8g", - options = "-m32 -M1x6 -t AFP", + options = "-m32 -M1x5 -t AFP", args = "-a -X large -X slow", afp = true, detect = Build_Log.Prop.build_tags.toString + " = " + SQL.string("AFP"))) diff -r 9f8034d29365 -r fbb38db0435d src/Pure/Build/export_theory.ML --- a/src/Pure/Build/export_theory.ML Mon Jul 29 15:26:03 2024 +0200 +++ b/src/Pure/Build/export_theory.ML Mon Jul 29 15:26:56 2024 +0200 @@ -65,17 +65,6 @@ in ([], triple int string int (ass, delim, pri)) end]; -(* free variables: not declared in the context *) - -val is_free = not oo Name.is_declared; - -fun add_frees used = - fold_aterms (fn Free (x, T) => is_free used x ? insert (op =) (x, T) | _ => I); - -fun add_tfrees used = - (fold_types o fold_atyps) (fn TFree (a, S) => is_free used a ? insert (op =) (a, S) | _ => I); - - (* locales *) fun locale_content thy loc = @@ -149,6 +138,22 @@ val enabled = export_enabled context; + (* recode *) + + val thy_cache = thy; (* FIXME tmp *) + + val ztyp_of = ZTerm.ztyp_cache thy_cache; + val zterm_of = ZTerm.zterm_of thy_cache; + val zproof_of = Proofterm.proof_to_zproof thy_cache; + + val encode_ztyp = ZTerm.encode_ztyp; + val encode_zterm = ZTerm.encode_zterm {typed_vars = true}; + val encode_term = encode_zterm o zterm_of; + + val encode_standard_zterm = ZTerm.encode_zterm {typed_vars = false}; + val encode_standard_zproof = ZTerm.encode_zproof {typed_vars = false}; + + (* strict parents *) val parents = Theory.parents_of thy; @@ -228,11 +233,9 @@ (* consts *) - val encode_term = Term_XML.Encode.term consts; - val encode_const = let open XML.Encode Term_XML.Encode - in pair encode_syntax (pair (list string) (pair typ (pair (option encode_term) bool))) end; + in pair encode_syntax (pair (list string) (pair typ (pair (option encode_zterm) bool))) end; val _ = export_entities "consts" Sign.const_space (#constants (Consts.dest consts)) @@ -242,7 +245,8 @@ val syntax = get_syntax_const thy_ctxt c; val U = Logic.unvarifyT_global T; val U0 = Term.strip_sortsT U; - val trim_abbrev = Proofterm.standard_vars_term Name.context #> Term.strip_sorts; + fun trim_abbrev t = + ZTerm.standard_vars Name.context (zterm_of t, NONE) |> #prop |> ZTerm.strip_sorts; val abbrev' = Option.map trim_abbrev abbrev; val args = map (#1 o dest_TFree) (Consts.typargs consts (c, U0)); val propositional = Object_Logic.is_propositional thy_ctxt (Term.body_type U0); @@ -253,19 +257,21 @@ fun standard_prop used extra_shyps raw_prop raw_proof = let - val (prop, proof) = Proofterm.standard_vars used (raw_prop, raw_proof); - val args = rev (add_frees used prop []); - val typargs = rev (add_tfrees used prop []); - val used_typargs = fold (Name.declare o #1) typargs used; + val {typargs, args, prop, proof} = + ZTerm.standard_vars used (zterm_of raw_prop, Option.map zproof_of raw_proof); + val is_free = not o Name.is_declared used; + val args' = args |> filter (is_free o #1); + val typargs' = typargs |> filter (is_free o #1); + val used_typargs = fold (Name.declare o #1) typargs' used; val sorts = Name.invent used_typargs Name.aT (length extra_shyps) ~~ extra_shyps; - in ((sorts @ typargs, args, prop), proof) end; + in ((sorts @ typargs', args', prop), proof) end; fun standard_prop_of thm = standard_prop Name.context (Thm.extra_shyps thm) (Thm.full_prop_of thm); val encode_prop = let open XML.Encode Term_XML.Encode - in triple (list (pair string sort)) (list (pair string typ)) encode_term end; + in triple (list (pair string sort)) (list (pair string encode_ztyp)) encode_zterm end; fun encode_axiom used prop = encode_prop (#1 (standard_prop used [] prop NONE)); @@ -283,8 +289,8 @@ val lookup_thm_id = Global_Theory.lookup_thm_id thy; fun expand_name thm_id (header: Proofterm.thm_header) = - if #serial header = #serial thm_id then Thm_Name.empty - else the_default Thm_Name.empty (lookup_thm_id (Proofterm.thm_header_id header)); + if #serial header = #serial thm_id then Thm_Name.none + else the_default Thm_Name.none (lookup_thm_id (Proofterm.thm_header_id header)); fun entity_markup_thm (serial, (name, i)) = let @@ -308,13 +314,11 @@ val _ = Thm.expose_proofs thy [thm]; in (prop, deps, proof) |> - let - open XML.Encode Term_XML.Encode; - val encode_proof = Proofterm.encode_standard_proof consts; - in triple encode_prop (list Thm_Name.encode) encode_proof end + let open XML.Encode Term_XML.Encode + in triple encode_prop (list Thm_Name.encode) encode_standard_zproof end end; - fun export_thm (thm_id, thm_name) = + fun export_thm (thm_id, (thm_name, _)) = let val markup = entity_markup_thm (#serial thm_id, thm_name); val body = diff -r 9f8034d29365 -r fbb38db0435d src/Pure/General/same.ML --- a/src/Pure/General/same.ML Mon Jul 29 15:26:03 2024 +0200 +++ b/src/Pure/General/same.ML Mon Jul 29 15:26:56 2024 +0200 @@ -12,6 +12,7 @@ type 'a operation = ('a, 'a) function (*exception SAME*) val same: ('a, 'b) function val commit: 'a operation -> 'a -> 'a + val commit_if: bool -> 'a operation -> 'a -> 'a val commit_id: 'a operation -> 'a -> 'a * bool val catch: ('a, 'b) function -> 'a -> 'b option val compose: 'a operation -> 'a operation -> 'a operation @@ -31,6 +32,7 @@ fun same _ = raise SAME; fun commit f x = f x handle SAME => x; +fun commit_if b f x = if b then commit f x else f x; fun commit_id f x = (f x, false) handle SAME => (x, true); fun catch f x = SOME (f x) handle SAME => NONE; diff -r 9f8034d29365 -r fbb38db0435d src/Pure/General/table.ML --- a/src/Pure/General/table.ML Mon Jul 29 15:26:03 2024 +0200 +++ b/src/Pure/General/table.ML Mon Jul 29 15:26:56 2024 +0200 @@ -66,7 +66,8 @@ val insert_set: key -> set -> set val remove_set: key -> set -> set val make_set: key list -> set - val unsynchronized_cache: (key -> 'a) -> key -> 'a + type 'a cache_ops = {apply: key -> 'a, reset: unit -> unit, size: unit -> int} + val unsynchronized_cache: (key -> 'a) -> 'a cache_ops end; functor Table(Key: KEY): TABLE = @@ -591,12 +592,13 @@ (* cache *) +type 'a cache_ops = {apply: key -> 'a, reset: unit -> unit, size: unit -> int}; + fun unsynchronized_cache f = let val cache1 = Unsynchronized.ref empty; val cache2 = Unsynchronized.ref empty; - in - fn x => + fun apply x = (case lookup (! cache1) x of SOME y => y | NONE => @@ -604,9 +606,11 @@ SOME exn => Exn.reraise exn | NONE => (case Exn.result f x of - Exn.Res y => (Unsynchronized.change cache1 (update (x, y)); y) - | Exn.Exn exn => (Unsynchronized.change cache2 (update (x, exn)); Exn.reraise exn)))) - end; + Exn.Res y => (Unsynchronized.change cache1 (default (x, y)); y) + | Exn.Exn exn => (Unsynchronized.change cache2 (default (x, exn)); Exn.reraise exn)))); + fun reset () = (cache2 := empty; cache1 := empty); + fun total_size () = size (! cache1) + size (! cache2); + in {apply = apply, reset = reset, size = total_size} end; (* ML pretty-printing *) diff -r 9f8034d29365 -r fbb38db0435d src/Pure/Proof/extraction.ML --- a/src/Pure/Proof/extraction.ML Mon Jul 29 15:26:03 2024 +0200 +++ b/src/Pure/Proof/extraction.ML Mon Jul 29 15:26:56 2024 +0200 @@ -511,9 +511,9 @@ val procs = maps (rev o fst o snd) types; val rtypes = map fst types; val typroc = typeof_proc []; - fun expand_name ({thm_name, ...}: Proofterm.thm_header) = + fun expand_name ({thm_name = (thm_name, _), ...}: Proofterm.thm_header) = if Thm_Name.is_empty thm_name orelse member (op =) expand thm_name - then SOME Thm_Name.empty else NONE; + then SOME Thm_Name.none else NONE; val prep = the_default (K I) prep thy' o Proof_Rewrite_Rules.elim_defs thy' false (map (Thm.transfer thy) defs) o Proofterm.expand_proof thy' expand_name; @@ -623,7 +623,7 @@ | corr d vs ts Ts hs cs _ (prf0 as PThm (thm_header as {types = SOME Ts', ...}, thm_body)) _ defs = let - val {pos, theory_name, thm_name, prop, ...} = thm_header; + val {command_pos, theory_name, thm_name = (thm_name, thm_pos), prop, ...} = thm_header; val prf = Proofterm.thm_body_proof_open thm_body; val (vs', tye) = find_inst prop Ts ts vs; val shyps = mk_shyps tye; @@ -650,8 +650,8 @@ val corr_prf = mkabsp shyps corr_prf0; val corr_prop = Proofterm.prop_of corr_prf; val corr_header = - Proofterm.thm_header (serial ()) pos theory_name - (corr_name thm_name vs', 0) corr_prop + Proofterm.thm_header (serial ()) command_pos theory_name + ((corr_name thm_name vs', 0), thm_pos) corr_prop (SOME (map TVar (Term.add_tvars corr_prop [] |> rev))); val corr_prf' = Proofterm.proof_combP @@ -727,7 +727,7 @@ | extr d vs ts Ts hs (prf0 as PThm (thm_header as {types = SOME Ts', ...}, thm_body)) defs = let - val {pos, theory_name, thm_name, prop, ...} = thm_header; + val {command_pos, theory_name, thm_name = (thm_name, thm_pos), prop, ...} = thm_header; val prf = Proofterm.thm_body_proof_open thm_body; val (vs', tye) = find_inst prop Ts ts vs; val shyps = mk_shyps tye; @@ -772,8 +772,8 @@ SOME (map TVar (Term.add_tvars eqn [] |> rev))))) %% corr_prf); val corr_prop = Proofterm.prop_of corr_prf'; val corr_header = - Proofterm.thm_header (serial ()) pos theory_name - (corr_name thm_name vs', 0) corr_prop + Proofterm.thm_header (serial ()) command_pos theory_name + ((corr_name thm_name vs', 0), thm_pos) corr_prop (SOME (map TVar (Term.add_tvars corr_prop [] |> rev))); val corr_prf'' = Proofterm.proof_combP (Proofterm.proof_combt diff -r 9f8034d29365 -r fbb38db0435d src/Pure/Proof/proof_checker.ML --- a/src/Pure/Proof/proof_checker.ML Mon Jul 29 15:26:03 2024 +0200 +++ b/src/Pure/Proof/proof_checker.ML Mon Jul 29 15:26:56 2024 +0200 @@ -88,7 +88,7 @@ (Thm.forall_intr_vars (Thm.forall_intr_frees thm')) end; - fun thm_of _ _ (PThm ({thm_name, prop = prop', types = SOME Ts, ...}, _)) = + fun thm_of _ _ (PThm ({thm_name = (thm_name, thm_pos), prop = prop', types = SOME Ts, ...}, _)) = let val thm = Thm.unconstrainT (Drule.implies_intr_hyps (lookup thm_name)); val prop = Thm.prop_of thm; @@ -96,9 +96,9 @@ if is_equal prop prop' then () else error ("Duplicate use of theorem name " ^ - quote (Global_Theory.print_thm_name thy thm_name) ^ "\n" ^ - Syntax.string_of_term_global thy prop ^ "\n\n" ^ - Syntax.string_of_term_global thy prop'); + quote (Global_Theory.print_thm_name thy thm_name) ^ Position.here thm_pos + ^ "\n" ^ Syntax.string_of_term_global thy prop + ^ "\n\n" ^ Syntax.string_of_term_global thy prop'); in thm_of_atom thm Ts end | thm_of _ _ (PAxm (name, _, SOME Ts)) = diff -r 9f8034d29365 -r fbb38db0435d src/Pure/Proof/proof_rewrite_rules.ML --- a/src/Pure/Proof/proof_rewrite_rules.ML Mon Jul 29 15:26:03 2024 +0200 +++ b/src/Pure/Proof/proof_rewrite_rules.ML Mon Jul 29 15:26:56 2024 +0200 @@ -39,12 +39,12 @@ val equal_elim_axm = ax Proofterm.equal_elim_axm []; val symmetric_axm = ax Proofterm.symmetric_axm [propT]; - fun rew' (PThm ({thm_name = ("Pure.protectD", 0), ...}, _) % _ %% - (PThm ({thm_name = ("Pure.protectI", 0), ...}, _) % _ %% prf)) = SOME prf - | rew' (PThm ({thm_name = ("Pure.conjunctionD1", 0), ...}, _) % _ % _ %% - (PThm ({thm_name = ("Pure.conjunctionI", 0), ...}, _) % _ % _ %% prf %% _)) = SOME prf - | rew' (PThm ({thm_name = ("Pure.conjunctionD2", 0), ...}, _) % _ % _ %% - (PThm ({thm_name = ("Pure.conjunctionI", 0), ...}, _) % _ % _ %% _ %% prf)) = SOME prf + fun rew' (PThm ({thm_name = (("Pure.protectD", 0), _), ...}, _) % _ %% + (PThm ({thm_name = (("Pure.protectI", 0), _), ...}, _) % _ %% prf)) = SOME prf + | rew' (PThm ({thm_name = (("Pure.conjunctionD1", 0), _), ...}, _) % _ % _ %% + (PThm ({thm_name = (("Pure.conjunctionI", 0), _), ...}, _) % _ % _ %% prf %% _)) = SOME prf + | rew' (PThm ({thm_name = (("Pure.conjunctionD2", 0), _), ...}, _) % _ % _ %% + (PThm ({thm_name = (("Pure.conjunctionI", 0), _), ...}, _) % _ % _ %% _ %% prf)) = SOME prf | rew' (PAxm ("Pure.equal_elim", _, _) % _ % _ %% (PAxm ("Pure.equal_intr", _, _) % _ % _ %% prf %% _)) = SOME prf | rew' (PAxm ("Pure.symmetric", _, _) % _ % _ %% @@ -54,14 +54,14 @@ | rew' (PAxm ("Pure.equal_elim", _, _) % SOME (_ $ A) % SOME (_ $ B) %% (PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.prop", _)) % _ % _ % _ %% (PAxm ("Pure.reflexive", _, _) % _) %% prf1) %% - ((tg as PThm ({thm_name = ("Pure.protectI", 0), ...}, _)) % _ %% prf2)) = + ((tg as PThm ({thm_name = (("Pure.protectI", 0), _), ...}, _)) % _ %% prf2)) = SOME (tg %> B %% (equal_elim_axm %> A %> B %% prf1 %% prf2)) | rew' (PAxm ("Pure.equal_elim", _, _) % SOME (_ $ A) % SOME (_ $ B) %% (PAxm ("Pure.symmetric", _, _) % _ % _ %% (PAxm ("Pure.combination", _, _) % SOME (Const ("Pure.prop", _)) % _ % _ % _ %% (PAxm ("Pure.reflexive", _, _) % _) %% prf1)) %% - ((tg as PThm ({thm_name = ("Pure.protectI", 0), ...}, _)) % _ %% prf2)) = + ((tg as PThm ({thm_name = (("Pure.protectI", 0), _), ...}, _)) % _ %% prf2)) = SOME (tg %> B %% (equal_elim_axm %> A %> B %% (symmetric_axm % ?? B % ?? A %% prf1) %% prf2)) @@ -245,7 +245,7 @@ (AbsP (s, t, fst (insert_refl defs Ts prf)), false) | insert_refl defs Ts prf = (case Proofterm.strip_combt prf of - (PThm ({thm_name, prop, types = SOME Ts, ...}, _), ts) => + (PThm ({thm_name = (thm_name, _), prop, types = SOME Ts, ...}, _), ts) => if member (op =) defs thm_name then let val vs = vars_of prop; @@ -271,11 +271,11 @@ let val cnames = map (fst o dest_Const o fst) defs'; val expand = Proofterm.fold_proof_atoms true - (fn PThm ({serial, thm_name, prop, ...}, _) => + (fn PThm ({serial, thm_name = (thm_name, _), prop, ...}, _) => if member (op =) defnames thm_name orelse not (exists_Const (member (op =) cnames o #1) prop) then I - else Inttab.update (serial, Thm_Name.empty) + else Inttab.update (serial, Thm_Name.none) | _ => I) [prf] Inttab.empty; in Proofterm.expand_proof thy (Inttab.lookup expand o #serial) prf end else prf; diff -r 9f8034d29365 -r fbb38db0435d src/Pure/Proof/proof_syntax.ML --- a/src/Pure/Proof/proof_syntax.ML Mon Jul 29 15:26:03 2024 +0200 +++ b/src/Pure/Proof/proof_syntax.ML Mon Jul 29 15:26:56 2024 +0200 @@ -16,7 +16,7 @@ val pretty_proof: Proof.context -> Proofterm.proof -> Pretty.T val pretty_proof_boxes_of: Proof.context -> {full: bool, preproc: theory -> proof -> proof} -> thm -> Pretty.T - val standard_proof_of: {full: bool, expand_name: Proofterm.thm_header -> Thm_Name.T option} -> + val standard_proof_of: {full: bool, expand_name: Proofterm.thm_header -> Thm_Name.P option} -> thm -> Proofterm.proof val pretty_standard_proof_of: Proof.context -> bool -> thm -> Pretty.T end; @@ -104,7 +104,7 @@ let val U = Term.itselfT T --> propT in Const ("Pure.PClass", U --> proofT) $ Const (Logic.const_of_class c, U) end; -fun term_of _ (PThm ({serial = i, thm_name = a, types = Ts, ...}, _)) = +fun term_of _ (PThm ({serial = i, thm_name = (a, _), types = Ts, ...}, _)) = fold AppT (these Ts) (Const (Long_Name.append "thm" @@ -227,7 +227,7 @@ fun proof_syntax prf = let val thm_names = Symset.dest (Proofterm.fold_proof_atoms true - (fn PThm ({thm_name, ...}, _) => + (fn PThm ({thm_name = (thm_name, _), ...}, _) => if Thm_Name.is_empty thm_name then I else Symset.insert (Thm_Name.short thm_name) | _ => I) [prf] Symset.empty); val axm_names = Symset.dest (Proofterm.fold_proof_atoms true @@ -264,7 +264,7 @@ excluded = is_some o Global_Theory.lookup_thm_id thy} in Proofterm.proof_boxes selection [Thm.proof_of thm] - |> map (fn ({serial = i, pos, prop, ...}, proof) => + |> map (fn ({serial = i, command_pos, prop, thm_name = (_, thm_pos), ...}, proof) => let val proof' = proof |> Proofterm.reconstruct_proof thy prop @@ -276,7 +276,7 @@ val name = Long_Name.append "thm" (string_of_int i); in Pretty.item - [Pretty.str (name ^ Position.here_list pos ^ ":"), Pretty.brk 1, + [Pretty.str (name ^ Position.here_list [command_pos, thm_pos] ^ ":"), Pretty.brk 1, Syntax.pretty_term ctxt prop', Pretty.fbrk, pretty_proof ctxt proof'] end) |> Pretty.chunks diff -r 9f8034d29365 -r fbb38db0435d src/Pure/global_theory.ML --- a/src/Pure/global_theory.ML Mon Jul 29 15:26:03 2024 +0200 +++ b/src/Pure/global_theory.ML Mon Jul 29 15:26:56 2024 +0200 @@ -13,13 +13,13 @@ val defined_fact: theory -> string -> bool val alias_fact: binding -> string -> theory -> theory val hide_fact: bool -> string -> theory -> theory - val dest_thms: bool -> theory list -> theory -> (Thm_Name.T * thm) list + val dest_thms: bool -> theory list -> theory -> (Thm_Name.P * thm) list val pretty_thm_name: theory -> Thm_Name.T -> Pretty.T val print_thm_name: theory -> Thm_Name.T -> string - val get_thm_names: theory -> Thm_Name.T Inttab.table - val dest_thm_names: theory -> (Proofterm.thm_id * Thm_Name.T) list - val lookup_thm_id: theory -> Proofterm.thm_id -> Thm_Name.T option - val lookup_thm: theory -> thm -> (Proofterm.thm_id * Thm_Name.T) option + val get_thm_names: theory -> Thm_Name.P Inttab.table + val dest_thm_names: theory -> (Proofterm.thm_id * Thm_Name.P) list + val lookup_thm_id: theory -> Proofterm.thm_id -> Thm_Name.P option + val lookup_thm: theory -> thm -> (Proofterm.thm_id * Thm_Name.P) option val get_thms: theory -> xstring -> thm list val get_thm: theory -> xstring -> thm val transfer_theories: theory -> thm -> thm @@ -63,7 +63,7 @@ structure Data = Theory_Data ( - type T = Facts.T * Thm_Name.T Inttab.table lazy option; + type T = Facts.T * Thm_Name.P Inttab.table lazy option; val empty: T = (Facts.empty, NONE); fun merge ((facts1, _), (facts2, _)) = (Facts.merge (facts1, facts2), NONE); ); @@ -86,7 +86,7 @@ fun dest_thms verbose prev_thys thy = Facts.dest_static verbose (map facts_of prev_thys) (facts_of thy) - |> maps (fn (name, thms) => Thm_Name.list (name, Position.none) thms |> map (apfst fst)); + |> maps (fn (name, thms) => Thm_Name.list (name, Position.none) thms); fun pretty_thm_name thy = Facts.pretty_thm_name (Context.Theory thy) (facts_of thy); val print_thm_name = Pretty.string_of oo pretty_thm_name; @@ -99,7 +99,7 @@ fun make_thm_names thy = (dest_thms true (Theory.parents_of thy) thy, Inttab.empty) - |-> fold (fn (thm_name, thm) => fn thm_names => + |-> fold (fn ((thm_name, thm_pos), thm) => fn thm_names => (case Thm.derivation_id (Thm.transfer thy thm) of NONE => thm_names | SOME {serial, theory_name} => @@ -107,11 +107,11 @@ raise THM ("Bad theory name for derivation", 0, [thm]) else (case Inttab.lookup thm_names serial of - SOME thm_name' => + SOME (thm_name', thm_pos') => raise THM ("Duplicate use of derivation identity for " ^ - print_thm_name thy thm_name ^ " vs. " ^ - print_thm_name thy thm_name', 0, [thm]) - | NONE => Inttab.update (serial, thm_name) thm_names))); + print_thm_name thy thm_name ^ Position.here thm_pos ^ " vs. " ^ + print_thm_name thy thm_name' ^ Position.here thm_pos', 0, [thm]) + | NONE => Inttab.update (serial, (thm_name, thm_pos)) thm_names))); fun lazy_thm_names thy = (case thm_names_of thy of @@ -256,7 +256,7 @@ No_Name_Flags => thm | Name_Flags {post, official} => thm - |> (official andalso (post orelse Thm_Name.is_empty (Thm.raw_derivation_name thm))) ? + |> (official andalso (post orelse Thm_Name.is_empty (#1 (Thm.raw_derivation_name thm)))) ? Thm.name_derivation (name, pos) |> (not (Thm_Name.is_empty name) andalso (post orelse not (Thm.has_name_hint thm))) ? Thm.put_name_hint name)); diff -r 9f8034d29365 -r fbb38db0435d src/Pure/name.ML --- a/src/Pure/name.ML Mon Jul 29 15:26:03 2024 +0200 +++ b/src/Pure/name.ML Mon Jul 29 15:26:56 2024 +0200 @@ -26,11 +26,13 @@ val build_context: (context -> context) -> context val make_context: string list -> context val declare: string -> context -> context + val declare_renamed: string * string -> context -> context val is_declared: context -> string -> bool val invent: context -> string -> int -> string list val invent_names: context -> string -> 'a list -> (string * 'a) list val invent_list: string list -> string -> int -> string list val variant: string -> context -> string * context + val variant_bound: string -> context -> string * context val variant_list: string list -> string list -> string list val enforce_case: bool -> string -> string val desymbolize: bool option -> string -> string @@ -101,6 +103,9 @@ fun declare_renaming (x, x') (Context tab) = Context (Symtab.update (clean x, SOME (clean x')) tab); +fun declare_renamed (x, x') = + clean x <> clean x' ? declare_renaming (x, x') #> declare x'; + fun is_declared (Context tab) = Symtab.defined tab; fun declared (Context tab) = Symtab.lookup tab; @@ -144,12 +149,12 @@ let val x0 = Symbol.bump_init x; val x' = vary x0; - val ctxt' = ctxt - |> x0 <> x' ? declare_renaming (x0, x') - |> declare x'; + val ctxt' = ctxt |> declare_renamed (x0, x'); in (x', ctxt') end; in (x' ^ replicate_string n "_", ctxt') end; +fun variant_bound name = variant (if is_bound name then "u" else name); + fun variant_list used names = #1 (make_context used |> fold_map variant names); diff -r 9f8034d29365 -r fbb38db0435d src/Pure/proofterm.ML --- a/src/Pure/proofterm.ML Mon Jul 29 15:26:03 2024 +0200 +++ b/src/Pure/proofterm.ML Mon Jul 29 15:26:56 2024 +0200 @@ -9,7 +9,7 @@ signature PROOFTERM = sig type thm_header = - {serial: serial, pos: Position.T list, theory_name: string, thm_name: Thm_Name.T, + {serial: serial, command_pos: Position.T, theory_name: string, thm_name: Thm_Name.P, prop: term, types: typ list option} type thm_body type thm_node @@ -38,7 +38,7 @@ val proof_of: proof_body -> proof val join_proof: proof_body future -> proof val map_proof_of: (proof -> proof) -> proof_body -> proof_body - val thm_header: serial -> Position.T list -> string -> Thm_Name.T -> term -> typ list option -> + val thm_header: serial -> Position.T -> string -> Thm_Name.P -> term -> typ list option -> thm_header val thm_body: proof_body -> thm_body val thm_body_proof_raw: thm_body -> proof @@ -65,12 +65,13 @@ val no_thm_proofs: proof -> proof val no_body_proofs: proof -> proof - val encode: Consts.T -> proof XML.Encode.T - val encode_body: Consts.T -> proof_body XML.Encode.T - val encode_standard_term: Consts.T -> term XML.Encode.T - val encode_standard_proof: Consts.T -> proof XML.Encode.T - val decode: Consts.T -> proof XML.Decode.T - val decode_body: Consts.T -> proof_body XML.Decode.T + val proof_to_zproof: theory -> proof -> zproof + val encode_standard_term: theory -> term XML.Encode.T + val encode_standard_proof: theory -> proof XML.Encode.T + val encode: theory -> proof XML.Encode.T + val encode_body: theory -> proof_body XML.Encode.T + val decode: theory -> proof XML.Decode.T + val decode_body: theory -> proof_body XML.Decode.T val %> : proof * term -> proof @@ -170,13 +171,8 @@ val reconstruct_proof: theory -> term -> proof -> proof val prop_of': term list -> proof -> term val prop_of: proof -> term - val expand_name_empty: thm_header -> Thm_Name.T option - val expand_proof: theory -> (thm_header -> Thm_Name.T option) -> proof -> proof - - val standard_vars: Name.context -> term * proof option -> term * proof option - val standard_vars_term: Name.context -> term -> term - val add_standard_vars: proof -> (string * typ) list -> (string * typ) list - val add_standard_vars_term: term -> (string * typ) list -> (string * typ) list + val expand_name_empty: thm_header -> Thm_Name.P option + val expand_proof: theory -> (thm_header -> Thm_Name.P option) -> proof -> proof val export_enabled: unit -> bool val export_standard_enabled: unit -> bool @@ -188,8 +184,8 @@ val unconstrain_thm_proof: theory -> sorts_proof -> sort list -> term -> (serial * proof_body future) list -> proof_body -> term * proof_body val get_identity: sort list -> term list -> term -> proof -> - {serial: serial, theory_name: string, thm_name: Thm_Name.T} option - val get_approximative_name: sort list -> term list -> term -> proof -> Thm_Name.T + {serial: serial, theory_name: string, thm_name: Thm_Name.P} option + val get_approximative_name: sort list -> term list -> term -> proof -> Thm_Name.P type thm_id = {serial: serial, theory_name: string} val make_thm_id: serial * string -> thm_id val thm_header_id: thm_header -> thm_id @@ -205,8 +201,10 @@ (** datatype proof **) +val proof_serial = Counter.make (); + type thm_header = - {serial: serial, pos: Position.T list, theory_name: string, thm_name: Thm_Name.T, + {serial: serial, command_pos: Position.T, theory_name: string, thm_name: Thm_Name.P, prop: term, types: typ list option}; datatype proof = @@ -253,8 +251,8 @@ fun no_proof (PBody {oracles, thms, zboxes, zproof, ...}) = PBody {oracles = oracles, thms = thms, zboxes = zboxes, zproof = zproof, proof = MinProof}; -fun thm_header serial pos theory_name thm_name prop types : thm_header = - {serial = serial, pos = pos, theory_name = theory_name, thm_name = thm_name, +fun thm_header serial command_pos theory_name thm_name prop types : thm_header = + {serial = serial, command_pos = command_pos, theory_name = theory_name, thm_name = thm_name, prop = prop, types = types}; fun thm_body body = Thm_Body {open_proof = I, body = Future.value body}; @@ -290,7 +288,9 @@ val no_export = Lazy.value (); -fun make_thm ({serial, theory_name, thm_name, prop, ...}: thm_header) (Thm_Body {body, ...}) = +fun make_thm + ({serial, theory_name, thm_name = (thm_name, _), prop, ...}: thm_header) + (Thm_Body {body, ...}) = (serial, make_thm_node theory_name thm_name prop body no_export); @@ -342,8 +342,8 @@ | no_thm_names (AbsP (x, t, prf)) = AbsP (x, t, no_thm_names prf) | no_thm_names (prf % t) = no_thm_names prf % t | no_thm_names (prf1 %% prf2) = no_thm_names prf1 %% no_thm_names prf2 - | no_thm_names (PThm ({serial, pos, theory_name, thm_name = _, prop, types}, thm_body)) = - PThm (thm_header serial pos theory_name Thm_Name.empty prop types, thm_body) + | no_thm_names (PThm ({serial, command_pos, theory_name, thm_name = _, prop, types}, thm_body)) = + PThm (thm_header serial command_pos theory_name Thm_Name.none prop types, thm_body) | no_thm_names a = a; fun no_thm_proofs (Abst (x, T, prf)) = Abst (x, T, no_thm_proofs prf) @@ -369,6 +369,39 @@ (** XML data representation **) +(* encode as zterm/zproof *) + +fun proof_to_zproof thy = + let + val {ztyp, zterm} = ZTerm.zterm_cache thy; + val ztvar = ztyp #> (fn ZTVar v => v); + + fun zproof_const name prop typargs = + let + val vs = rev ((fold_types o fold_atyps) (insert (op =) o ztvar) prop []); + val Ts = map ztyp typargs + in ZConstp ((name, zterm prop), (ZTVars.make (vs ~~ Ts), ZVars.empty)) end; + + fun zproof MinProof = ZNop + | zproof (PBound i) = ZBoundp i + | zproof (Abst (a, SOME T, p)) = ZAbst (a, ztyp T, zproof p) + | zproof (AbsP (a, SOME t, p)) = ZAbsp (a, zterm t, zproof p) + | zproof (p % SOME t) = ZAppt (zproof p, zterm t) + | zproof (p %% q) = ZAppp (zproof p, zproof q) + | zproof (Hyp t) = ZHyp (zterm t) + | zproof (PAxm (a, prop, SOME Ts)) = zproof_const (ZAxiom a) prop Ts + | zproof (PClass (T, c)) = OFCLASSp (ztyp T, c) + | zproof (Oracle (a, prop, SOME Ts)) = zproof_const (ZOracle a) prop Ts + | zproof (PThm ({serial, theory_name, thm_name, prop, types = SOME Ts, ...}, _)) = + let val proof_name = + ZThm {theory_name = theory_name, thm_name = thm_name, serial = serial}; + in zproof_const proof_name prop Ts end; + in zproof end; + +fun encode_standard_term thy = ZTerm.zterm_of thy #> ZTerm.encode_zterm {typed_vars = false}; +fun encode_standard_proof thy = proof_to_zproof thy #> ZTerm.encode_zproof {typed_vars = false}; + + (* encode *) local @@ -386,11 +419,13 @@ fn PAxm (a, b, c) => ([a], pair (term consts) (option (list typ)) (b, c)), fn PClass (a, b) => ([b], typ a), fn Oracle (a, b, c) => ([a], pair (term consts) (option (list typ)) (b, c)), - fn PThm ({serial, pos, theory_name, thm_name, prop, types}, Thm_Body {open_proof, body, ...}) => - ([int_atom serial, theory_name, #1 thm_name, int_atom (#2 thm_name)], - pair (list properties) (pair (term consts) (pair (option (list typ)) (proof_body consts))) - (map Position.properties_of pos, - (prop, (types, map_proof_of open_proof (Future.join body)))))] + fn PThm ({serial, command_pos, theory_name, thm_name, prop, types}, Thm_Body {open_proof, body, ...}) => + ([int_atom serial, theory_name], + pair (properties o Position.properties_of) + (pair Thm_Name.encode_pos + (pair (term consts) + (pair (option (list typ)) (proof_body consts)))) + (command_pos, (thm_name, (prop, (types, map_proof_of open_proof (Future.join body))))))] and proof_body consts (PBody {oracles, thms, zboxes = _, zproof = _, proof = prf}) = triple (list (pair (pair string (properties o Position.properties_of)) (option (term consts)))) (list (thm consts)) (proof consts) (oracles, thms, prf) @@ -399,39 +434,10 @@ (a, (thm_node_theory_name thm_node, (thm_node_name thm_node, (thm_node_prop thm_node, (Future.join (thm_node_body thm_node)))))); -fun standard_term consts t = t |> variant - [fn Const (a, b) => ([a], list typ (Consts.typargs consts (a, b))), - fn Free (a, _) => ([a], []), - fn Var (a, _) => (indexname a, []), - fn Bound a => ([], int a), - fn Abs (a, b, c) => ([a], pair typ (standard_term consts) (b, c)), - fn t as op $ a => - if can Logic.match_of_class t then raise Match - else ([], pair (standard_term consts) (standard_term consts) a), - fn t => - let val (T, c) = Logic.match_of_class t - in ([c], typ T) end]; - -fun standard_proof consts prf = prf |> variant - [fn MinProof => ([], []), - fn PBound a => ([], int a), - fn Abst (a, SOME b, c) => ([a], pair typ (standard_proof consts) (b, c)), - fn AbsP (a, SOME b, c) => ([a], pair (standard_term consts) (standard_proof consts) (b, c)), - fn a % SOME b => ([], pair (standard_proof consts) (standard_term consts) (a, b)), - fn a %% b => ([], pair (standard_proof consts) (standard_proof consts) (a, b)), - fn Hyp a => ([], standard_term consts a), - fn PAxm (name, _, SOME Ts) => ([name], list typ Ts), - fn PClass (T, c) => ([c], typ T), - fn Oracle (name, prop, SOME Ts) => ([name], pair (standard_term consts) (list typ) (prop, Ts)), - fn PThm ({serial, theory_name, thm_name, types = SOME Ts, ...}, _) => - ([int_atom serial, theory_name, #1 thm_name, int_atom (#2 thm_name)], list typ Ts)]; - in -val encode = proof; -val encode_body = proof_body; -val encode_standard_term = standard_term; -val encode_standard_proof = standard_proof; +val encode = proof o Sign.consts_of; +val encode_body = proof_body o Sign.consts_of; end; @@ -453,12 +459,14 @@ fn ([a], b) => let val (c, d) = pair (term consts) (option (list typ)) b in PAxm (a, c, d) end, fn ([b], a) => PClass (typ a, b), fn ([a], b) => let val (c, d) = pair (term consts) (option (list typ)) b in Oracle (a, c, d) end, - fn ([a, b, c, d], e) => + fn ([ser, theory_name], b) => let - val ((x, (y, (z, w)))) = - pair (list properties) (pair (term consts) (pair (option (list typ)) (proof_body consts))) e; - val header = thm_header (int_atom a) (map Position.of_properties x) b (c, int_atom d) y z; - in PThm (header, thm_body w) end] + val ((command_pos, (thm_name, (prop, (types, body))))) = + pair (Position.of_properties o properties) + (pair Thm_Name.decode_pos + (pair (term consts) (pair (option (list typ)) (proof_body consts)))) b; + val header = thm_header (int_atom ser) command_pos theory_name thm_name prop types; + in PThm (header, thm_body body) end] and proof_body consts x = let val (a, b, c) = @@ -473,8 +481,8 @@ in -val decode = proof; -val decode_body = proof_body; +val decode = proof o Sign.consts_of; +val decode_body = proof_body o Sign.consts_of; end; @@ -562,8 +570,8 @@ | proof (PAxm (a, prop, SOME Ts)) = PAxm (a, prop, SOME (typs Ts)) | proof (PClass C) = ofclass C | proof (Oracle (a, prop, SOME Ts)) = Oracle (a, prop, SOME (typs Ts)) - | proof (PThm ({serial, pos, theory_name, thm_name, prop, types = SOME Ts}, thm_body)) = - PThm (thm_header serial pos theory_name thm_name prop (SOME (typs Ts)), thm_body) + | proof (PThm ({serial, command_pos, theory_name, thm_name, prop, types = SOME Ts}, thm_body)) = + PThm (thm_header serial command_pos theory_name thm_name prop (SOME (typs Ts)), thm_body) | proof _ = raise Same.SAME; in proof end; @@ -603,8 +611,8 @@ fun change_types types (PAxm (name, prop, _)) = PAxm (name, prop, types) | change_types (SOME [T]) (PClass (_, c)) = PClass (T, c) | change_types types (Oracle (name, prop, _)) = Oracle (name, prop, types) - | change_types types (PThm ({serial, pos, theory_name, thm_name, prop, types = _}, thm_body)) = - PThm (thm_header serial pos theory_name thm_name prop types, thm_body) + | change_types types (PThm ({serial, command_pos, theory_name, thm_name, prop, types = _}, thm_body)) = + PThm (thm_header serial command_pos theory_name thm_name prop types, thm_body) | change_types _ prf = prf; @@ -768,7 +776,7 @@ PClass (norm_type_same T, c) | norm (Oracle (a, prop, Ts)) = Oracle (a, prop, Same.map_option norm_types_same Ts) - | norm (PThm ({serial = i, pos = p, theory_name, thm_name, prop = t, types = Ts}, thm_body)) = + | norm (PThm ({serial = i, command_pos = p, theory_name, thm_name, prop = t, types = Ts}, thm_body)) = PThm (thm_header i p theory_name thm_name t (Same.map_option norm_types_same Ts), thm_body) | norm _ = raise Same.SAME; @@ -1029,8 +1037,8 @@ | proof _ _ (PAxm (a, prop, Ts)) = PAxm (a, prop, typs Ts) | proof _ _ (PClass (T, c)) = PClass (typ T, c) | proof _ _ (Oracle (a, prop, Ts)) = Oracle (a, prop, typs Ts) - | proof _ _ (PThm ({serial, pos, theory_name, thm_name, prop, types}, thm_body)) = - PThm (thm_header serial pos theory_name thm_name prop (typs types), thm_body) + | proof _ _ (PThm ({serial, command_pos, theory_name, thm_name, prop, types}, thm_body)) = + PThm (thm_header serial command_pos theory_name thm_name prop (typs types), thm_body) | proof _ _ _ = raise Same.SAME; val k = length prems; @@ -1452,7 +1460,7 @@ | subst _ _ (PAxm (id, prop, Ts)) = PAxm (id, prop, Same.map_option substTs Ts) | subst _ _ (PClass (T, c)) = PClass (substT T, c) | subst _ _ (Oracle (id, prop, Ts)) = Oracle (id, prop, Same.map_option substTs Ts) - | subst _ _ (PThm ({serial = i, pos = p, theory_name, thm_name, prop, types}, thm_body)) = + | subst _ _ (PThm ({serial = i, command_pos = p, theory_name, thm_name, prop, types}, thm_body)) = PThm (thm_header i p theory_name thm_name prop (Same.map_option substTs types), thm_body) | subst _ _ _ = raise Same.SAME; in fn t => subst 0 0 t handle Same.SAME => t end; @@ -1627,7 +1635,7 @@ | guess_name (prf %% PClass _) = guess_name prf | guess_name (prf % NONE) = guess_name prf | guess_name (prf % SOME (Var _)) = guess_name prf - | guess_name _ = Thm_Name.empty; + | guess_name _ = Thm_Name.none; (* generate constraints for proof term *) @@ -1739,7 +1747,8 @@ | SOME Ts => (Ts, env)); val prop' = subst_atomic_types (prop_types ~~ Ts) (forall_intr_variables_term prop) handle ListPair.UnequalLengths => - error ("Wrong number of type arguments for " ^ quote (Thm_Name.print (guess_name prf))) + error ("Wrong number of type arguments for " ^ + quote (Thm_Name.print_pos (guess_name prf))) in (prop', change_types (SOME Ts) prf, [], env', vTs) end; fun head_norm (prop, prf, cnstrts, env, vTs) = @@ -1927,7 +1936,7 @@ (* expand and reconstruct subproofs *) fun expand_name_empty (header: thm_header) = - if Thm_Name.is_empty (#thm_name header) then SOME Thm_Name.empty else NONE; + if Thm_Name.is_empty (#1 (#thm_name header)) then SOME Thm_Name.none else NONE; fun expand_proof thy expand_name prf = let @@ -1946,10 +1955,10 @@ let val (seen', maxidx', prf') = expand seen maxidx prf in (seen', maxidx', prf' % t) end | expand seen maxidx (prf as PThm (header, thm_body)) = - let val {serial, pos, theory_name, thm_name, prop, types} = header in + let val {serial, command_pos, theory_name, thm_name, prop, types} = header in (case expand_name header of SOME thm_name' => - if #1 thm_name' = "" andalso is_some types then + if #1 (#1 thm_name') = "" andalso is_some types then let val (seen', maxidx', prf') = (case Inttab.lookup seen serial of @@ -1968,7 +1977,7 @@ in (seen', maxidx' + maxidx + 1, prf'') end else if thm_name' <> thm_name then (seen, maxidx, - PThm (thm_header serial pos theory_name thm_name' prop types, thm_body)) + PThm (thm_header serial command_pos theory_name thm_name' prop types, thm_body)) else (seen, maxidx, prf) | NONE => (seen, maxidx, prf)) end @@ -2020,110 +2029,6 @@ (** theorems **) -(* standardization of variables for export: only frees and named bounds *) - -local - -val declare_names_term = Term.declare_term_frees; -val declare_names_term' = fn SOME t => declare_names_term t | NONE => I; -val declare_names_proof = fold_proof_terms declare_names_term; - -fun variant names bs x = - #1 (Name.variant x (fold Name.declare bs names)); - -fun variant_term bs (Abs (x, T, t)) = - let - val x' = variant (declare_names_term t Name.context) bs x; - val t' = variant_term (x' :: bs) t; - in Abs (x', T, t') end - | variant_term bs (t $ u) = variant_term bs t $ variant_term bs u - | variant_term _ t = t; - -fun variant_proof bs (Abst (x, T, prf)) = - let - val x' = variant (declare_names_proof prf Name.context) bs x; - val prf' = variant_proof (x' :: bs) prf; - in Abst (x', T, prf') end - | variant_proof bs (AbsP (x, t, prf)) = - let - val x' = variant (declare_names_term' t (declare_names_proof prf Name.context)) bs x; - val t' = Option.map (variant_term bs) t; - val prf' = variant_proof (x' :: bs) prf; - in AbsP (x', t', prf') end - | variant_proof bs (prf % t) = variant_proof bs prf % Option.map (variant_term bs) t - | variant_proof bs (prf1 %% prf2) = variant_proof bs prf1 %% variant_proof bs prf2 - | variant_proof bs (Hyp t) = Hyp (variant_term bs t) - | variant_proof _ prf = prf; - -val used_frees_type = fold_atyps (fn TFree (a, _) => Name.declare a | _ => I); -fun used_frees_term t = fold_types used_frees_type t #> Term.declare_term_frees t; -val used_frees_proof = fold_proof_terms_types used_frees_term used_frees_type; - -val unvarifyT = Term.map_atyps (fn TVar ((a, _), S) => TFree (a, S) | T => T); -val unvarify = Term.map_aterms (fn Var ((x, _), T) => Free (x, T) | t => t) #> map_types unvarifyT; -val unvarify_proof = map_proof_terms unvarify unvarifyT; - -fun hidden_types prop proof = - let - val visible = (fold_types o fold_atyps) (insert (op =)) prop []; - val add_hiddenT = fold_atyps (fn T => not (member (op =) visible T) ? insert (op =) T); - in rev (fold_proof_terms_types (fold_types add_hiddenT) add_hiddenT proof []) end; - -fun standard_hidden_types term proof = - let - val hidden = hidden_types term proof; - val idx = Term.maxidx_term term (maxidx_proof proof ~1) + 1; - fun smash T = - if member (op =) hidden T then - (case Type.sort_of_atyp T of - [] => dummyT - | S => TVar (("'", idx), S)) - else T; - val smashT = map_atyps smash; - in map_proof_terms (map_types smashT) smashT proof end; - -fun standard_hidden_terms term proof = - let - fun add_unless excluded x = - ((is_Free x orelse is_Var x) andalso not (member (op =) excluded x)) ? insert (op =) x; - val visible = fold_aterms (add_unless []) term []; - val hidden = fold_proof_terms (fold_aterms (add_unless visible)) proof []; - val dummy_term = Term.map_aterms (fn x => - if member (op =) hidden x then Term.dummy_pattern (Term.fastype_of x) else x); - in proof |> not (null hidden) ? map_proof_terms dummy_term I end; - -in - -fun standard_vars used (term, opt_proof) = - let - val proofs = opt_proof - |> Option.map (standard_hidden_types term #> standard_hidden_terms term) |> the_list; - val proof_terms = rev (fold (fold_proof_terms_types cons (cons o Logic.mk_type)) proofs []); - val used_frees = used - |> used_frees_term term - |> fold used_frees_proof proofs; - val inst = Term_Subst.zero_var_indexes_inst used_frees (term :: proof_terms); - val term' = term |> Term_Subst.instantiate inst |> unvarify |> variant_term []; - val proofs' = proofs |> map (instantiate inst #> unvarify_proof #> variant_proof []); - in (term', try hd proofs') end; - -fun standard_vars_term used t = #1 (standard_vars used (t, NONE)); - -val add_standard_vars_term = fold_aterms - (fn Free (x, T) => - (fn env => - (case AList.lookup (op =) env x of - NONE => (x, T) :: env - | SOME T' => - if T = T' then env - else raise TYPE ("standard_vars_env: type conflict for variable " ^ quote x, [T, T'], []))) - | _ => I); - -val add_standard_vars = fold_proof_terms add_standard_vars_term; - -end; - - (* PThm nodes *) fun prune_body body = @@ -2154,20 +2059,19 @@ local -fun export_proof thy i prop prf = +fun export_proof thy i prop0 proof0 = let - val (prop', SOME prf') = (prop, SOME prf) |> standard_vars Name.context; - val args = [] |> add_standard_vars_term prop' |> add_standard_vars prf' |> rev; - val typargs = [] |> Term.add_tfrees prop' |> fold_proof_terms Term.add_tfrees prf' |> rev; - - val consts = Sign.consts_of thy; - val xml = (typargs, (args, (prop', no_thm_names prf'))) |> + val {typargs, args, prop, proof} = + ZTerm.standard_vars Name.context + (ZTerm.zterm_of thy prop0, SOME (proof_to_zproof thy (no_thm_names proof0))); + val xml = (typargs, (args, (prop, the proof))) |> let open XML.Encode Term_XML.Encode; - val encode_vars = list (pair string typ); - val encode_term = encode_standard_term consts; - val encode_proof = encode_standard_proof consts; - in pair (list (pair string sort)) (pair encode_vars (pair encode_term encode_proof)) end; + val encode_tfrees = list (pair string sort); + val encode_frees = list (pair string ZTerm.encode_ztyp); + val encode_term = ZTerm.encode_zterm {typed_vars = false}; + val encode_proof = ZTerm.encode_zproof {typed_vars = false}; + in pair encode_tfrees (pair encode_frees (pair encode_term encode_proof)) end; in Export.export_params (Context.Theory thy) {theory_name = Context.theory_long_name thy, @@ -2199,7 +2103,7 @@ fun new_prf () = let - val i = serial (); + val i = proof_serial (); val unconstrainT = unconstrainT_proof (Sign.classes_of thy) sorts_proof ucontext; val postproc = map_proof_of (unconstrainT #> named ? rew_proof thy); val body1 = @@ -2211,11 +2115,11 @@ if export_enabled () then new_prf () else (case strip_combt (proof_head_of proof0) of - (PThm ({serial = ser, thm_name = a, prop = prop', types = NONE, ...}, thm_body'), args') => + (PThm ({serial = ser, thm_name = (a, _), prop = prop', types = NONE, ...}, thm_body'), args') => if (#1 a = "" orelse a = name) andalso prop' = prop1 andalso args' = args then let val Thm_Body {body = body', ...} = thm_body'; - val i = if #1 a = "" andalso named then serial () else ser; + val i = if #1 a = "" andalso named then proof_serial () else ser; in (i, body' |> ser <> i ? Future.map (map_proof_of (rew_proof thy))) end else new_prf () | _ => new_prf ()); @@ -2244,7 +2148,7 @@ val theory_name = Context.theory_long_name thy; val thm = (i, make_thm_node theory_name name prop1 thm_body export); - val header = thm_header i ([pos, Position.thread_data ()]) theory_name name prop1 NONE; + val header = thm_header i (Position.thread_data ()) theory_name (name, pos) prop1 NONE; val head = PThm (header, Thm_Body {open_proof = open_proof, body = thm_body}); val argst = if unconstrain then @@ -2283,7 +2187,7 @@ end; fun get_approximative_name shyps hyps prop prf = - Option.map #thm_name (get_identity shyps hyps prop prf) |> the_default Thm_Name.empty; + Option.map #thm_name (get_identity shyps hyps prop prf) |> the_default Thm_Name.none; (* thm_id *) @@ -2302,7 +2206,7 @@ fun get_id shyps hyps prop prf : thm_id option = (case get_identity shyps hyps prop prf of NONE => NONE - | SOME {serial, theory_name, thm_name, ...} => + | SOME {serial, theory_name, thm_name = (thm_name, _), ...} => if Thm_Name.is_empty thm_name then NONE else SOME (make_thm_id (serial, theory_name))); fun this_id NONE _ = false diff -r 9f8034d29365 -r fbb38db0435d src/Pure/term_items.ML --- a/src/Pure/term_items.ML Mon Jul 29 15:26:03 2024 +0200 +++ b/src/Pure/term_items.ML Mon Jul 29 15:26:56 2024 +0200 @@ -34,7 +34,8 @@ val make1: key * 'a -> 'a table val make2: key * 'a -> key * 'a -> 'a table val make3: key * 'a -> key * 'a -> key * 'a -> 'a table - val unsynchronized_cache: (key -> 'a) -> key -> 'a + type 'a cache_ops = {apply: key -> 'a, reset: unit -> unit, size: unit -> int}; + val unsynchronized_cache: (key -> 'a) -> 'a cache_ops type set = int table val add_set: key -> set -> set val make_set: key list -> set @@ -86,6 +87,7 @@ fun make2 e1 e2 = build (add e1 #> add e2); fun make3 e1 e2 e3 = build (add e1 #> add e2 #> add e3); +type 'a cache_ops = 'a Table.cache_ops; val unsynchronized_cache = Table.unsynchronized_cache; diff -r 9f8034d29365 -r fbb38db0435d src/Pure/term_subst.ML --- a/src/Pure/term_subst.ML Mon Jul 29 15:26:03 2024 +0200 +++ b/src/Pure/term_subst.ML Mon Jul 29 15:26:56 2024 +0200 @@ -230,8 +230,10 @@ (* zero var indexes *) fun zero_var_inst ins mk (v as ((x, i), X)) (inst, used) = - let val (x', used') = Name.variant (if Name.is_bound x then "u" else x) used - in if x = x' andalso i = 0 then (inst, used') else (ins (v, mk ((x', 0), X)) inst, used') end; + let + val (x', used') = Name.variant_bound x used; + val inst' = if x = x' andalso i = 0 then inst else ins (v, mk ((x', 0), X)) inst; + in (inst', used') end; fun zero_var_indexes_inst used ts = let diff -r 9f8034d29365 -r fbb38db0435d src/Pure/term_xml.ML --- a/src/Pure/term_xml.ML Mon Jul 29 15:26:03 2024 +0200 +++ b/src/Pure/term_xml.ML Mon Jul 29 15:26:56 2024 +0200 @@ -12,7 +12,6 @@ val sort: sort T val typ: typ T val term_raw: term T - val typ_body: typ T val term: Consts.T -> term T end @@ -47,12 +46,12 @@ fn Abs (a, b, c) => ([a], pair typ term_raw (b, c)), fn op $ a => ([], pair term_raw term_raw a)]; -fun typ_body T = if T = dummyT then [] else typ T; +fun var_type T = if T = dummyT then [] else typ T; fun term consts t = t |> variant [fn Const (a, b) => ([a], list typ (Consts.typargs consts (a, b))), - fn Free (a, b) => ([a], typ_body b), - fn Var (a, b) => (indexname a, typ_body b), + fn Free (a, b) => ([a], var_type b), + fn Var (a, b) => (indexname a, var_type b), fn Bound a => ([], int a), fn Abs (a, b, c) => ([a], pair typ (term consts) (b, c)), fn t as op $ a => @@ -87,13 +86,13 @@ fn ([a], b) => let val (c, d) = pair typ term_raw b in Abs (a, c, d) end, fn ([], a) => op $ (pair term_raw term_raw a)]; -fun typ_body [] = dummyT - | typ_body body = typ body; +fun var_type [] = dummyT + | var_type body = typ body; fun term consts body = body |> variant [fn ([a], b) => Const (a, Consts.instance consts (a, list typ b)), - fn ([a], b) => Free (a, typ_body b), - fn (a, b) => Var (indexname a, typ_body b), + fn ([a], b) => Free (a, var_type b), + fn (a, b) => Var (indexname a, var_type b), fn ([], a) => Bound (int a), fn ([a], b) => let val (c, d) = pair typ (term consts) b in Abs (a, c, d) end, fn ([], a) => op $ (pair (term consts) (term consts) a), diff -r 9f8034d29365 -r fbb38db0435d src/Pure/term_xml.scala --- a/src/Pure/term_xml.scala Mon Jul 29 15:26:03 2024 +0200 +++ b/src/Pure/term_xml.scala Mon Jul 29 15:26:56 2024 +0200 @@ -25,13 +25,13 @@ { case TFree(a, b) => (List(a), sort(b)) }, { case TVar(a, b) => (indexname(a), sort(b)) })) - val typ_body: T[Typ] = (t: Typ) => if (t == dummyT) Nil else typ(t) + private val var_type: T[Typ] = (t: Typ) => if (t == dummyT) Nil else typ(t) def term: T[Term] = variant[Term](List( { case Const(a, b) => (List(a), list(typ)(b)) }, - { case Free(a, b) => (List(a), typ_body(b)) }, - { case Var(a, b) => (indexname(a), typ_body(b)) }, + { case Free(a, b) => (List(a), var_type(b)) }, + { case Var(a, b) => (indexname(a), var_type(b)) }, { case Bound(a) => (Nil, int(a)) }, { case Abs(a, b, c) => (List(a), pair(typ, term)(b, c)) }, { case App(a, b) => (Nil, pair(term, term)(a, b)) }, @@ -53,13 +53,13 @@ { case (List(a), b) => TFree(a, sort(b)) }, { case (a, b) => TVar(indexname(a), sort(b)) })) - val typ_body: T[Typ] = { case Nil => dummyT case body => typ(body) } + private val var_type: T[Typ] = { case Nil => dummyT case body => typ(body) } def term: T[Term] = variant[Term](List( { case (List(a), b) => Const(a, list(typ)(b)) }, - { case (List(a), b) => Free(a, typ_body(b)) }, - { case (a, b) => Var(indexname(a), typ_body(b)) }, + { case (List(a), b) => Free(a, var_type(b)) }, + { case (a, b) => Var(indexname(a), var_type(b)) }, { case (Nil, a) => Bound(int(a)) }, { case (List(a), b) => val (c, d) = pair(typ, term)(b); Abs(a, c, d) }, { case (Nil, a) => val (b, c) = pair(term, term)(a); App(b, c) }, @@ -72,8 +72,8 @@ def term: T[Term] = variant[Term](List( { case (List(a), b) => Const(a, list(typ)(b)) }, - { case (List(a), b) => Free(a, env_type(a, typ_body(b))) }, - { case (a, b) => Var(indexname(a), typ_body(b)) }, + { case (List(a), b) => Free(a, env_type(a, var_type(b))) }, + { case (a, b) => Var(indexname(a), var_type(b)) }, { case (Nil, a) => Bound(int(a)) }, { case (List(a), b) => val (c, d) = pair(typ, term)(b); Abs(a, c, d) }, { case (Nil, a) => val (b, c) = pair(term, term)(a); App(b, c) }, diff -r 9f8034d29365 -r fbb38db0435d src/Pure/thm.ML --- a/src/Pure/thm.ML Mon Jul 29 15:26:03 2024 +0200 +++ b/src/Pure/thm.ML Mon Jul 29 15:26:56 2024 +0200 @@ -120,8 +120,8 @@ val derivation_closed: thm -> bool val derivation_name: thm -> Thm_Name.T val derivation_id: thm -> Proofterm.thm_id option - val raw_derivation_name: thm -> Thm_Name.T - val expand_name: thm -> Proofterm.thm_header -> Thm_Name.T option + val raw_derivation_name: thm -> Thm_Name.P + val expand_name: thm -> Proofterm.thm_header -> Thm_Name.P option val name_derivation: Thm_Name.P -> thm -> thm val close_derivation: Position.T -> thm -> thm val trim_context: thm -> thm @@ -1100,8 +1100,8 @@ SOME T' => T' | NONE => raise Fail "strip_shyps: bad type variable in proof term")); val map_ztyp = - ZTypes.unsynchronized_cache - (ZTerm.subst_type_same (ZTerm.ztyp_of o map_atyp o ZTerm.typ_of o ZTVar)); + #apply (ZTypes.unsynchronized_cache + (ZTerm.subst_type_same (ZTerm.ztyp_of o map_atyp o ZTerm.typ_of o ZTVar))); val zproof' = ZTerm.map_proof_types {hyps = true} map_ztyp zproof; val proof' = Proofterm.map_proof_types (Term.map_atyps_same map_atyp) proof; @@ -1131,13 +1131,13 @@ NONE => K false | SOME {serial, ...} => fn (header: Proofterm.thm_header) => serial = #serial header); fun expand header = - if self_id header orelse Thm_Name.is_empty (#thm_name header) - then SOME Thm_Name.empty else NONE; + if self_id header orelse Thm_Name.is_empty (#1 (#thm_name header)) + then SOME Thm_Name.none else NONE; in expand end; (*deterministic name of finished proof*) fun derivation_name (thm as Thm (_, {shyps, hyps, prop, ...})) = - Proofterm.get_approximative_name shyps hyps prop (proof_of thm); + #1 (Proofterm.get_approximative_name shyps hyps prop (proof_of thm)); (*identified PThm node*) fun derivation_id (thm as Thm (_, {shyps, hyps, prop, ...})) = diff -r 9f8034d29365 -r fbb38db0435d src/Pure/thm_deps.ML --- a/src/Pure/thm_deps.ML Mon Jul 29 15:26:03 2024 +0200 +++ b/src/Pure/thm_deps.ML Mon Jul 29 15:26:56 2024 +0200 @@ -59,7 +59,7 @@ else let val thm_id = Proofterm.thm_id (i, thm_node) in (case lookup thm_id of - SOME thm_name => + SOME (thm_name, _) => Inttab.update (i, SOME (thm_id, thm_name)) res | NONE => Inttab.update (i, NONE) res diff -r 9f8034d29365 -r fbb38db0435d src/Pure/thm_name.ML --- a/src/Pure/thm_name.ML Mon Jul 29 15:26:03 2024 +0200 +++ b/src/Pure/thm_name.ML Mon Jul 29 15:26:56 2024 +0200 @@ -26,9 +26,12 @@ val print_prefix: Context.generic -> Name_Space.T -> T -> Markup.T * string val print_suffix: T -> string val print: T -> string + val print_pos: P -> string val short: T -> string val encode: T XML.Encode.T val decode: T XML.Decode.T + val encode_pos: P XML.Encode.T + val decode_pos: P XML.Decode.T end; structure Thm_Name: THM_NAME = @@ -106,6 +109,8 @@ fun print (name, index) = name ^ print_suffix (name, index); +fun print_pos (thm_name, pos) = print thm_name ^ Position.here pos; + fun short (name, index) = if name = "" orelse index = 0 then name else name ^ "_" ^ string_of_int index; @@ -116,4 +121,7 @@ val encode = let open XML.Encode in pair string int end; val decode = let open XML.Decode in pair string int end; +val encode_pos = let open XML.Encode in pair encode (properties o Position.properties_of) end; +val decode_pos = let open XML.Decode in pair decode (Position.of_properties o properties) end; + end; diff -r 9f8034d29365 -r fbb38db0435d src/Pure/zterm.ML --- a/src/Pure/zterm.ML Mon Jul 29 15:26:03 2024 +0200 +++ b/src/Pure/zterm.ML Mon Jul 29 15:26:56 2024 +0200 @@ -52,6 +52,53 @@ | fold_types _ _ = I; +(* map *) + +fun map_tvars_same f = + let + fun typ (ZTVar v) = f v + | typ (ZFun (T, U)) = + (ZFun (typ T, Same.commit typ U) + handle Same.SAME => ZFun (T, typ U)) + | typ ZProp = raise Same.SAME + | typ (ZType0 _) = raise Same.SAME + | typ (ZType1 (a, T)) = ZType1 (a, typ T) + | typ (ZType (a, Ts)) = ZType (a, Same.map typ Ts); + in typ end; + +fun map_aterms_same f = + let + fun term (ZAbs (x, T, t)) = ZAbs (x, T, term t) + | term (ZApp (t, u)) = + (ZApp (term t, Same.commit term u) + handle Same.SAME => ZApp (t, term u)) + | term a = f a; + in term end; + +fun map_vars_same f = map_aterms_same (fn ZVar v => f v | _ => raise Same.SAME); + +fun map_types_same f = + let + fun term (ZVar (xi, T)) = ZVar (xi, f T) + | term (ZBound _) = raise Same.SAME + | term (ZConst0 _) = raise Same.SAME + | term (ZConst1 (c, T)) = ZConst1 (c, f T) + | term (ZConst (c, Ts)) = ZConst (c, Same.map f Ts) + | term (ZAbs (x, T, t)) = + (ZAbs (x, f T, Same.commit term t) + handle Same.SAME => ZAbs (x, T, term t)) + | term (ZApp (t, u)) = + (ZApp (term t, Same.commit term u) + handle Same.SAME => ZApp (t, term u)) + | term (OFCLASS (T, c)) = OFCLASS (f T, c); + in term end; + +val map_tvars = Same.commit o map_tvars_same; +val map_aterms = Same.commit o map_aterms_same; +val map_vars = Same.commit o map_vars_same; +val map_types = Same.commit o map_types_same; + + (* type ordering *) local @@ -224,6 +271,10 @@ val fold_aterms: (zterm -> 'a -> 'a) -> zterm -> 'a -> 'a val fold_vars: (indexname * ztyp -> 'a -> 'a) -> zterm -> 'a -> 'a val fold_types: (ztyp -> 'a -> 'a) -> zterm -> 'a -> 'a + val map_tvars: (indexname * sort -> ztyp) -> ztyp -> ztyp + val map_aterms: (zterm -> zterm) -> zterm -> zterm + val map_vars: (indexname * ztyp -> zterm) -> zterm -> zterm + val map_types: (ztyp -> ztyp) -> zterm -> zterm val ztyp_ord: ztyp ord val fast_zterm_ord: zterm ord val aconv_zterm: zterm * zterm -> bool @@ -231,6 +282,8 @@ val ZAbsps: zterm list -> zproof -> zproof val ZAppts: zproof * zterm list -> zproof val ZAppps: zproof * zproof list -> zproof + val strip_sortsT: ztyp -> ztyp + val strip_sorts: zterm -> zterm val incr_bv_same: int -> int -> zterm Same.operation val incr_proof_bv_same: int -> int -> int -> int -> zproof Same.operation val incr_bv: int -> int -> zterm -> zterm @@ -247,14 +300,28 @@ val subst_term_same: ztyp Same.operation -> (indexname * ztyp, zterm) Same.function -> zterm Same.operation exception BAD_INST of ((indexname * ztyp) * zterm) list + val fold_proof: {hyps: bool} -> (ztyp -> 'a -> 'a) -> (zterm -> 'a -> 'a) -> zproof -> 'a -> 'a + val fold_proof_types: {hyps: bool} -> (ztyp -> 'a -> 'a) -> zproof -> 'a -> 'a val map_proof: {hyps: bool} -> ztyp Same.operation -> zterm Same.operation -> zproof -> zproof val map_proof_types: {hyps: bool} -> ztyp Same.operation -> zproof -> zproof val map_class_proof: (ztyp * class, zproof) Same.function -> zproof -> zproof + val maxidx_type: ztyp -> int -> int + val maxidx_term: zterm -> int -> int + val maxidx_proof: zproof -> int -> int val ztyp_of: typ -> ztyp + val ztyp_dummy: ztyp + val typ_of_tvar: indexname * sort -> typ + val typ_of: ztyp -> typ + val init_cache: theory -> theory option + val exit_cache: theory -> theory option + val reset_cache: theory -> theory + val check_cache: theory -> {ztyp: int, typ: int} option + val ztyp_cache: theory -> typ -> ztyp + val typ_cache: theory -> ztyp -> typ val zterm_cache: theory -> {zterm: term -> zterm, ztyp: typ -> ztyp} val zterm_of: theory -> term -> zterm - val typ_of_tvar: indexname * sort -> typ - val typ_of: ztyp -> typ + val zterm_dummy_pattern: ztyp -> zterm + val zterm_type: ztyp -> zterm val term_of: theory -> zterm -> term val beta_norm_term_same: zterm Same.operation val beta_norm_proof_same: zproof Same.operation @@ -262,7 +329,7 @@ val beta_norm_term: zterm -> zterm val beta_norm_proof: zproof -> zproof val beta_norm_prooft: zproof -> zproof - val norm_type: Type.tyenv -> ztyp -> ztyp + val norm_type: theory -> Type.tyenv -> ztyp -> ztyp val norm_term: theory -> Envir.env -> zterm -> zterm val norm_proof: theory -> Envir.env -> term list -> zproof -> zproof val zproof_const_typargs: zproof_const -> ((indexname * sort) * ztyp) list @@ -310,6 +377,11 @@ val of_sort_proof: Sorts.algebra -> sorts_proof -> (typ * class -> zproof) -> typ * sort -> zproof list val unconstrainT_proof: theory -> sorts_proof -> Logic.unconstrain_context -> zproof -> zproof + val encode_ztyp: ztyp XML.Encode.T + val encode_zterm: {typed_vars: bool} -> zterm XML.Encode.T + val encode_zproof: {typed_vars: bool} -> zproof XML.Encode.T + val standard_vars: Name.context -> zterm * zproof option -> + {typargs: (string * sort) list, args: (string * ztyp) list, prop: zterm, proof: zproof option} end; structure ZTerm: ZTERM = @@ -337,6 +409,12 @@ fun combound (t, n, k) = if k > 0 then ZApp (combound (t, n + 1, k - 1), ZBound n) else t; +val strip_sortsT_same = map_tvars_same (fn (_, []) => raise Same.SAME | (a, _) => ZTVar (a, [])); +val strip_sorts_same = map_types_same strip_sortsT_same; + +val strip_sortsT = Same.commit strip_sortsT_same; +val strip_sorts = Same.commit strip_sorts_same; + (* increment bound variables *) @@ -470,7 +548,7 @@ fun instantiate_type_same instT = if ZTVars.is_empty instT then Same.same - else ZTypes.unsynchronized_cache (subst_type_same (Same.function (ZTVars.lookup instT))); + else #apply (ZTypes.unsynchronized_cache (subst_type_same (Same.function (ZTVars.lookup instT)))); fun instantiate_term_same typ inst = subst_term_same typ (Same.function (ZVars.lookup inst)); @@ -577,7 +655,18 @@ in Same.commit proof end; -(* convert ztyp / zterm vs. regular typ / term *) +(* maximum index of variables *) + +val maxidx_type = fold_tvars (fn ((_, i), _) => Integer.max i); + +fun maxidx_term t = + fold_types maxidx_type t #> + fold_aterms (fn ZVar ((_, i), _) => Integer.max i | _ => I) t; + +val maxidx_proof = fold_proof {hyps = false} maxidx_type maxidx_term; + + +(* convert ztyp vs. regular typ *) fun ztyp_of (TFree (a, S)) = ZTVar ((a, ~1), S) | ztyp_of (TVar v) = ZTVar v @@ -587,13 +676,75 @@ | ztyp_of (Type (c, [T])) = ZType1 (c, ztyp_of T) | ztyp_of (Type (c, ts)) = ZType (c, map ztyp_of ts); -fun ztyp_cache () = Typtab.unsynchronized_cache ztyp_of; +val ztyp_dummy = ztyp_of dummyT; + +fun typ_of_tvar ((a, ~1), S) = TFree (a, S) + | typ_of_tvar v = TVar v; + +fun typ_of (ZTVar v) = typ_of_tvar v + | typ_of (ZFun (T, U)) = typ_of T --> typ_of U + | typ_of ZProp = propT + | typ_of (ZType0 c) = Type (c, []) + | typ_of (ZType1 (c, T)) = Type (c, [typ_of T]) + | typ_of (ZType (c, Ts)) = Type (c, map typ_of Ts); + + +(* cache within theory context *) + +local + +structure Data = Theory_Data +( + type T = (ztyp Typtab.cache_ops * typ ZTypes.cache_ops) option; + val empty = NONE; + val merge = K NONE; +); + +fun make_ztyp_cache () = Typtab.unsynchronized_cache ztyp_of; +fun make_typ_cache () = ZTypes.unsynchronized_cache typ_of; + +in + +fun init_cache thy = + if is_some (Data.get thy) then NONE + else SOME (Data.put (SOME (make_ztyp_cache (), make_typ_cache ())) thy); + +fun exit_cache thy = + (case Data.get thy of + SOME (cache1, cache2) => (#reset cache1 (); #reset cache2 (); SOME (Data.put NONE thy)) + | NONE => NONE); + +val _ = Theory.setup (Theory.at_begin init_cache #> Theory.at_end exit_cache); + +fun reset_cache thy = + (case Data.get thy of + SOME (cache1, cache2) => (#reset cache1 (); #reset cache2 (); thy) + | NONE => thy); + +fun check_cache thy = + Data.get thy + |> Option.map (fn (cache1, cache2) => {ztyp = #size cache1 (), typ = #size cache2 ()}); + +fun ztyp_cache thy = + (case Data.get thy of + SOME (cache, _) => cache + | NONE => make_ztyp_cache ()) |> #apply; + +fun typ_cache thy = + (case Data.get thy of + SOME (_, cache) => cache + | NONE => make_typ_cache ()) |> #apply; + +end; + + +(* convert zterm vs. regular term *) fun zterm_cache thy = let val typargs = Consts.typargs (Sign.consts_of thy); - val ztyp = ztyp_cache (); + val ztyp = ztyp_cache thy; fun zterm (Free (x, T)) = ZVar ((x, ~1), ztyp T) | zterm (Var (xi, T)) = ZVar (xi, ztyp T) @@ -612,17 +763,8 @@ val zterm_of = #zterm o zterm_cache; -fun typ_of_tvar ((a, ~1), S) = TFree (a, S) - | typ_of_tvar v = TVar v; - -fun typ_of (ZTVar v) = typ_of_tvar v - | typ_of (ZFun (T, U)) = typ_of T --> typ_of U - | typ_of ZProp = propT - | typ_of (ZType0 c) = Type (c, []) - | typ_of (ZType1 (c, T)) = Type (c, [typ_of T]) - | typ_of (ZType (c, Ts)) = Type (c, map typ_of Ts); - -fun typ_cache () = ZTypes.unsynchronized_cache typ_of; +fun zterm_dummy_pattern T = ZConst1 ("Pure.dummy_pattern", T); +fun zterm_type T = ZConst1 ("Pure.type", T); fun term_of thy = let @@ -717,11 +859,13 @@ | norm (ZType (a, Ts)) = ZType (a, Same.map norm Ts); in norm end; -fun norm_term_same {ztyp, zterm, typ} (envir as Envir.Envir {tyenv, tenv, ...}) = +fun norm_term_same {ztyp, zterm} (envir as Envir.Envir {tyenv, tenv, ...}) = let val lookup = if Vartab.is_empty tenv then K NONE - else ZVars.unsynchronized_cache (apsnd typ #> Envir.lookup envir #> Option.map zterm); + else + #apply (ZVars.unsynchronized_cache + (apsnd typ_of #> Envir.lookup envir #> Option.map zterm)); val normT = norm_type_same ztyp tyenv; @@ -776,15 +920,9 @@ in beta_norm_prooft (map_proof {hyps = false} inst_typ norm_term prf) end end; -fun norm_cache thy = - let - val {ztyp, zterm} = zterm_cache thy; - val typ = typ_cache (); - in {ztyp = ztyp, zterm = zterm, typ = typ} end; - -fun norm_type tyenv = Same.commit (norm_type_same (ztyp_cache ()) tyenv); -fun norm_term thy envir = Same.commit (norm_term_same (norm_cache thy) envir); -fun norm_proof thy envir = norm_proof_cache (norm_cache thy) envir; +fun norm_type thy tyenv = Same.commit (norm_type_same (ztyp_cache thy) tyenv); +fun norm_term thy envir = Same.commit (norm_term_same (zterm_cache thy) envir); +fun norm_proof thy envir = norm_proof_cache (zterm_cache thy) envir; @@ -870,7 +1008,8 @@ val typ_operation = #typ_operation ucontext {strip_sorts = true}; val unconstrain_typ = Same.commit typ_operation; val unconstrain_ztyp = - ZTypes.unsynchronized_cache (Same.function_eq (op =) (typ_of #> unconstrain_typ #> ztyp_of)); + #apply (ZTypes.unsynchronized_cache + (Same.function_eq (op =) (typ_of #> unconstrain_typ #> ztyp_of))); val unconstrain_zterm = zterm o Term.map_types typ_operation; val unconstrain_proof = Same.commit (map_proof_types {hyps = true} unconstrain_ztyp); @@ -1087,10 +1226,10 @@ val typ = if Names.is_empty tfrees then Same.same else - ZTypes.unsynchronized_cache + #apply (ZTypes.unsynchronized_cache (subst_type_same (fn ((a, i), S) => if i = ~1 andalso Names.defined tfrees a then ZTVar ((a, idx), S) - else raise Same.SAME)); + else raise Same.SAME))); val term = subst_term_same typ (fn ((x, i), T) => if i = ~1 andalso Names.defined frees x then ZVar ((x, idx), T) @@ -1112,10 +1251,10 @@ let val tab = ZTVars.build (names |> fold (fn ((a, S), b) => ZTVars.add (((a, ~1), S), b))); val typ = - ZTypes.unsynchronized_cache (subst_type_same (fn v => + #apply (ZTypes.unsynchronized_cache (subst_type_same (fn v => (case ZTVars.lookup tab v of NONE => raise Same.SAME - | SOME w => ZTVar w))); + | SOME w => ZTVar w)))); in map_proof_types {hyps = false} typ prf end; fun legacy_freezeT_proof t prf = @@ -1124,7 +1263,7 @@ | SOME f => let val tvar = ztyp_of o Same.function f; - val typ = ZTypes.unsynchronized_cache (subst_type_same tvar); + val typ = #apply (ZTypes.unsynchronized_cache (subst_type_same tvar)); in map_proof_types {hyps = false} typ prf end); @@ -1158,7 +1297,7 @@ if inc = 0 then Same.same else let fun tvar ((a, i), S) = if i >= 0 then ZTVar ((a, i + inc), S) else raise Same.SAME - in ZTypes.unsynchronized_cache (subst_type_same tvar) end; + in #apply (ZTypes.unsynchronized_cache (subst_type_same tvar)) end; fun incr_indexes_proof inc prf = let @@ -1240,7 +1379,7 @@ fun assumption_proof thy envir Bs Bi n visible prf = let - val cache as {zterm, ...} = norm_cache thy; + val cache as {zterm, ...} = zterm_cache thy; val normt = zterm #> Same.commit (norm_term_same cache envir); in ZAbsps (map normt Bs) @@ -1258,7 +1397,7 @@ fun bicompose_proof thy env smax flatten Bs As A oldAs n m visible rprf sprf = let - val cache as {zterm, ...} = norm_cache thy; + val cache as {zterm, ...} = zterm_cache thy; val normt = zterm #> Same.commit (norm_term_same cache env); val normp = norm_proof_cache cache env visible; @@ -1292,12 +1431,14 @@ fun unconstrainT_proof thy sorts_proof (ucontext: Logic.unconstrain_context) = let - val cache = norm_cache thy; val algebra = Sign.classes_of thy; + val cache = zterm_cache thy; + val typ_cache = typ_cache thy; + val typ = - ZTypes.unsynchronized_cache - (typ_of #> #typ_operation ucontext {strip_sorts = true} #> ztyp_of); + #apply (ZTypes.unsynchronized_cache + (typ_of #> #typ_operation ucontext {strip_sorts = true} #> ztyp_of)); val typ_sort = #typ_operation ucontext {strip_sorts = false}; @@ -1307,11 +1448,206 @@ | NONE => raise Fail "unconstrainT_proof: missing constraint"); fun class (T, c) = - let val T' = Same.commit typ_sort (#typ cache T) + let val T' = Same.commit typ_sort (typ_cache T) in the_single (of_sort_proof algebra sorts_proof hyps (T', [c])) end; in map_proof_types {hyps = false} typ #> map_class_proof class #> beta_norm_prooft #> fold_rev (implies_intr_proof' o #zterm cache o #2) (#constraints ucontext) end; + + +(** XML data representation **) + +(* encode *) + +local + +open XML.Encode Term_XML.Encode; + +fun ztyp T = T |> variant + [fn ZFun args => (["fun"], pair ztyp ztyp args) + | ZProp => (["prop"], []) + | ZType0 a => ([a], []) + | ZType1 (a, b) => ([a], list ztyp [b]) + | ZType (a, b) => ([a], list ztyp b), + fn ZTVar ((a, ~1), b) => ([a], sort b), + fn ZTVar (a, b) => (indexname a, sort b)]; + +fun zvar_type {typed_vars} T = + if typed_vars andalso T <> ztyp_dummy then ztyp T else []; + +fun zterm flag t = t |> variant + [fn ZConst0 a => ([a], []) + | ZConst1 (a, b) => ([a], list ztyp [b]) + | ZConst (a, b) => ([a], list ztyp b), + fn ZVar ((a, ~1), b) => ([a], zvar_type flag b), + fn ZVar (a, b) => (indexname a, zvar_type flag b), + fn ZBound a => ([], int a), + fn ZAbs (a, b, c) => ([a], pair ztyp (zterm flag) (b, c)), + fn ZApp a => ([], pair (zterm flag) (zterm flag) a), + fn OFCLASS (a, b) => ([b], ztyp a)]; + +fun zproof flag prf = prf |> variant + [fn ZNop => ([], []), + fn ZBoundp a => ([], int a), + fn ZAbst (a, b, c) => ([a], pair ztyp (zproof flag) (b, c)), + fn ZAbsp (a, b, c) => ([a], pair (zterm flag) (zproof flag) (b, c)), + fn ZAppt a => ([], pair (zproof flag) (zterm flag) a), + fn ZAppp a => ([], pair (zproof flag) (zproof flag) a), + fn ZHyp a => ([], zterm flag a), + fn ZConstp (c as ((ZAxiom a, b), _)) => ([a], list (ztyp o #2) (zproof_const_typargs c)), + fn OFCLASSp (a, b) => ([b], ztyp a), + fn ZConstp (c as ((ZOracle a, b), _)) => ([a], list (ztyp o #2) (zproof_const_typargs c)), + fn ZConstp (c as ((ZThm {theory_name, thm_name = (name, _), serial}, b), _)) => + ([int_atom serial, theory_name, #1 name, int_atom (#2 name)], + list (ztyp o #2) (zproof_const_typargs c))]; + +in + +val encode_ztyp = ztyp; +val encode_zterm = zterm; +val encode_zproof = zproof; + end; + + +(* standardization of variables for export: only frees and named bounds *) + +local + +fun declare_frees_term t = fold_vars (fn ((x, ~1), _) => Name.declare x | _ => I) t; +val declare_frees_proof = fold_proof {hyps = true} (K I) declare_frees_term; + +val (variant_abs_term, variant_abs_proof) = + let + fun term bs (ZAbs (x, T, t)) = + let + val x' = #1 (Name.variant x (declare_frees_term t bs)); + val bs' = Name.declare x' bs; + in ZAbs (x', T, Same.commit_if (x <> x') (term bs') t) end + | term bs (ZApp (t, u)) = + (ZApp (term bs t, Same.commit (term bs) u) + handle Same.SAME => ZApp (t, term bs u)) + | term _ _ = raise Same.SAME; + + fun proof bs (ZAbst (x, T, p)) = + let + val x' = #1 (Name.variant x (declare_frees_proof p bs)); + val bs' = Name.declare x' bs; + in ZAbst (x', T, Same.commit_if (x <> x') (proof bs') p) end + | proof bs (ZAbsp (x, t, p)) = + let + val x' = #1 (Name.variant x (declare_frees_term t (declare_frees_proof p bs))); + val (t', term_eq) = Same.commit_id (term bs) t; + val bs' = Name.declare x' bs; + in ZAbsp (x', t', Same.commit_if (x <> x' orelse not term_eq) (proof bs') p) end + | proof bs (ZAppt (p, t)) = + (ZAppt (proof bs p, Same.commit (term bs) t) + handle Same.SAME => ZAppt (p, term bs t)) + | proof bs (ZAppp (p, q)) = + (ZAppp (proof bs p, Same.commit (proof bs) q) + handle Same.SAME => ZAppp (p, proof bs q)) + | proof bs (ZHyp t) = ZHyp (term bs t) + | proof _ _ = raise Same.SAME; + in (term Name.context, proof Name.context) end; + +val used_frees_type = fold_tvars (fn ((a, ~1), _) => Name.declare a | _ => I); +fun used_frees_term t = fold_types used_frees_type t #> declare_frees_term t; +val used_frees_proof = fold_proof {hyps = true} used_frees_type used_frees_term; + +fun hidden_types prop proof = + let + val visible = (fold_types o fold_tvars) (insert (op =)) prop []; + val add_hiddenT = fold_tvars (fn v => not (member (op =) visible v) ? insert (op =) v); + in rev (fold_proof {hyps = true} add_hiddenT (fold_types add_hiddenT) proof []) end; + +fun standard_hidden_types prop proof = + let + val hidden = hidden_types prop proof; + val idx = maxidx_term prop (maxidx_proof proof ~1) + 1; + fun zvar (v as (_, S)) = + if not (member (op =) hidden v) then raise Same.SAME + else if null S then ztyp_dummy + else ZTVar (("'", idx), S); + val typ = map_tvars_same zvar; + in proof |> not (null hidden) ? map_proof {hyps = true} typ (map_types typ) end; + +fun standard_hidden_terms prop proof = + let + fun add_unless excluded (ZVar v) = not (member (op =) excluded v) ? insert (op =) v + | add_unless _ _ = I; + val visible = fold_aterms (add_unless []) prop []; + val hidden = fold_proof {hyps = true} (K I) (fold_aterms (add_unless visible)) proof []; + fun var (v as (_, T)) = + if member (op =) hidden v then zterm_dummy_pattern T else raise Same.SAME; + val term = map_vars_same var; + in proof |> not (null hidden) ? map_proof {hyps = true} Same.same term end; + +fun standard_inst add mk (v as ((x, i), T)) (inst, used) = + let + val (x', used') = Name.variant_bound x used; + val inst' = if x = x' andalso i = ~1 then inst else add (v, mk ((x', ~1), T)) inst; + in (inst', used') end; + +fun standard_inst_type used prf = + let + val add_tvars = fold_tvars (fn v => ZTVars.add (v, ())); + val (instT, _) = + (ZTVars.empty, used) |> ZTVars.fold (standard_inst ZTVars.add ZTVar o #1) + (TVars.build (prf |> fold_proof {hyps = true} add_tvars (fold_types add_tvars))); + in instantiate_type_same instT end; + +fun standard_inst_term used inst_type prf = + let + val add_vars = fold_vars (fn (x, T) => ZVars.add ((x, Same.commit inst_type T), ())); + val (inst, _) = + (ZVars.empty, used) |> ZVars.fold (standard_inst ZVars.add ZVar o #1) + (ZVars.build (prf |> fold_proof {hyps = true} (K I) add_vars)); + in instantiate_term_same inst_type inst end; + +val typargs_type = fold_tvars (fn ((a, ~1), S) => TFrees.add_set (a, S) | _ => I); +val typargs_term = fold_types typargs_type; +val typargs_proof = fold_proof {hyps = true} typargs_type typargs_term; + +val add_standard_vars_term = fold_aterms + (fn ZVar ((x, ~1), T) => + (fn env => + (case AList.lookup (op =) env x of + NONE => (x, T) :: env + | SOME T' => + if T = T' then env + else + raise TYPE ("standard_vars_env: type conflict for variable " ^ quote x, + [typ_of T, typ_of T'], []))) + | _ => I); + +val add_standard_vars = fold_proof {hyps = true} (K I) add_standard_vars_term; + +in + +fun standard_vars used (prop, opt_prf) = + let + val prf = the_default ZNop opt_prf + |> standard_hidden_types prop + |> standard_hidden_terms prop; + val prop_prf = ZAppp (ZHyp prop, prf); + + val used' = used |> used_frees_proof prop_prf; + val inst_type = standard_inst_type used' prop_prf; + val inst_term = standard_inst_term used' inst_type prop_prf; + val inst_proof = map_proof_same {hyps = true} inst_type inst_term; + + val prop' = prop |> Same.commit (Same.compose variant_abs_term inst_term); + val opt_proof' = + if is_none opt_prf then NONE + else SOME (prf |> Same.commit (Same.compose variant_abs_proof inst_proof)); + val proofs' = the_list opt_proof'; + + val args = build_rev (add_standard_vars_term prop' #> fold add_standard_vars proofs'); + val typargs = TFrees.list_set (TFrees.build (typargs_term prop' #> fold typargs_proof proofs')); + in {typargs = typargs, args = args, prop = prop', proof = opt_proof'} end; + +end; + +end; diff -r 9f8034d29365 -r fbb38db0435d src/Tools/Haskell/Haskell.thy --- a/src/Tools/Haskell/Haskell.thy Mon Jul 29 15:26:03 2024 +0200 +++ b/src/Tools/Haskell/Haskell.thy Mon Jul 29 15:26:56 2024 +0200 @@ -2175,7 +2175,7 @@ Name, uu, uu_, aT, clean_index, clean, internal, skolem, is_internal, is_skolem, dest_internal, dest_skolem, - Context, declare, declare_renaming, is_declared, declared, context, make_context, + Context, declare, declare_renamed, is_declared, declared, context, make_context, variant, variant_list ) where @@ -2242,6 +2242,10 @@ declare_renaming (x, x') (Context names) = Context (Map.insert (clean x) (Just (clean x')) names) +declare_renamed :: (Name, Name) -> Context -> Context +declare_renamed (x, x') = + (if clean x /= clean x' then declare_renaming (x, x') else id) #> declare x' + is_declared :: Context -> Name -> Bool is_declared (Context names) x = Map.member x names @@ -2289,9 +2293,7 @@ let x0 = bump_init x x' = vary x0 - ctxt' = ctxt - |> (if x0 /= x' then declare_renaming (x0, x') else id) - |> declare x' + ctxt' = ctxt |> declare_renamed (x0, x') in (x', ctxt') in (x' <> Bytes.pack (replicate n underscore), ctxt') @@ -2651,7 +2653,7 @@ {-# LANGUAGE LambdaCase #-} -module Isabelle.Term_XML.Encode (indexname, sort, typ, typ_body, term) +module Isabelle.Term_XML.Encode (indexname, sort, typ, term) where import Isabelle.Library @@ -2671,15 +2673,15 @@ \case { TFree (a, b) -> Just ([a], sort b); _ -> Nothing }, \case { TVar (a, b) -> Just (indexname a, sort b); _ -> Nothing }] -typ_body :: T Typ -typ_body ty = if is_dummyT ty then [] else typ ty +var_type :: T Typ +var_type ty = if is_dummyT ty then [] else typ ty term :: T Term term t = t |> variant [\case { Const (a, b) -> Just ([a], list typ b); _ -> Nothing }, - \case { Free (a, b) -> Just ([a], typ_body b); _ -> Nothing }, - \case { Var (a, b) -> Just (indexname a, typ_body b); _ -> Nothing }, + \case { Free (a, b) -> Just ([a], var_type b); _ -> Nothing }, + \case { Var (a, b) -> Just (indexname a, var_type b); _ -> Nothing }, \case { Bound a -> Just ([], int a); _ -> Nothing }, \case { Abs (a, b, c) -> Just ([a], pair typ term (b, c)); _ -> Nothing }, \case { App a -> Just ([], pair term term a); _ -> Nothing }, @@ -2698,7 +2700,7 @@ {-# OPTIONS_GHC -fno-warn-incomplete-patterns #-} -module Isabelle.Term_XML.Decode (indexname, sort, typ, typ_body, term) +module Isabelle.Term_XML.Decode (indexname, sort, typ, term) where import Isabelle.Library @@ -2720,16 +2722,16 @@ \([a], b) -> TFree (a, sort b), \(a, b) -> TVar (indexname a, sort b)] -typ_body :: T Typ -typ_body [] = dummyT -typ_body body = typ body +var_type :: T Typ +var_type [] = dummyT +var_type body = typ body term :: T Term term t = t |> variant [\([a], b) -> Const (a, list typ b), - \([a], b) -> Free (a, typ_body b), - \(a, b) -> Var (indexname a, typ_body b), + \([a], b) -> Free (a, var_type b), + \(a, b) -> Var (indexname a, var_type b), \([], a) -> Bound (int a), \([a], b) -> let (c, d) = pair typ term b in Abs (a, c, d), \([], a) -> App (pair term term a),