--- a/src/HOL/Library/BNF_Corec.thy Sun Nov 03 21:12:50 2024 +0100
+++ b/src/HOL/Library/BNF_Corec.thy Sun Nov 03 22:29:07 2024 +0100
@@ -139,11 +139,9 @@
lemma cong_gen_cong: "cong (gen_cong R)"
proof -
- { fix R' x y
- have "rel (gen_cong R) x y \<Longrightarrow> R \<le> R' \<Longrightarrow> cong R' \<Longrightarrow> R' (eval x) (eval y)"
- by (force simp: rel_fun_def gen_cong_def cong_def dest: spec[of _ R']
+ have "rel (gen_cong R) x y \<Longrightarrow> R \<le> R' \<Longrightarrow> cong R' \<Longrightarrow> R' (eval x) (eval y)" for R' x y
+ by (force simp: rel_fun_def gen_cong_def cong_def dest: spec[of _ R']
predicate2D[OF rel_mono, rotated -1, of _ _ _ R'])
- }
then show "cong (gen_cong R)" by (auto simp: equivp_gen_cong rel_fun_def gen_cong_def cong_def)
qed
--- a/src/HOL/Library/Discrete.thy Sun Nov 03 21:12:50 2024 +0100
+++ b/src/HOL/Library/Discrete.thy Sun Nov 03 22:29:07 2024 +0100
@@ -37,7 +37,7 @@
with \<open>n \<ge> 2\<close> show ?thesis by (rule double)
qed
qed
-
+
lemma log_zero [simp]: "log 0 = 0"
by (simp add: log.simps)
@@ -177,11 +177,8 @@
fixes n :: nat
shows "finite {m. m\<^sup>2 \<le> n}" and "{m. m\<^sup>2 \<le> n} \<noteq> {}"
proof -
- { fix m
- assume "m\<^sup>2 \<le> n"
- then have "m \<le> n"
- by (cases m) (simp_all add: power2_eq_square)
- } note ** = this
+ have **: "m \<le> n" if "m\<^sup>2 \<le> n" for m
+ using that by (cases m) (simp_all add: power2_eq_square)
then have "{m. m\<^sup>2 \<le> n} \<subseteq> {m. m \<le> n}" by auto
then show "finite {m. m\<^sup>2 \<le> n}" by (rule finite_subset) rule
have "0\<^sup>2 \<le> n" by simp
@@ -295,14 +292,14 @@
using sqrt_aux [of n] by (auto simp add: sqrt_def intro: power2_nat_le_imp_le)
text \<open>Additional facts about the discrete square root, thanks to Julian Biendarra, Manuel Eberl\<close>
-
+
lemma Suc_sqrt_power2_gt: "n < (Suc (Discrete.sqrt n))^2"
using Max_ge[OF Discrete.sqrt_aux(1), of "Discrete.sqrt n + 1" n]
by (cases "n < (Suc (Discrete.sqrt n))^2") (simp_all add: Discrete.sqrt_def)
lemma le_sqrt_iff: "x \<le> Discrete.sqrt y \<longleftrightarrow> x^2 \<le> y"
proof -
- have "x \<le> Discrete.sqrt y \<longleftrightarrow> (\<exists>z. z\<^sup>2 \<le> y \<and> x \<le> z)"
+ have "x \<le> Discrete.sqrt y \<longleftrightarrow> (\<exists>z. z\<^sup>2 \<le> y \<and> x \<le> z)"
using Max_ge_iff[OF Discrete.sqrt_aux, of x y] by (simp add: Discrete.sqrt_def)
also have "\<dots> \<longleftrightarrow> x^2 \<le> y"
proof safe
@@ -311,7 +308,7 @@
qed auto
finally show ?thesis .
qed
-
+
lemma le_sqrtI: "x^2 \<le> y \<Longrightarrow> x \<le> Discrete.sqrt y"
by (simp add: le_sqrt_iff)
@@ -321,14 +318,14 @@
lemma sqrt_leI:
"(\<And>z. z^2 \<le> y \<Longrightarrow> z \<le> x) \<Longrightarrow> Discrete.sqrt y \<le> x"
by (simp add: sqrt_le_iff)
-
+
lemma sqrt_Suc:
"Discrete.sqrt (Suc n) = (if \<exists>m. Suc n = m^2 then Suc (Discrete.sqrt n) else Discrete.sqrt n)"
proof cases
assume "\<exists> m. Suc n = m^2"
then obtain m where m_def: "Suc n = m^2" by blast
then have lhs: "Discrete.sqrt (Suc n) = m" by simp
- from m_def sqrt_power2_le[of n]
+ from m_def sqrt_power2_le[of n]
have "(Discrete.sqrt n)^2 < m^2" by linarith
with power2_less_imp_less have lt_m: "Discrete.sqrt n < m" by blast
from m_def Suc_sqrt_power2_gt[of "n"]
@@ -340,7 +337,7 @@
next
assume asm: "\<not> (\<exists> m. Suc n = m^2)"
hence "Suc n \<noteq> (Discrete.sqrt (Suc n))^2" by simp
- with sqrt_power2_le[of "Suc n"]
+ with sqrt_power2_le[of "Suc n"]
have "Discrete.sqrt (Suc n) \<le> Discrete.sqrt n" by (intro le_sqrtI) linarith
moreover have "Discrete.sqrt (Suc n) \<ge> Discrete.sqrt n"
by (intro monoD[OF mono_sqrt]) simp_all
--- a/src/HOL/Library/Extended_Nat.thy Sun Nov 03 21:12:50 2024 +0100
+++ b/src/HOL/Library/Extended_Nat.thy Sun Nov 03 22:29:07 2024 +0100
@@ -656,20 +656,21 @@
definition Sup_enat :: "enat set \<Rightarrow> enat" where
"Sup_enat A = (if A = {} then 0 else if finite A then Max A else \<infinity>)"
+
instance
proof
fix x :: "enat" and A :: "enat set"
- { assume "x \<in> A" then show "Inf A \<le> x"
- unfolding Inf_enat_def by (auto intro: Least_le) }
- { assume "\<And>y. y \<in> A \<Longrightarrow> x \<le> y" then show "x \<le> Inf A"
- unfolding Inf_enat_def
- by (cases "A = {}") (auto intro: LeastI2_ex) }
- { assume "x \<in> A" then show "x \<le> Sup A"
- unfolding Sup_enat_def by (cases "finite A") auto }
- { assume "\<And>y. y \<in> A \<Longrightarrow> y \<le> x" then show "Sup A \<le> x"
- unfolding Sup_enat_def using finite_enat_bounded by auto }
-qed (simp_all add:
- inf_enat_def sup_enat_def bot_enat_def top_enat_def Inf_enat_def Sup_enat_def)
+ show "x \<in> A \<Longrightarrow> Inf A \<le> x"
+ unfolding Inf_enat_def by (auto intro: Least_le)
+ show "(\<And>y. y \<in> A \<Longrightarrow> x \<le> y) \<Longrightarrow> x \<le> Inf A"
+ unfolding Inf_enat_def
+ by (cases "A = {}") (auto intro: LeastI2_ex)
+ show "x \<in> A \<Longrightarrow> x \<le> Sup A"
+ unfolding Sup_enat_def by (cases "finite A") auto
+ show "(\<And>y. y \<in> A \<Longrightarrow> y \<le> x) \<Longrightarrow> Sup A \<le> x"
+ unfolding Sup_enat_def using finite_enat_bounded by auto
+qed (simp_all add: inf_enat_def sup_enat_def bot_enat_def top_enat_def Inf_enat_def Sup_enat_def)
+
end
instance enat :: complete_linorder ..
--- a/src/HOL/Library/Extended_Nonnegative_Real.thy Sun Nov 03 21:12:50 2024 +0100
+++ b/src/HOL/Library/Extended_Nonnegative_Real.thy Sun Nov 03 22:29:07 2024 +0100
@@ -37,14 +37,15 @@
proof -
have "(\<lambda>k. (\<Sum>n<k. f (n + i)) + (\<Sum>j<i. f j)) \<longlonglongrightarrow> l + (\<Sum>j<i. f j)"
using assms by (auto intro!: tendsto_add simp: sums_def)
- moreover
- { fix k :: nat
+ moreover have "(\<Sum>j<k + i. f j) = (\<Sum>n<k. f (n + i)) + (\<Sum>j<i. f j)" for k :: nat
+ proof -
have "(\<Sum>j<k + i. f j) = (\<Sum>j=i..<k + i. f j) + (\<Sum>j=0..<i. f j)"
by (subst sum.union_disjoint[symmetric]) (auto intro!: sum.cong)
also have "(\<Sum>j=i..<k + i. f j) = (\<Sum>j\<in>(\<lambda>n. n + i)`{0..<k}. f j)"
unfolding image_add_atLeastLessThan by simp
- finally have "(\<Sum>j<k + i. f j) = (\<Sum>n<k. f (n + i)) + (\<Sum>j<i. f j)"
- by (auto simp: inj_on_def atLeast0LessThan sum.reindex) }
+ finally show ?thesis
+ by (auto simp: inj_on_def atLeast0LessThan sum.reindex)
+ qed
ultimately have "(\<lambda>k. (\<Sum>n<k + i. f n)) \<longlonglongrightarrow> l + (\<Sum>j<i. f j)"
by simp
then show ?thesis
@@ -222,7 +223,7 @@
instance
by standard
- (transfer ; auto simp: Inf_lower Inf_greatest Sup_upper Sup_least le_max_iff_disj max.absorb1)+
+ (transfer; auto simp: Inf_lower Inf_greatest Sup_upper Sup_least le_max_iff_disj max.absorb1)+
end
@@ -236,8 +237,7 @@
begin
definition infinity_ennreal :: ennreal
-where
- [simp]: "\<infinity> = (top::ennreal)"
+ where [simp]: "\<infinity> = (top::ennreal)"
instance ..
@@ -293,7 +293,7 @@
instance ennreal :: ordered_comm_semiring
by standard
- (transfer ; auto intro: add_mono mult_mono mult_ac ereal_left_distrib ereal_mult_left_mono)+
+ (transfer; auto intro: add_mono mult_mono mult_ac ereal_left_distrib ereal_mult_left_mono)+
instance ennreal :: linordered_nonzero_semiring
proof
@@ -788,7 +788,7 @@
declare [[coercion ennreal]]
-lemma ennreal_cong: "x = y \<Longrightarrow> ennreal x = ennreal y"
+lemma ennreal_cong: "x = y \<Longrightarrow> ennreal x = ennreal y"
by simp
lemma ennreal_cases[cases type: ennreal]:
@@ -975,7 +975,7 @@
case (real r) then show ?thesis
proof (cases "x = 0")
case False then show ?thesis
- by (smt (verit, best) ennreal_0 ennreal_power inverse_ennreal
+ by (smt (verit, best) ennreal_0 ennreal_power inverse_ennreal
inverse_nonnegative_iff_nonnegative power_inverse real zero_less_power)
qed (simp add: top_power_ennreal)
qed
@@ -1288,7 +1288,7 @@
assumes *: "\<forall>\<^sub>F x in F. 0 \<le> f x" and x: "0 \<le> x"
shows "(f \<longlongrightarrow> x) F"
proof -
- have "((\<lambda>x. enn2ereal (ennreal (f x))) \<longlongrightarrow> enn2ereal (ennreal x)) F
+ have "((\<lambda>x. enn2ereal (ennreal (f x))) \<longlongrightarrow> enn2ereal (ennreal x)) F
\<longleftrightarrow> (f \<longlongrightarrow> enn2ereal (ennreal x)) F"
using "*" eventually_mono
by (intro tendsto_cong) fastforce
@@ -1322,7 +1322,7 @@
lemma ennreal_tendsto_0_iff: "(\<And>n. f n \<ge> 0) \<Longrightarrow> ((\<lambda>n. ennreal (f n)) \<longlonglongrightarrow> 0) \<longleftrightarrow> (f \<longlonglongrightarrow> 0)"
by (metis (mono_tags) ennreal_0 eventuallyI order_refl tendsto_ennreal_iff)
-
+
lemma continuous_on_add_ennreal:
fixes f g :: "'a::topological_space \<Rightarrow> ennreal"
shows "continuous_on A f \<Longrightarrow> continuous_on A g \<Longrightarrow> continuous_on A (\<lambda>x. f x + g x)"
--- a/src/HOL/Library/Extended_Real.thy Sun Nov 03 21:12:50 2024 +0100
+++ b/src/HOL/Library/Extended_Real.thy Sun Nov 03 22:29:07 2024 +0100
@@ -1163,7 +1163,7 @@
case False
then obtain p q where "x = ereal p" "y = ereal q"
by (metis MInfty_eq_minfinity ereal.distinct(3) uminus_ereal.elims)
- then show ?thesis
+ then show ?thesis
by (metis assms field_le_epsilon ereal_less(2) ereal_less_eq(3) plus_ereal.simps(1))
qed
@@ -1857,7 +1857,7 @@
by (metis MInfty_eq_minfinity ereal.distinct(3) uminus_ereal.elims)
moreover have "\<And>x. \<lbrakk>q < x; x < p\<rbrakk> \<Longrightarrow> x \<in> real_of_ereal ` {ereal q<..<ereal p}"
by (metis greaterThanLessThan_iff imageI less_ereal.simps(1) real_of_ereal.simps(1))
- ultimately show ?thesis
+ ultimately show ?thesis
by (auto elim!: less_ereal.elims)
qed
@@ -2026,17 +2026,15 @@
lemma tendsto_cmult_ereal[tendsto_intros, simp, intro]:
assumes c: "\<bar>c\<bar> \<noteq> \<infinity>" and f: "(f \<longlongrightarrow> x) F" shows "((\<lambda>x. c * f x::ereal) \<longlongrightarrow> c * x) F"
proof -
- { fix c :: ereal assume "0 < c" "c < \<infinity>"
- then have "((\<lambda>x. c * f x::ereal) \<longlongrightarrow> c * x) F"
- apply (intro tendsto_compose[OF _ f])
- apply (auto intro!: order_tendstoI simp: eventually_at_topological)
- apply (rule_tac x="{a/c <..}" in exI)
- apply (auto split: ereal.split simp: ereal_divide_less_iff mult.commute) []
- apply (rule_tac x="{..< a/c}" in exI)
- apply (auto split: ereal.split simp: ereal_less_divide_iff mult.commute) []
- done }
- note * = this
-
+ have *: "((\<lambda>x. c * f x::ereal) \<longlongrightarrow> c * x) F" if "0 < c" "c < \<infinity>" for c :: ereal
+ using that
+ apply (intro tendsto_compose[OF _ f])
+ apply (auto intro!: order_tendstoI simp: eventually_at_topological)
+ apply (rule_tac x="{a/c <..}" in exI)
+ apply (auto split: ereal.split simp: ereal_divide_less_iff mult.commute) []
+ apply (rule_tac x="{..< a/c}" in exI)
+ apply (auto split: ereal.split simp: ereal_less_divide_iff mult.commute) []
+ done
have "((0 < c \<and> c < \<infinity>) \<or> (-\<infinity> < c \<and> c < 0) \<or> c = 0)"
using c by (cases c) auto
then show ?thesis
@@ -2267,9 +2265,8 @@
assumes "I \<noteq> {}" and "\<bar>c\<bar> \<noteq> \<infinity>"
shows "(INF i\<in>I. c - f i) = c - (SUP i\<in>I. f i::ereal)"
proof -
- { fix b have "(-c) + b = - (c - b)"
- using \<open>\<bar>c\<bar> \<noteq> \<infinity>\<close> by (cases c b rule: ereal2_cases) auto }
- note * = this
+ have *: "(- c) + b = - (c - b)" for b
+ using \<open>\<bar>c\<bar> \<noteq> \<infinity>\<close> by (cases c b rule: ereal2_cases) auto
show ?thesis
using SUP_ereal_add_right[OF \<open>I \<noteq> {}\<close>, of "-c" f] \<open>\<bar>c\<bar> \<noteq> \<infinity>\<close>
by (auto simp add: * ereal_SUP_uminus_eq)
@@ -2359,10 +2356,8 @@
proof -
have INF_less: "(INF i. f i) < \<infinity>" "(INF i. g i) < \<infinity>"
using assms unfolding INF_less_iff by auto
- { fix a b :: ereal assume "a \<noteq> \<infinity>" "b \<noteq> \<infinity>"
- then have "- ((- a) + (- b)) = a + b"
- by (cases a b rule: ereal2_cases) auto }
- note * = this
+ have *: "- ((- a) + (- b)) = a + b" if "a \<noteq> \<infinity>" "b \<noteq> \<infinity>" for a b :: ereal
+ using that by (cases a b rule: ereal2_cases) auto
have "(INF i. f i + g i) = (INF i. - ((- f i) + (- g i)))"
by (simp add: fin *)
also have "\<dots> = Inf (f ` UNIV) + Inf (g ` UNIV)"
@@ -2953,8 +2948,8 @@
fixes x y :: "ereal"
assumes "x \<noteq> -\<infinity>" "y \<noteq> -\<infinity>" "(f \<longlongrightarrow> x) F" "(g \<longlongrightarrow> y) F"
shows "((\<lambda>x. f x + g x) \<longlongrightarrow> x + y) F"
-proof cases
- assume "x = \<infinity> \<or> y = \<infinity>"
+proof (cases "x = \<infinity> \<or> y = \<infinity>")
+ case True
moreover
{ fix y :: ereal and f g :: "'a \<Rightarrow> ereal" assume "y \<noteq> -\<infinity>" "(f \<longlongrightarrow> \<infinity>) F" "(g \<longlongrightarrow> y) F"
then obtain y' where "-\<infinity> < y'" "y' < y"
@@ -2974,7 +2969,7 @@
ultimately show ?thesis
using assms by (auto simp: add_ac)
next
- assume "\<not> (x = \<infinity> \<or> y = \<infinity>)"
+ case False
with assms tendsto_add_ereal[of x y f F g]
show ?thesis
by auto
@@ -3939,7 +3934,7 @@
fixes f :: "'a::t2_space \<Rightarrow> ereal"
assumes "\<And>x. x \<in> A \<Longrightarrow> \<bar>f x\<bar> \<noteq> \<infinity>"
shows "continuous_on A f \<longleftrightarrow> continuous_on A (real_of_ereal \<circ> f)"
-proof
+proof
assume L: "continuous_on A f"
have "f ` A \<subseteq> UNIV - {\<infinity>, -\<infinity>}"
using assms by force
@@ -3950,7 +3945,7 @@
then have "continuous_on A (ereal \<circ> (real_of_ereal \<circ> f))"
by (meson continuous_at_iff_ereal continuous_on_eq_continuous_within)
then show "continuous_on A f"
- using assms ereal_real' by auto
+ using assms ereal_real' by auto
qed
lemma continuous_uminus_ereal [continuous_intros]: "continuous_on (A :: ereal set) uminus"
--- a/src/HOL/Library/Landau_Symbols.thy Sun Nov 03 21:12:50 2024 +0100
+++ b/src/HOL/Library/Landau_Symbols.thy Sun Nov 03 22:29:07 2024 +0100
@@ -659,7 +659,7 @@
using assms by (intro big_power_increasing[OF small_imp_big]) auto
finally show ?thesis by simp
qed
-
+
sublocale big: landau_symbol L L' Lr
proof
have L: "L = bigo \<or> L = bigomega"
@@ -1522,15 +1522,16 @@
assumes "f \<in> o[F](h)" "g \<in> o[F](h)"
shows "(\<lambda>k. max (f k) (g k)) \<in> o[F](h)" "(\<lambda>k. min (f k) (g k)) \<in> o[F](h)"
proof -
- { fix c::real
- assume "c>0"
- with assms smallo_def
+ have "\<forall>\<^sub>F x in F. norm (max (f x) (g x)) \<le> c * norm(h x) \<and> norm (min (f x) (g x)) \<le> c * norm(h x)"
+ if "c>0" for c::real
+ proof -
+ from assms smallo_def that
have "\<forall>\<^sub>F x in F. norm (f x) \<le> c * norm(h x)" "\<forall>\<^sub>F x in F. norm(g x) \<le> c * norm(h x)"
by (auto simp: smallo_def)
- then have "\<forall>\<^sub>F x in F. norm (max (f x) (g x)) \<le> c * norm(h x) \<and> norm (min (f x) (g x)) \<le> c * norm(h x)"
+ then show ?thesis
by (smt (verit) eventually_elim2 max_def min_def)
- } with assms
- show "(\<lambda>x. max (f x) (g x)) \<in> o[F](h)" "(\<lambda>x. min (f x) (g x)) \<in> o[F](h)"
+ qed
+ with assms show "(\<lambda>x. max (f x) (g x)) \<in> o[F](h)" "(\<lambda>x. min (f x) (g x)) \<in> o[F](h)"
by (smt (verit) eventually_elim2 landau_o.smallI)+
qed
--- a/src/HOL/Library/Liminf_Limsup.thy Sun Nov 03 21:12:50 2024 +0100
+++ b/src/HOL/Library/Liminf_Limsup.thy Sun Nov 03 22:29:07 2024 +0100
@@ -431,16 +431,14 @@
assumes c: "continuous_on UNIV f" and am: "mono f" and F: "F \<noteq> bot"
shows "Liminf F (\<lambda>n. f (g n)) = f (Liminf F g)"
proof -
- { fix P assume "eventually P F"
- have "\<exists>x. P x"
- proof (rule ccontr)
- assume "\<not> (\<exists>x. P x)" then have "P = (\<lambda>x. False)"
- by auto
- with \<open>eventually P F\<close> F show False
- by auto
- qed }
- note * = this
-
+ have *: "\<exists>x. P x" if "eventually P F" for P
+ proof (rule ccontr)
+ assume "\<not> ?thesis"
+ then have "P = (\<lambda>x. False)"
+ by auto
+ with \<open>eventually P F\<close> F show False
+ by auto
+ qed
have "f (SUP P\<in>{P. eventually P F}. Inf (g ` Collect P)) =
Sup (f ` (\<lambda>P. Inf (g ` Collect P)) ` {P. eventually P F})"
using am continuous_on_imp_continuous_within [OF c]
@@ -449,7 +447,7 @@
by (simp add: Liminf_def image_comp)
also have "\<dots> = (SUP P \<in> {P. eventually P F}. Inf (f ` (g ` Collect P)))"
using * continuous_at_Inf_mono [OF am continuous_on_imp_continuous_within [OF c]]
- by auto
+ by auto
finally show ?thesis by (auto simp: Liminf_def image_comp)
qed
@@ -458,16 +456,14 @@
assumes c: "continuous_on UNIV f" and am: "mono f" and F: "F \<noteq> bot"
shows "Limsup F (\<lambda>n. f (g n)) = f (Limsup F g)"
proof -
- { fix P assume "eventually P F"
- have "\<exists>x. P x"
- proof (rule ccontr)
- assume "\<not> (\<exists>x. P x)" then have "P = (\<lambda>x. False)"
- by auto
- with \<open>eventually P F\<close> F show False
- by auto
- qed }
- note * = this
-
+ have *: "\<exists>x. P x" if "eventually P F" for P
+ proof (rule ccontr)
+ assume "\<not> ?thesis"
+ then have "P = (\<lambda>x. False)"
+ by auto
+ with \<open>eventually P F\<close> F show False
+ by auto
+ qed
have "f (INF P\<in>{P. eventually P F}. Sup (g ` Collect P)) =
Inf (f ` (\<lambda>P. Sup (g ` Collect P)) ` {P. eventually P F})"
using am continuous_on_imp_continuous_within [OF c]
@@ -513,16 +509,13 @@
assumes c: "continuous_on UNIV f" and am: "antimono f" and F: "F \<noteq> bot"
shows "Limsup F (\<lambda>n. f (g n)) = f (Liminf F g)"
proof -
- { fix P assume "eventually P F"
- have "\<exists>x. P x"
- proof (rule ccontr)
- assume "\<not> (\<exists>x. P x)" then have "P = (\<lambda>x. False)"
- by auto
- with \<open>eventually P F\<close> F show False
- by auto
- qed }
- note * = this
-
+ have *: "\<exists>x. P x" if "eventually P F" for P
+ proof (rule ccontr)
+ assume "\<not> (\<exists>x. P x)" then have "P = (\<lambda>x. False)"
+ by auto
+ with \<open>eventually P F\<close> F show False
+ by auto
+ qed
have "f (SUP P\<in>{P. eventually P F}. Inf (g ` Collect P)) =
Inf (f ` (\<lambda>P. Inf (g ` Collect P)) ` {P. eventually P F})"
using am continuous_on_imp_continuous_within [OF c]
--- a/src/HOL/Library/Multiset.thy Sun Nov 03 21:12:50 2024 +0100
+++ b/src/HOL/Library/Multiset.thy Sun Nov 03 22:29:07 2024 +0100
@@ -636,7 +636,7 @@
lemma mset_subset_eq_insertD:
assumes "add_mset x A \<subseteq># B"
shows "x \<in># B \<and> A \<subset># B"
-proof
+proof
show "x \<in># B"
using assms by (simp add: mset_subset_eqD)
have "A \<subseteq># add_mset x A"
@@ -728,7 +728,7 @@
by (simp add: union_mset_def)
global_interpretation subset_mset: semilattice_neutr_order \<open>(\<union>#)\<close> \<open>{#}\<close> \<open>(\<supseteq>#)\<close> \<open>(\<supset>#)\<close>
-proof
+proof
show "\<And>a b. (b \<subseteq># a) = (a = a \<union># b)"
by (simp add: Diff_eq_empty_iff_mset union_mset_def)
show "\<And>a b. (b \<subset># a) = (a = a \<union># b \<and> a \<noteq> b)"
@@ -2224,7 +2224,7 @@
lemma count_mset_set': "count (mset_set A) x = (if finite A \<and> x \<in> A then 1 else 0)"
by auto
-lemma subset_imp_msubset_mset_set:
+lemma subset_imp_msubset_mset_set:
assumes "A \<subseteq> B" "finite B"
shows "mset_set A \<subseteq># mset_set B"
proof (rule mset_subset_eqI)
@@ -2958,7 +2958,7 @@
show ?case
proof (cases "x \<in> set_mset M")
case True
- have "(\<Prod>y\<in>set_mset (add_mset x M). f y ^ count (add_mset x M) y) =
+ have "(\<Prod>y\<in>set_mset (add_mset x M). f y ^ count (add_mset x M) y) =
(\<Prod>y\<in>set_mset M. (if y = x then f x else 1) * f y ^ count M y)"
using True add by (intro prod.cong) auto
also have "\<dots> = f x * (\<Prod>y\<in>set_mset M. f y ^ count M y)"
@@ -2970,7 +2970,7 @@
hence "(\<Prod>y\<in>set_mset (add_mset x M). f y ^ count (add_mset x M) y) =
f x * (\<Prod>y\<in>set_mset M. f y ^ count (add_mset x M) y)"
by (auto simp: not_in_iff)
- also have "(\<Prod>y\<in>set_mset M. f y ^ count (add_mset x M) y) =
+ also have "(\<Prod>y\<in>set_mset M. f y ^ count (add_mset x M) y) =
(\<Prod>y\<in>set_mset M. f y ^ count M y)"
using False by (intro prod.cong) auto
also note add.IH [symmetric]
@@ -3942,9 +3942,8 @@
using one_step_implies_mult[of "M - M \<inter># N" "N - M \<inter># N" R "M \<inter># N"] *
by (auto simp: multp_code_def Let_def)
next
- { fix I J K :: "'a multiset" assume "(I + J) \<inter># (I + K) = {#}"
- then have "I = {#}" by (metis inter_union_distrib_right union_eq_empty)
- } note [dest!] = this
+ have [dest!]: "I = {#}" if "(I + J) \<inter># (I + K) = {#}" for I J K
+ using that by (metis inter_union_distrib_right union_eq_empty)
assume ?R thus ?L
using mult_cancel_max
using mult_implies_one_step[OF assms(2), of "N - M \<inter># N" "M - M \<inter># N"]
@@ -3975,11 +3974,14 @@
assumes "irrefl_on (set_mset N \<inter> set_mset M) R" and "trans R" and "\<And>x y. P x y \<longleftrightarrow> (x, y) \<in> R"
shows "multeqp_code P N M \<longleftrightarrow> (N, M) \<in> (mult R)\<^sup>="
proof -
- { assume "N \<noteq> M" "M - M \<inter># N = {#}"
- then obtain y where "count N y \<noteq> count M y" by (auto simp flip: count_inject)
- then have "\<exists>y. count M y < count N y" using \<open>M - M \<inter># N = {#}\<close>
+ have "\<exists>y. count M y < count N y" if "N \<noteq> M" "M - M \<inter># N = {#}"
+ proof -
+ from that obtain y where "count N y \<noteq> count M y"
+ by (auto simp flip: count_inject)
+ then show ?thesis
+ using \<open>M - M \<inter># N = {#}\<close>
by (auto simp flip: count_inject dest!: le_neq_implies_less fun_cong[of _ _ y])
- }
+ qed
then have "multeqp_code P N M \<longleftrightarrow> multp_code P N M \<or> N = M"
by (auto simp: multeqp_code_def multp_code_def Let_def in_diff_count)
thus ?thesis
@@ -4538,7 +4540,7 @@
lemma ex_mset: "\<exists>xs. mset xs = X"
by (induct X) (simp, metis mset.simps(2))
-inductive pred_mset :: "('a \<Rightarrow> bool) \<Rightarrow> 'a multiset \<Rightarrow> bool"
+inductive pred_mset :: "('a \<Rightarrow> bool) \<Rightarrow> 'a multiset \<Rightarrow> bool"
where
"pred_mset P {#}"
| "\<lbrakk>P a; pred_mset P M\<rbrakk> \<Longrightarrow> pred_mset P (add_mset a M)"
--- a/src/HOL/Library/Order_Continuity.thy Sun Nov 03 21:12:50 2024 +0100
+++ b/src/HOL/Library/Order_Continuity.thy Sun Nov 03 22:29:07 2024 +0100
@@ -133,12 +133,12 @@
proof (rule lfp_lowerbound)
have "mono (\<lambda>i::nat. (F ^^ i) bot)"
proof -
- { fix i::nat have "(F ^^ i) bot \<le> (F ^^ (Suc i)) bot"
- proof (induct i)
- case 0 show ?case by simp
- next
- case Suc thus ?case using monoD[OF mono Suc] by auto
- qed }
+ have "(F ^^ i) bot \<le> (F ^^ (Suc i)) bot" for i::nat
+ proof (induct i)
+ case 0 show ?case by simp
+ next
+ case Suc thus ?case using monoD[OF mono Suc] by auto
+ qed
thus ?thesis by (auto simp add: mono_iff_le_Suc)
qed
hence "F ?U = (SUP i. (F ^^ Suc i) bot)"
@@ -307,12 +307,14 @@
proof (rule gfp_upperbound)
have *: "antimono (\<lambda>i::nat. (F ^^ i) top)"
proof -
- { fix i::nat have "(F ^^ Suc i) top \<le> (F ^^ i) top"
- proof (induct i)
- case 0 show ?case by simp
- next
- case Suc thus ?case using monoD[OF mono Suc] by auto
- qed }
+ have "(F ^^ Suc i) top \<le> (F ^^ i) top" for i::nat
+ proof (induct i)
+ case 0
+ show ?case by simp
+ next
+ case Suc
+ thus ?case using monoD[OF mono Suc] by auto
+ qed
thus ?thesis by (auto simp add: antimono_iff_le_Suc)
qed
have "?U \<le> (INF i. (F ^^ Suc i) top)"
--- a/src/HOL/Library/Poly_Mapping.thy Sun Nov 03 21:12:50 2024 +0100
+++ b/src/HOL/Library/Poly_Mapping.thy Sun Nov 03 22:29:07 2024 +0100
@@ -471,11 +471,12 @@
proof -
have *: "finite {k. (\<exists>l. f1 l \<noteq> 0 \<and> (\<exists>q. f2 q \<noteq> 0 \<and> k = l + q))}"
using assms by simp
- { fix k l
+ have aux: "sum f2 {q. f2 q \<noteq> 0 \<and> k = l + q} = (\<Sum>q. (f2 q when k = l + q))" for k l
+ proof -
have "{q. (f2 q when k = l + q) \<noteq> 0} \<subseteq> {q. f2 q \<noteq> 0 \<and> k = l + q}" by auto
- with fin2 have "sum f2 {q. f2 q \<noteq> 0 \<and> k = l + q} = (\<Sum>q. (f2 q when k = l + q))"
- by (simp add: Sum_any.expand_superset [of "{q. f2 q \<noteq> 0 \<and> k = l + q}"]) }
- note aux = this
+ with fin2 show ?thesis
+ by (simp add: Sum_any.expand_superset [of "{q. f2 q \<noteq> 0 \<and> k = l + q}"])
+ qed
have "{k. (\<Sum>l. f1 l * sum f2 {q. f2 q \<noteq> 0 \<and> k = l + q}) \<noteq> 0}
\<subseteq> {k. (\<exists>l. f1 l * sum f2 {q. f2 q \<noteq> 0 \<and> k = l + q} \<noteq> 0)}"
by (auto elim!: Sum_any.not_neutral_obtains_not_neutral)
@@ -526,7 +527,7 @@
have "?lhs k = (\<Sum>(ab, c). (\<Sum>(a, b). f a * g b when ab = a + b) * h c when k = ab + c)"
by (simp add: prod_fun_unfold_prod)
also have "\<dots> = (\<Sum>(ab, c). (\<Sum>(a, b). f a * g b * h c when k = ab + c when ab = a + b))"
- using fin_fg
+ using fin_fg
apply (simp add: Sum_any_left_distrib split_def flip: Sum_any_when_independent)
apply (simp add: when_when when_mult mult_when conj_commute)
done
@@ -635,7 +636,7 @@
assume fin_g: "finite {k. g k \<noteq> 0}"
assume fin_h: "finite {k. h k \<noteq> 0}"
show "prod_fun (\<lambda>k. f k + g k) h = (\<lambda>k. prod_fun f h k + prod_fun g h k)"
- by (auto simp: prod_fun_def fun_eq_iff algebra_simps
+ by (auto simp: prod_fun_def fun_eq_iff algebra_simps
Sum_any.distrib fin_f fin_g finite_mult_not_eq_zero_rightI)
qed
qed
@@ -921,14 +922,15 @@
proof (intro_classes, transfer)
fix f g h :: "'a \<Rightarrow> 'b"
assume *: "less_fun f g \<or> f = g"
- { assume "less_fun f g"
- then obtain k where "f k < g k" "(\<And>k'. k' < k \<Longrightarrow> f k' = g k')"
+ have "less_fun (\<lambda>k. h k + f k) (\<lambda>k. h k + g k)" if "less_fun f g"
+ proof -
+ from that obtain k where "f k < g k" "(\<And>k'. k' < k \<Longrightarrow> f k' = g k')"
by (blast elim!: less_funE)
then have "h k + f k < h k + g k" "(\<And>k'. k' < k \<Longrightarrow> h k' + f k' = h k' + g k')"
by simp_all
- then have "less_fun (\<lambda>k. h k + f k) (\<lambda>k. h k + g k)"
+ then show ?thesis
by (blast intro: less_funI)
- }
+ qed
with * show "less_fun (\<lambda>k. h k + f k) (\<lambda>k. h k + g k) \<or> (\<lambda>k. h k + f k) = (\<lambda>k. h k + g k)"
by (auto simp: fun_eq_iff)
qed
@@ -1273,12 +1275,14 @@
proof(transfer fixing: s)
fix p :: "'a \<Rightarrow> 'b"
assume *: "finite {x. p x \<noteq> 0}"
- { fix x
- have "prod_fun (\<lambda>k'. s when 0 = k') p x =
- (\<Sum>l :: 'a. if l = 0 then s * (\<Sum>q. p q when x = q) else 0)"
- by(auto simp: prod_fun_def when_def intro: Sum_any.cong simp del: Sum_any.delta)
- also have "\<dots> = (\<lambda>k. s * p k when p k \<noteq> 0) x" by(simp add: when_def)
- also note calculation }
+ have "prod_fun (\<lambda>k'. s when 0 = k') p x = (\<lambda>k. s * p k when p k \<noteq> 0) x" (is "?lhs = ?rhs") for x
+ proof -
+ have "?lhs = (\<Sum>l :: 'a. if l = 0 then s * (\<Sum>q. p q when x = q) else 0)"
+ by (auto simp: prod_fun_def when_def intro: Sum_any.cong simp del: Sum_any.delta)
+ also have "\<dots> = ?rhs"
+ by (simp add: when_def)
+ finally show ?thesis .
+ qed
then show "(\<lambda>k. s * p k when p k \<noteq> 0) = prod_fun (\<lambda>k'. s when 0 = k') p"
by(simp add: fun_eq_iff)
qed
@@ -1345,17 +1349,18 @@
"keys (nth xs) = fst ` {(n, v) \<in> set (enumerate 0 xs). v \<noteq> 0}"
proof transfer
fix xs :: "'a list"
- { fix n
- assume "nth_default 0 xs n \<noteq> 0"
- then have "n < length xs" and "xs ! n \<noteq> 0"
+ have "n \<in> fst ` {(n, v). (n, v) \<in> set (enumerate 0 xs) \<and> v \<noteq> 0}"
+ if "nth_default 0 xs n \<noteq> 0" for n
+ proof -
+ from that have "n < length xs" and "xs ! n \<noteq> 0"
by (auto simp: nth_default_def split: if_splits)
then have "(n, xs ! n) \<in> {(n, v). (n, v) \<in> set (enumerate 0 xs) \<and> v \<noteq> 0}" (is "?x \<in> ?A")
by (auto simp: in_set_conv_nth enumerate_eq_zip)
then have "fst ?x \<in> fst ` ?A"
by blast
- then have "n \<in> fst ` {(n, v). (n, v) \<in> set (enumerate 0 xs) \<and> v \<noteq> 0}"
+ then show ?thesis
by simp
- }
+ qed
then show "{k. nth_default 0 xs k \<noteq> 0} = fst ` {(n, v). (n, v) \<in> set (enumerate 0 xs) \<and> v \<noteq> 0}"
by (auto simp: in_enumerate_iff_nth_default_eq)
qed
@@ -1535,7 +1540,7 @@
case False
then show ?thesis
by (simp add: Poly_Mapping.poly_mapping_size_def in_keys_iff)
-qed
+qed
lemma poly_mapping_size_estimation:
"k \<in> keys m \<Longrightarrow> y \<le> f k + g (lookup m k) \<Longrightarrow> y < poly_mapping_size m"
@@ -1654,7 +1659,7 @@
lemma keys_cmul: "Poly_Mapping.keys(frag_cmul c a) \<subseteq> Poly_Mapping.keys a"
using finite_cmul_nonzero [of c a]
by (metis lookup_frag_cmul mult_zero_right not_in_keys_iff_lookup_eq_zero subsetI)
-
+
lemma keys_cmul_iff [iff]: "i \<in> Poly_Mapping.keys (frag_cmul c x) \<longleftrightarrow> i \<in> Poly_Mapping.keys x \<and> c \<noteq> 0"
by (metis in_keys_iff lookup_frag_cmul mult_eq_0_iff)
@@ -1662,7 +1667,7 @@
lemma keys_minus [simp]: "Poly_Mapping.keys(-a) = Poly_Mapping.keys a"
by (metis (no_types, opaque_lifting) in_keys_iff lookup_uminus neg_equal_0_iff_equal subsetI subset_antisym)
-lemma keys_diff:
+lemma keys_diff:
"Poly_Mapping.keys(a - b) \<subseteq> Poly_Mapping.keys a \<union> Poly_Mapping.keys b"
by (auto simp: in_keys_iff lookup_minus)
@@ -1726,15 +1731,15 @@
lemma frag_extend_add:
"frag_extend f (a+b) = (frag_extend f a) + (frag_extend f b)"
proof -
- have *: "(\<Sum>i\<in>Poly_Mapping.keys a. frag_cmul (poly_mapping.lookup a i) (f i))
+ have *: "(\<Sum>i\<in>Poly_Mapping.keys a. frag_cmul (poly_mapping.lookup a i) (f i))
= (\<Sum>i\<in>Poly_Mapping.keys a \<union> Poly_Mapping.keys b. frag_cmul (poly_mapping.lookup a i) (f i))"
- "(\<Sum>i\<in>Poly_Mapping.keys b. frag_cmul (poly_mapping.lookup b i) (f i))
+ "(\<Sum>i\<in>Poly_Mapping.keys b. frag_cmul (poly_mapping.lookup b i) (f i))
= (\<Sum>i\<in>Poly_Mapping.keys a \<union> Poly_Mapping.keys b. frag_cmul (poly_mapping.lookup b i) (f i))"
by (auto simp: in_keys_iff intro: sum.mono_neutral_cong_left)
have "frag_extend f (a+b) = (\<Sum>i\<in>Poly_Mapping.keys (a + b).
frag_cmul (poly_mapping.lookup a i) (f i) + frag_cmul (poly_mapping.lookup b i) (f i)) "
by (auto simp: frag_extend_def Poly_Mapping.lookup_add frag_cmul_distrib)
- also have "... = (\<Sum>i \<in> Poly_Mapping.keys a \<union> Poly_Mapping.keys b. frag_cmul (poly_mapping.lookup a i) (f i)
+ also have "... = (\<Sum>i \<in> Poly_Mapping.keys a \<union> Poly_Mapping.keys b. frag_cmul (poly_mapping.lookup a i) (f i)
+ frag_cmul (poly_mapping.lookup b i) (f i))"
proof (rule sum.mono_neutral_cong_left)
show "\<forall>i\<in>keys a \<union> keys b - keys (a + b).
@@ -1830,7 +1835,7 @@
fixes c :: "'a \<Rightarrow>\<^sub>0 int"
assumes "Poly_Mapping.keys c \<subseteq> S \<union> T"
obtains d e where "Poly_Mapping.keys d \<subseteq> S" "Poly_Mapping.keys e \<subseteq> T" "d + e = c"
-proof
+proof
let ?d = "frag_extend (\<lambda>f. if f \<in> S then frag_of f else 0) c"
let ?e = "frag_extend (\<lambda>f. if f \<in> S then 0 else frag_of f) c"
show "Poly_Mapping.keys ?d \<subseteq> S" "Poly_Mapping.keys ?e \<subseteq> T"
--- a/src/HOL/Library/Ramsey.thy Sun Nov 03 21:12:50 2024 +0100
+++ b/src/HOL/Library/Ramsey.thy Sun Nov 03 22:29:07 2024 +0100
@@ -5,13 +5,13 @@
section \<open>Ramsey's Theorem\<close>
theory Ramsey
- imports Infinite_Set Equipollence FuncSet
+ imports Infinite_Set Equipollence FuncSet
begin
subsection \<open>Preliminary definitions\<close>
abbreviation strict_sorted :: "'a::linorder list \<Rightarrow> bool" where
- "strict_sorted \<equiv> sorted_wrt (<)"
+ "strict_sorted \<equiv> sorted_wrt (<)"
subsubsection \<open>The $n$-element subsets of a set $A$\<close>
@@ -89,12 +89,13 @@
shows "[A]\<^bsup>4\<^esup> = {U. \<exists>u x y z. U = {u,x,y,z} \<and> u \<in> A \<and> x \<in> A \<and> y \<in> A \<and> z \<in> A \<and> u < x \<and> x < y \<and> y < z}"
(is "_ = Collect ?RHS")
proof -
- { fix U
- assume "U \<in> [A]\<^bsup>4\<^esup>"
- then obtain l where "strict_sorted l" "List.set l = U" "length l = 4" "U \<subseteq> A"
+ have "?RHS U" if "U \<in> [A]\<^bsup>4\<^esup>" for U
+ proof -
+ from that obtain l where "strict_sorted l" "List.set l = U" "length l = 4" "U \<subseteq> A"
by (simp add: nsets_def) (metis finite_set_strict_sorted)
- then have "?RHS U"
- unfolding numeral_nat length_Suc_conv by auto blast }
+ then show ?thesis
+ unfolding numeral_nat length_Suc_conv by auto blast
+ qed
moreover
have "Collect ?RHS \<subseteq> [A]\<^bsup>4\<^esup>"
apply (clarsimp simp add: nsets_def eval_nat_numeral)
@@ -109,13 +110,14 @@
shows "[A]\<^bsup>5\<^esup> = {U. \<exists>u v x y z. U = {u,v,x,y,z} \<and> u \<in> A \<and> v \<in> A \<and> x \<in> A \<and> y \<in> A \<and> z \<in> A \<and> u < v \<and> v < x \<and> x < y \<and> y < z}"
(is "_ = Collect ?RHS")
proof -
- { fix U
- assume "U \<in> [A]\<^bsup>5\<^esup>"
- then obtain l where "strict_sorted l" "List.set l = U" "length l = 5" "U \<subseteq> A"
- apply (simp add: nsets_def)
- by (metis finite_set_strict_sorted)
- then have "?RHS U"
- unfolding numeral_nat length_Suc_conv by auto blast }
+ have "?RHS U" if "U \<in> [A]\<^bsup>5\<^esup>" for U
+ proof -
+ from that obtain l where "strict_sorted l" "List.set l = U" "length l = 5" "U \<subseteq> A"
+ apply (simp add: nsets_def)
+ by (metis finite_set_strict_sorted)
+ then show ?thesis
+ unfolding numeral_nat length_Suc_conv by auto blast
+ qed
moreover
have "Collect ?RHS \<subseteq> [A]\<^bsup>5\<^esup>"
apply (clarsimp simp add: nsets_def eval_nat_numeral)
@@ -338,7 +340,7 @@
using partn_lst_def monochromatic_def assms by metis
lemma partn_lst_less:
- assumes M: "partn_lst \<beta> \<alpha> n" and eq: "length \<alpha>' = length \<alpha>"
+ assumes M: "partn_lst \<beta> \<alpha> n" and eq: "length \<alpha>' = length \<alpha>"
and le: "\<And>i. i < length \<alpha> \<Longrightarrow> \<alpha>'!i \<le> \<alpha>!i "
shows "partn_lst \<beta> \<alpha>' n"
proof (clarsimp simp: partn_lst_def)
@@ -385,7 +387,7 @@
text \<open>The Erdős--Szekeres bound, essentially extracted from the proof\<close>
fun ES :: "[nat,nat,nat] \<Rightarrow> nat"
where "ES 0 k l = max k l"
- | "ES (Suc r) k l =
+ | "ES (Suc r) k l =
(if r=0 then k+l-1
else if k=0 \<or> l=0 then 1 else Suc (ES r (ES (Suc r) (k-1) l) (ES (Suc r) k (l-1))))"
@@ -685,7 +687,7 @@
using Suc.prems
proof (induct k \<equiv> "q1 + q2" arbitrary: q1 q2)
case 0
- with partn_lst_0 show ?case by auto
+ with partn_lst_0 show ?case by auto
next
case (Suc k)
consider "q1 = 0 \<or> q2 = 0" | "q1 \<noteq> 0" "q2 \<noteq> 0" by auto
--- a/src/HOL/Library/Stream.thy Sun Nov 03 21:12:50 2024 +0100
+++ b/src/HOL/Library/Stream.thy Sun Nov 03 22:29:07 2024 +0100
@@ -470,11 +470,12 @@
next
case False
hence "x = flat (stl s) !! (y - length (shd s))" by (metis less(2,3) flat_snth)
- moreover
- { from less(2) have *: "length (shd s) > 0" by (cases s) simp_all
+ moreover have "y - length (shd s) < y"
+ proof -
+ from less(2) have *: "length (shd s) > 0" by (cases s) simp_all
with False have "y > 0" by (cases y) simp_all
- with * have "y - length (shd s) < y" by simp
- }
+ with * show ?thesis by simp
+ qed
moreover have "\<forall>xs \<in> sset (stl s). xs \<noteq> []" using less(2) by (cases s) auto
ultimately have "\<exists>n m'. x = stl s !! n ! m' \<and> m' < length (stl s !! n)" by (intro less(1)) auto
thus ?thesis by (metis snth.simps(2))
@@ -482,7 +483,9 @@
qed
thus "x \<in> ?R" by (auto simp: sset_range dest!: nth_mem)
next
- fix x xs assume "xs \<in> sset s" ?P "x \<in> set xs" thus "x \<in> ?L"
+ fix x xs
+ assume "xs \<in> sset s" ?P "x \<in> set xs"
+ thus "x \<in> ?L"
by (induct rule: sset_induct)
(metis UnI1 flat_unfold shift.simps(1) sset_shift,
metis UnI2 flat_unfold shd_sset stl_sset sset_shift)
--- a/src/HOL/Library/Sublist.thy Sun Nov 03 21:12:50 2024 +0100
+++ b/src/HOL/Library/Sublist.thy Sun Nov 03 22:29:07 2024 +0100
@@ -80,7 +80,7 @@
next
assume "xs = ys @ [y] \<or> prefix xs ys"
then show "prefix xs (ys @ [y])"
- by auto (metis append.assoc prefix_def)
+ by auto (metis append.assoc prefix_def)
qed
lemma Cons_prefix_Cons [simp]: "prefix (x # xs) (y # ys) = (x = y \<and> prefix xs ys)"
@@ -96,7 +96,7 @@
by (induct xs) simp_all
lemma same_prefix_nil [simp]: "prefix (xs @ ys) xs = (ys = [])"
- by (simp add: prefix_def)
+ by (simp add: prefix_def)
lemma prefix_prefix [simp]: "prefix xs ys \<Longrightarrow> prefix xs (ys @ zs)"
unfolding prefix_def by fastforce
@@ -295,7 +295,7 @@
lemma length_prefixes[simp]: "length (prefixes xs) = length xs+1"
by (induction xs) auto
-
+
lemma distinct_prefixes [intro]: "distinct (prefixes xs)"
by (induction xs) (auto simp: distinct_map)
@@ -310,8 +310,8 @@
lemma last_prefixes [simp]: "last (prefixes xs) = xs"
by (induction xs) (simp_all add: last_map)
-
-lemma prefixes_append:
+
+lemma prefixes_append:
"prefixes (xs @ ys) = prefixes xs @ map (\<lambda>ys'. xs @ ys') (tl (prefixes ys))"
proof (induction xs)
case Nil
@@ -323,7 +323,7 @@
(ys = [] \<and> xs = [] \<or> (\<exists>z zs. ys = zs@[z] \<and> xs = prefixes zs)) \<and> x = ys"
by (cases ys rule: rev_cases) auto
-lemma prefixes_tailrec [code]:
+lemma prefixes_tailrec [code]:
"prefixes xs = rev (snd (foldl (\<lambda>(acc1, acc2) x. (x#acc1, rev (x#acc1)#acc2)) ([],[[]]) xs))"
proof -
have "foldl (\<lambda>(acc1, acc2) x. (x#acc1, rev (x#acc1)#acc2)) (ys, rev ys # zs) xs =
@@ -335,14 +335,14 @@
qed simp_all
from this [of "[]" "[]"] show ?thesis by simp
qed
-
+
lemma set_prefixes_eq: "set (prefixes xs) = {ys. prefix ys xs}"
by auto
lemma card_set_prefixes [simp]: "card (set (prefixes xs)) = Suc (length xs)"
by (subst distinct_card) auto
-lemma set_prefixes_append:
+lemma set_prefixes_append:
"set (prefixes (xs @ ys)) = set (prefixes xs) \<union> {xs @ ys' |ys'. ys' \<in> set (prefixes ys)}"
by (subst prefixes_append, cases ys) auto
@@ -376,13 +376,14 @@
by - (rule Least_equality, fastforce+)
have 2: "?L \<noteq> {}" using \<open>x # xs \<in> L\<close> by auto
from Suc.hyps(1)[OF 1[symmetric] 2] obtain ps where IH: "?P ?L ps" ..
- { fix qs
- assume "\<forall>qs. (\<forall>xa. x # xa \<in> L \<longrightarrow> prefix qs xa) \<longrightarrow> length qs \<le> length ps"
- and "\<forall>xs\<in>L. prefix qs xs"
- hence "length (tl qs) \<le> length ps"
- by (metis Cons_prefix_Cons hd_Cons_tl list.sel(2) Nil_prefix)
- hence "length qs \<le> Suc (length ps)" by auto
- }
+ have "length qs \<le> Suc (length ps)"
+ if "\<forall>qs. (\<forall>xa. x # xa \<in> L \<longrightarrow> prefix qs xa) \<longrightarrow> length qs \<le> length ps"
+ and "\<forall>xs\<in>L. prefix qs xs" for qs
+ proof -
+ from that have "length (tl qs) \<le> length ps"
+ by (metis Cons_prefix_Cons hd_Cons_tl list.sel(2) Nil_prefix)
+ thus ?thesis by auto
+ qed
hence "?P L (x#ps)" using True IH by auto
thus ?thesis ..
next
@@ -544,7 +545,7 @@
assumes "suffix xs ys"
obtains zs where "ys = zs @ xs"
using assms unfolding suffix_def by blast
-
+
lemma suffix_tl [simp]: "suffix (tl xs) xs"
by (induct xs) (auto simp: suffix_def)
@@ -599,7 +600,7 @@
then have "ys = rev zs @ xs" by simp
then show "suffix xs ys" ..
qed
-
+
lemma strict_suffix_to_prefix [code]: "strict_suffix xs ys \<longleftrightarrow> strict_prefix (rev xs) (rev ys)"
by (auto simp: suffix_to_prefix strict_suffix_def strict_prefix_def)
@@ -645,7 +646,7 @@
theorem suffix_Cons: "suffix xs (y # ys) \<longleftrightarrow> xs = y # ys \<or> suffix xs ys"
unfolding suffix_def by (auto simp: Cons_eq_append_conv)
-theorem suffix_append:
+theorem suffix_append:
"suffix xs (ys @ zs) \<longleftrightarrow> suffix xs zs \<or> (\<exists>xs'. xs = xs' @ zs \<and> suffix xs' ys)"
by (auto simp: suffix_def append_eq_append_conv2)
@@ -672,7 +673,7 @@
then show ?case by (cases ys) simp_all
next
case (Suc n)
- then show ?case
+ then show ?case
by (cases xs) (auto intro: Suc dest: suffix_ConsD' suffix_order.less_imp_le)
qed
@@ -808,7 +809,7 @@
lemma last_suffixes [simp]: "last (suffixes xs) = xs"
by (cases xs) simp_all
-lemma suffixes_append:
+lemma suffixes_append:
"suffixes (xs @ ys) = suffixes ys @ map (\<lambda>xs'. xs' @ ys) (tl (suffixes xs))"
proof (induction ys rule: rev_induct)
case Nil
@@ -824,7 +825,7 @@
(ys = [] \<and> xs = [] \<or> (\<exists>z zs. ys = z#zs \<and> xs = suffixes zs)) \<and> x = ys"
by (cases ys) auto
-lemma suffixes_tailrec [code]:
+lemma suffixes_tailrec [code]:
"suffixes xs = rev (snd (foldl (\<lambda>(acc1, acc2) x. (x#acc1, (x#acc1)#acc2)) ([],[[]]) (rev xs)))"
proof -
have "foldl (\<lambda>(acc1, acc2) x. (x#acc1, (x#acc1)#acc2)) (ys, ys # zs) (rev xs) =
@@ -836,14 +837,14 @@
qed simp_all
from this [of "[]" "[]"] show ?thesis by simp
qed
-
+
lemma set_suffixes_eq: "set (suffixes xs) = {ys. suffix ys xs}"
by auto
-
+
lemma card_set_suffixes [simp]: "card (set (suffixes xs)) = Suc (length xs)"
by (subst distinct_card) auto
-
-lemma set_suffixes_append:
+
+lemma set_suffixes_append:
"set (suffixes (xs @ ys)) = set (suffixes ys) \<union> {xs' @ ys |xs'. xs' \<in> set (suffixes xs)}"
by (subst suffixes_append, cases xs rule: rev_cases) auto
@@ -853,10 +854,10 @@
lemma prefixes_conv_suffixes: "prefixes xs = map rev (suffixes (rev xs))"
by (induction xs) auto
-
+
lemma prefixes_rev: "prefixes (rev xs) = map rev (suffixes xs)"
by (induction xs) auto
-
+
lemma suffixes_rev: "suffixes (rev xs) = map rev (prefixes xs)"
by (induction xs) auto
@@ -870,13 +871,13 @@
| list_emb_Cons [intro] : "list_emb P xs ys \<Longrightarrow> list_emb P xs (y#ys)"
| list_emb_Cons2 [intro]: "P x y \<Longrightarrow> list_emb P xs ys \<Longrightarrow> list_emb P (x#xs) (y#ys)"
-lemma list_emb_mono:
+lemma list_emb_mono:
assumes "\<And>x y. P x y \<longrightarrow> Q x y"
shows "list_emb P xs ys \<longrightarrow> list_emb Q xs ys"
-proof
- assume "list_emb P xs ys"
+proof
+ assume "list_emb P xs ys"
then show "list_emb Q xs ys" by (induct) (auto simp: assms)
-qed
+qed
lemma list_emb_Nil2 [simp]:
assumes "list_emb P xs []" shows "xs = []"
@@ -888,13 +889,11 @@
using assms by (induct xs) auto
lemma list_emb_Cons_Nil [simp]: "list_emb P (x#xs) [] = False"
-proof -
- { assume "list_emb P (x#xs) []"
- from list_emb_Nil2 [OF this] have False by simp
- } moreover {
- assume False
- then have "list_emb P (x#xs) []" by simp
- } ultimately show ?thesis by blast
+proof
+ show False if "list_emb P (x#xs) []"
+ using list_emb_Nil2 [OF that] by simp
+ show "list_emb P (x#xs) []" if False
+ using that ..
qed
lemma list_emb_append2 [intro]: "list_emb P xs ys \<Longrightarrow> list_emb P xs (zs @ ys)"
@@ -1002,13 +1001,13 @@
"list_emb P (x#xs) [] \<longleftrightarrow> False"
"list_emb P (x#xs) (y#ys) \<longleftrightarrow> (if P x y then list_emb P xs ys else list_emb P (x#xs) ys)"
by simp_all
-
+
subsection \<open>Subsequences (special case of homeomorphic embedding)\<close>
abbreviation subseq :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool"
where "subseq xs ys \<equiv> list_emb (=) xs ys"
-
+
definition strict_subseq where "strict_subseq xs ys \<longleftrightarrow> xs \<noteq> ys \<and> subseq xs ys"
lemma subseq_Cons2: "subseq xs ys \<Longrightarrow> subseq (x#xs) (x#ys)" by auto
@@ -1073,7 +1072,7 @@
thus "subseq xs ys"
by (induction ys arbitrary: xs) (auto simp: Let_def)
next
- have [simp]: "[] \<in> set (subseqs ys)" for ys :: "'a list"
+ have [simp]: "[] \<in> set (subseqs ys)" for ys :: "'a list"
by (induction ys) (auto simp: Let_def)
assume "subseq xs ys"
thus "xs \<in> set (subseqs ys)"
@@ -1105,33 +1104,39 @@
lemma subseq_append [simp]:
"subseq (xs @ zs) (ys @ zs) \<longleftrightarrow> subseq xs ys" (is "?l = ?r")
proof
- { fix xs' ys' xs ys zs :: "'a list" assume "subseq xs' ys'"
- then have "xs' = xs @ zs \<and> ys' = ys @ zs \<longrightarrow> subseq xs ys"
- proof (induct arbitrary: xs ys zs)
- case list_emb_Nil show ?case by simp
- next
- case (list_emb_Cons xs' ys' x)
- { assume "ys=[]" then have ?case using list_emb_Cons(1) by auto }
- moreover
- { fix us assume "ys = x#us"
- then have ?case using list_emb_Cons(2) by(simp add: list_emb.list_emb_Cons) }
- ultimately show ?case by (auto simp:Cons_eq_append_conv)
- next
- case (list_emb_Cons2 x y xs' ys')
- { assume "xs=[]" then have ?case using list_emb_Cons2(1) by auto }
- moreover
- { fix us vs assume "xs=x#us" "ys=x#vs" then have ?case using list_emb_Cons2 by auto}
- moreover
- { fix us assume "xs=x#us" "ys=[]" then have ?case using list_emb_Cons2(2) by bestsimp }
- ultimately show ?case using \<open>(=) x y\<close> by (auto simp: Cons_eq_append_conv)
- qed }
- moreover assume ?l
- ultimately show ?r by blast
-next
- assume ?r then show ?l by (metis list_emb_append_mono subseq_order.order_refl)
+ have "xs' = xs @ zs \<and> ys' = ys @ zs \<longrightarrow> subseq xs ys"
+ if "subseq xs' ys'" for xs' ys' xs ys zs :: "'a list"
+ using that
+ proof (induct arbitrary: xs ys zs)
+ case list_emb_Nil
+ show ?case by simp
+ next
+ case (list_emb_Cons xs' ys' x)
+ have ?case if "ys = []"
+ using list_emb_Cons(1) that by auto
+ moreover
+ have ?case if "ys = x#us" for us
+ using list_emb_Cons(2) that by (simp add: list_emb.list_emb_Cons)
+ ultimately show ?case
+ by (auto simp: Cons_eq_append_conv)
+ next
+ case (list_emb_Cons2 x y xs' ys')
+ have ?case if "xs = []"
+ using list_emb_Cons2(1) that by auto
+ moreover
+ have ?case if "xs = x#us" "ys = x#vs" for us vs
+ using list_emb_Cons2 that by auto
+ moreover
+ have ?case if "xs = x#us" "ys = []" for us
+ using list_emb_Cons2(2) that by bestsimp
+ ultimately show ?case
+ using \<open>x = y\<close> by (auto simp: Cons_eq_append_conv)
+ qed
+ then show "?l \<Longrightarrow> ?r" by blast
+ show "?r \<Longrightarrow> ?l" by (metis list_emb_append_mono subseq_order.order_refl)
qed
-lemma subseq_append_iff:
+lemma subseq_append_iff:
"subseq xs (ys @ zs) \<longleftrightarrow> (\<exists>xs1 xs2. xs = xs1 @ xs2 \<and> subseq xs1 ys \<and> subseq xs2 zs)"
(is "?lhs = ?rhs")
proof
@@ -1139,7 +1144,7 @@
proof (induction xs "ys @ zs" arbitrary: ys zs rule: list_emb.induct)
case (list_emb_Cons xs ws y ys zs)
from list_emb_Cons(2)[of "tl ys" zs] and list_emb_Cons(2)[of "[]" "tl zs"] and list_emb_Cons(1,3)
- show ?case by (cases ys) auto
+ show ?case by (cases ys) auto
next
case (list_emb_Cons2 x y xs ws ys zs)
from list_emb_Cons2(3)[of "tl ys" zs] and list_emb_Cons2(3)[of "[]" "tl zs"]
@@ -1148,7 +1153,7 @@
qed auto
qed (auto intro: list_emb_append_mono)
-lemma subseq_appendE [case_names append]:
+lemma subseq_appendE [case_names append]:
assumes "subseq xs (ys @ zs)"
obtains xs1 xs2 where "xs = xs1 @ xs2" "subseq xs1 ys" "subseq xs2 zs"
using assms by (subst (asm) subseq_append_iff) auto
@@ -1173,13 +1178,13 @@
assumes "subseq xs ys" shows "subseq (filter P xs) (filter P ys)"
using assms by induct auto
-lemma subseq_conv_nths:
- "subseq xs ys \<longleftrightarrow> (\<exists>N. xs = nths ys N)" (is "?L = ?R")
+lemma subseq_conv_nths: "subseq xs ys \<longleftrightarrow> (\<exists>N. xs = nths ys N)"
+ (is "?L = ?R")
proof
- assume ?L
- then show ?R
+ show ?R if ?L using that
proof (induct)
- case list_emb_Nil show ?case by (metis nths_empty)
+ case list_emb_Nil
+ show ?case by (metis nths_empty)
next
case (list_emb_Cons xs ys x)
then obtain N where "xs = nths ys N" by blast
@@ -1194,27 +1199,30 @@
moreover from list_emb_Cons2 have "x = y" by simp
ultimately show ?case by blast
qed
-next
- assume ?R
- then obtain N where "xs = nths ys N" ..
- moreover have "subseq (nths ys N) ys"
- proof (induct ys arbitrary: N)
- case Nil show ?case by simp
- next
- case Cons then show ?case by (auto simp: nths_Cons)
+ show ?L if ?R
+ proof -
+ from that obtain N where "xs = nths ys N" ..
+ moreover have "subseq (nths ys N) ys"
+ proof (induct ys arbitrary: N)
+ case Nil
+ show ?case by simp
+ next
+ case Cons
+ then show ?case by (auto simp: nths_Cons)
+ qed
+ ultimately show ?thesis by simp
qed
- ultimately show ?L by simp
qed
-
-
+
+
subsection \<open>Contiguous sublists\<close>
subsubsection \<open>\<open>sublist\<close>\<close>
-definition sublist :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" where
+definition sublist :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" where
"sublist xs ys = (\<exists>ps ss. ys = ps @ xs @ ss)"
-
-definition strict_sublist :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" where
+
+definition strict_sublist :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" where
"strict_sublist xs ys \<longleftrightarrow> sublist xs ys \<and> xs \<noteq> ys"
interpretation sublist_order: order sublist strict_sublist
@@ -1227,38 +1235,37 @@
thus "sublist xs zs" unfolding sublist_def by blast
next
fix xs ys :: "'a list"
- {
- assume "sublist xs ys" "sublist ys xs"
- then obtain as bs cs ds
- where xs: "xs = as @ ys @ bs" and ys: "ys = cs @ xs @ ds"
+ show "xs = ys" if "sublist xs ys" "sublist ys xs"
+ proof -
+ from that obtain as bs cs ds where xs: "xs = as @ ys @ bs" and ys: "ys = cs @ xs @ ds"
by (auto simp: sublist_def)
have "xs = as @ cs @ xs @ ds @ bs" by (subst xs, subst ys) auto
- also have "length \<dots> = length as + length cs + length xs + length bs + length ds"
+ also have "length \<dots> = length as + length cs + length xs + length bs + length ds"
by simp
finally have "as = []" "bs = []" by simp_all
- with xs show "xs = ys" by simp
- }
- thus "strict_sublist xs ys \<longleftrightarrow> (sublist xs ys \<and> \<not>sublist ys xs)"
+ with xs show ?thesis by simp
+ qed
+ thus "strict_sublist xs ys \<longleftrightarrow> (sublist xs ys \<and> \<not> sublist ys xs)"
by (auto simp: strict_sublist_def)
qed (auto simp: strict_sublist_def sublist_def intro: exI[of _ "[]"])
-
+
lemma sublist_Nil_left [simp, intro]: "sublist [] ys"
by (auto simp: sublist_def)
-
+
lemma sublist_Cons_Nil [simp]: "\<not>sublist (x#xs) []"
by (auto simp: sublist_def)
-
+
lemma sublist_Nil_right [simp]: "sublist xs [] \<longleftrightarrow> xs = []"
by (cases xs) auto
-
+
lemma sublist_appendI [simp, intro]: "sublist xs (ps @ xs @ ss)"
by (auto simp: sublist_def)
-
+
lemma sublist_append_leftI [simp, intro]: "sublist xs (ps @ xs)"
by (auto simp: sublist_def intro: exI[of _ "[]"])
-
+
lemma sublist_append_rightI [simp, intro]: "sublist xs (xs @ ss)"
- by (auto simp: sublist_def intro: exI[of _ "[]"])
+ by (auto simp: sublist_def intro: exI[of _ "[]"])
lemma sublist_altdef: "sublist xs ys \<longleftrightarrow> (\<exists>ys'. prefix ys' ys \<and> suffix xs ys')"
proof safe
@@ -1271,7 +1278,7 @@
assume "prefix ys' ys" "suffix xs ys'"
thus "sublist xs ys" by (auto simp: prefix_def suffix_def)
qed
-
+
lemma sublist_altdef': "sublist xs ys \<longleftrightarrow> (\<exists>ys'. suffix ys' ys \<and> prefix xs ys')"
proof safe
assume "sublist xs ys"
@@ -1286,7 +1293,7 @@
lemma sublist_Cons_right: "sublist xs (y # ys) \<longleftrightarrow> prefix xs (y # ys) \<or> sublist xs ys"
by (auto simp: sublist_def prefix_def Cons_eq_append_conv)
-
+
lemma sublist_code [code]:
"sublist [] ys \<longleftrightarrow> True"
"sublist (x # xs) [] \<longleftrightarrow> False"
@@ -1294,7 +1301,7 @@
by (simp_all add: sublist_Cons_right)
lemma sublist_append:
- "sublist xs (ys @ zs) \<longleftrightarrow>
+ "sublist xs (ys @ zs) \<longleftrightarrow>
sublist xs ys \<or> sublist xs zs \<or> (\<exists>xs1 xs2. xs = xs1 @ xs2 \<and> suffix xs1 ys \<and> prefix xs2 zs)"
by (auto simp: sublist_altdef prefix_append suffix_append)
@@ -1315,10 +1322,10 @@
lemma set_mono_sublist: "sublist xs ys \<Longrightarrow> set xs \<subseteq> set ys"
by (auto simp add: sublist_def)
-
+
lemma prefix_imp_sublist [simp, intro]: "prefix xs ys \<Longrightarrow> sublist xs ys"
by (auto simp: sublist_def prefix_def intro: exI[of _ "[]"])
-
+
lemma suffix_imp_sublist [simp, intro]: "suffix xs ys \<Longrightarrow> sublist xs ys"
by (auto simp: sublist_def suffix_def intro: exI[of _ "[]"])
@@ -1333,13 +1340,13 @@
lemma sublist_dropWhile [simp, intro]: "sublist (dropWhile P xs) xs"
by (rule suffix_imp_sublist[OF suffix_dropWhile])
-
+
lemma sublist_tl [simp, intro]: "sublist (tl xs) xs"
by (rule suffix_imp_sublist) (simp_all add: suffix_drop)
-
+
lemma sublist_butlast [simp, intro]: "sublist (butlast xs) xs"
by (rule prefix_imp_sublist) (simp_all add: prefixeq_butlast)
-
+
lemma sublist_rev [simp]: "sublist (rev xs) (rev ys) = sublist xs ys"
proof
assume "sublist (rev xs) (rev ys)"
@@ -1354,15 +1361,15 @@
also have "rev \<dots> = rev bs @ rev xs @ rev as" by simp
finally show "sublist (rev xs) (rev ys)" by simp
qed
-
+
lemma sublist_rev_left: "sublist (rev xs) ys = sublist xs (rev ys)"
by (subst sublist_rev [symmetric]) (simp only: rev_rev_ident)
-
+
lemma sublist_rev_right: "sublist xs (rev ys) = sublist (rev xs) ys"
by (subst sublist_rev [symmetric]) (simp only: rev_rev_ident)
-lemma snoc_sublist_snoc:
- "sublist (xs @ [x]) (ys @ [y]) \<longleftrightarrow>
+lemma snoc_sublist_snoc:
+ "sublist (xs @ [x]) (ys @ [y]) \<longleftrightarrow>
(x = y \<and> suffix xs ys \<or> sublist (xs @ [x]) ys) "
by (subst (1 2) sublist_rev [symmetric])
(simp del: sublist_rev add: sublist_Cons_right suffix_to_prefix)
@@ -1370,8 +1377,8 @@
lemma sublist_snoc:
"sublist xs (ys @ [y]) \<longleftrightarrow> suffix xs (ys @ [y]) \<or> sublist xs ys"
by (subst (1 2) sublist_rev [symmetric])
- (simp del: sublist_rev add: sublist_Cons_right suffix_to_prefix)
-
+ (simp del: sublist_rev add: sublist_Cons_right suffix_to_prefix)
+
lemma sublist_imp_subseq [intro]: "sublist xs ys \<Longrightarrow> subseq xs ys"
by (auto simp: sublist_def)
@@ -1415,7 +1422,7 @@
"sublists [] = [[]]"
| "sublists (x # xs) = sublists xs @ map ((#) x) (prefixes xs)"
-lemma in_set_sublists [simp]: "xs \<in> set (sublists ys) \<longleftrightarrow> sublist xs ys"
+lemma in_set_sublists [simp]: "xs \<in> set (sublists ys) \<longleftrightarrow> sublist xs ys"
by (induction ys arbitrary: xs) (auto simp: sublist_Cons_right prefix_Cons)
lemma set_sublists_eq: "set (sublists xs) = {ys. sublist ys xs}"
@@ -1428,8 +1435,8 @@
subsection \<open>Parametricity\<close>
context includes lifting_syntax
-begin
-
+begin
+
private lemma prefix_primrec:
"prefix = rec_list (\<lambda>xs. True) (\<lambda>x xs xsa ys.
case ys of [] \<Rightarrow> False | y # ys \<Rightarrow> x = y \<and> xsa ys)"
@@ -1446,7 +1453,7 @@
qed
private lemma list_emb_primrec:
- "list_emb = (\<lambda>uu uua uuaa. rec_list (\<lambda>P xs. List.null xs) (\<lambda>y ys ysa P xs. case xs of [] \<Rightarrow> True
+ "list_emb = (\<lambda>uu uua uuaa. rec_list (\<lambda>P xs. List.null xs) (\<lambda>y ys ysa P xs. case xs of [] \<Rightarrow> True
| x # xs \<Rightarrow> if P x y then ysa P xs else ysa P (x # xs)) uuaa uu uua)"
proof (intro ext, goal_cases)
case (1 P xs ys)
@@ -1457,12 +1464,12 @@
lemma prefix_transfer [transfer_rule]:
assumes [transfer_rule]: "bi_unique A"
- shows "(list_all2 A ===> list_all2 A ===> (=)) prefix prefix"
+ shows "(list_all2 A ===> list_all2 A ===> (=)) prefix prefix"
unfolding prefix_primrec by transfer_prover
-
+
lemma suffix_transfer [transfer_rule]:
assumes [transfer_rule]: "bi_unique A"
- shows "(list_all2 A ===> list_all2 A ===> (=)) suffix suffix"
+ shows "(list_all2 A ===> list_all2 A ===> (=)) suffix suffix"
unfolding suffix_to_prefix [abs_def] by transfer_prover
lemma sublist_transfer [transfer_rule]:
@@ -1474,7 +1481,7 @@
assumes [transfer_rule]: "bi_unique A"
shows "(list_all2 A ===> list_all2 A ===> (=)) parallel parallel"
unfolding parallel_def by transfer_prover
-
+
lemma list_emb_transfer [transfer_rule]:
@@ -1483,34 +1490,34 @@
lemma strict_prefix_transfer [transfer_rule]:
assumes [transfer_rule]: "bi_unique A"
- shows "(list_all2 A ===> list_all2 A ===> (=)) strict_prefix strict_prefix"
+ shows "(list_all2 A ===> list_all2 A ===> (=)) strict_prefix strict_prefix"
unfolding strict_prefix_def by transfer_prover
-
+
lemma strict_suffix_transfer [transfer_rule]:
assumes [transfer_rule]: "bi_unique A"
- shows "(list_all2 A ===> list_all2 A ===> (=)) strict_suffix strict_suffix"
+ shows "(list_all2 A ===> list_all2 A ===> (=)) strict_suffix strict_suffix"
unfolding strict_suffix_def by transfer_prover
-
+
lemma strict_subseq_transfer [transfer_rule]:
assumes [transfer_rule]: "bi_unique A"
- shows "(list_all2 A ===> list_all2 A ===> (=)) strict_subseq strict_subseq"
+ shows "(list_all2 A ===> list_all2 A ===> (=)) strict_subseq strict_subseq"
unfolding strict_subseq_def by transfer_prover
-
+
lemma strict_sublist_transfer [transfer_rule]:
assumes [transfer_rule]: "bi_unique A"
- shows "(list_all2 A ===> list_all2 A ===> (=)) strict_sublist strict_sublist"
+ shows "(list_all2 A ===> list_all2 A ===> (=)) strict_sublist strict_sublist"
unfolding strict_sublist_def by transfer_prover
lemma prefixes_transfer [transfer_rule]:
assumes [transfer_rule]: "bi_unique A"
shows "(list_all2 A ===> list_all2 (list_all2 A)) prefixes prefixes"
unfolding prefixes_def by transfer_prover
-
+
lemma suffixes_transfer [transfer_rule]:
assumes [transfer_rule]: "bi_unique A"
shows "(list_all2 A ===> list_all2 (list_all2 A)) suffixes suffixes"
unfolding suffixes_def by transfer_prover
-
+
lemma sublists_transfer [transfer_rule]:
assumes [transfer_rule]: "bi_unique A"
shows "(list_all2 A ===> list_all2 (list_all2 A)) sublists sublists"