--- a/src/HOL/List.thy Sun Oct 27 12:54:58 2024 +0100
+++ b/src/HOL/List.thy Mon Oct 28 18:48:14 2024 +0100
@@ -1230,6 +1230,17 @@
then show ?case by (cases ys) auto
qed
+lemma rev_eq_append_conv: "rev xs = ys @ zs \<longleftrightarrow> xs = rev zs @ rev ys"
+by (metis rev_append rev_rev_ident)
+
+lemma append_eq_rev_conv: "xs @ ys = rev zs \<longleftrightarrow> rev ys @ rev xs = zs"
+using rev_eq_append_conv[of zs xs ys] by auto
+
+lemma rev_eq_Cons_iff[iff]: "(rev xs = y#ys) = (xs = rev ys @ [y])"
+by (simp add: rev_swap)
+
+lemmas Cons_eq_rev_iff = rev_eq_Cons_iff[THEN eq_iff_swap]
+
lemma inj_on_rev[iff]: "inj_on rev A"
by(simp add:inj_on_def)
@@ -1262,9 +1273,6 @@
qed
qed simp
-lemma rev_eq_Cons_iff[iff]: "(rev xs = y#ys) = (xs = rev ys @ [y])"
-by(rule rev_cases[of xs]) auto
-
lemma length_Suc_conv_rev: "(length xs = Suc n) = (\<exists>y ys. xs = ys @ [y] \<and> length ys = n)"
by (induct xs rule: rev_induct) auto
--- a/src/HOL/Transitive_Closure.thy Sun Oct 27 12:54:58 2024 +0100
+++ b/src/HOL/Transitive_Closure.thy Mon Oct 28 18:48:14 2024 +0100
@@ -690,6 +690,9 @@
lemma trancl_unfold_left: "r\<^sup>+ = r O r\<^sup>*"
by (auto dest: tranclD intro: rtrancl_into_trancl2)
+lemma tranclp_unfold_left: "r^++ = r OO r^**"
+by (auto intro!: ext dest: tranclpD intro: rtranclp_into_tranclp2)
+
lemma trancl_insert: "(insert (y, x) r)\<^sup>+ = r\<^sup>+ \<union> {(a, b). (a, y) \<in> r\<^sup>* \<and> (x, b) \<in> r\<^sup>*}"
\<comment> \<open>primitive recursion for \<open>trancl\<close> over finite relations\<close>
proof -
@@ -930,6 +933,8 @@
end
+lemmas relpowp_Suc_right = relpowp.simps(2)
+
lemma relpowp_relpow_eq [pred_set_conv]:
"(\<lambda>x y. (x, y) \<in> R) ^^ n = (\<lambda>x y. (x, y) \<in> R ^^ n)" for R :: "'a rel"
by (induct n) (simp_all add: relcompp_relcomp_eq)
@@ -963,6 +968,10 @@
for P :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
by (fact relpow_1 [to_pred])
+lemma relpowp_Suc_0 [simp]: "P ^^ (Suc 0) = P"
+ for P :: "'a \<Rightarrow> 'a \<Rightarrow> bool"
+ by (auto)
+
lemma relpow_0_I: "(x, x) \<in> R ^^ 0"
by simp
@@ -972,13 +981,13 @@
lemma relpow_Suc_I: "(x, y) \<in> R ^^ n \<Longrightarrow> (y, z) \<in> R \<Longrightarrow> (x, z) \<in> R ^^ Suc n"
by auto
-lemma relpowp_Suc_I: "(P ^^ n) x y \<Longrightarrow> P y z \<Longrightarrow> (P ^^ Suc n) x z"
+lemma relpowp_Suc_I[trans]: "(P ^^ n) x y \<Longrightarrow> P y z \<Longrightarrow> (P ^^ Suc n) x z"
by (fact relpow_Suc_I [to_pred])
lemma relpow_Suc_I2: "(x, y) \<in> R \<Longrightarrow> (y, z) \<in> R ^^ n \<Longrightarrow> (x, z) \<in> R ^^ Suc n"
by (induct n arbitrary: z) (simp, fastforce)
-lemma relpowp_Suc_I2: "P x y \<Longrightarrow> (P ^^ n) y z \<Longrightarrow> (P ^^ Suc n) x z"
+lemma relpowp_Suc_I2[trans]: "P x y \<Longrightarrow> (P ^^ n) y z \<Longrightarrow> (P ^^ Suc n) x z"
by (fact relpow_Suc_I2 [to_pred])
lemma relpow_0_E: "(x, y) \<in> R ^^ 0 \<Longrightarrow> (x = y \<Longrightarrow> P) \<Longrightarrow> P"
@@ -1064,6 +1073,11 @@
qed
qed
+lemma relpowp_mono:
+ fixes x y :: 'a
+ shows "(\<And>x y. R x y \<Longrightarrow> S x y) \<Longrightarrow> (R ^^ n) x y \<Longrightarrow> (S ^^ n) x y"
+by (induction n arbitrary: y) auto
+
lemma relpow_trans[trans]: "(x, y) \<in> R ^^ i \<Longrightarrow> (y, z) \<in> R ^^ j \<Longrightarrow> (x, z) \<in> R ^^ (i + j)"
using relpowp_trans[to_set] .
@@ -1137,6 +1151,9 @@
lemma relpowp_commute: "P OO P ^^ n = P ^^ n OO P"
by (fact relpow_commute [to_pred])
+lemma relpowp_Suc_left: "R ^^ Suc n = R OO (R ^^ n)"
+by (simp add: relpowp_commute)
+
lemma relpow_empty: "0 < n \<Longrightarrow> ({} :: ('a \<times> 'a) set) ^^ n = {}"
by (cases n) auto