--- a/NEWS Mon Jul 04 16:12:47 2022 +0000
+++ b/NEWS Tue Jul 05 17:54:52 2022 +0200
@@ -90,6 +90,11 @@
totalp_on_subset
totalp_on_total_on_eq[pred_set_conv]
+* Theory "HOL.Transitive_Closure":
+ - Added lemmas.
+ total_on_trancl
+ totalp_on_tranclp
+
* Theory "HOL-Library.Multiset":
- Consolidated operation and fact names.
multp ~> multp_code
--- a/src/HOL/Transitive_Closure.thy Mon Jul 04 16:12:47 2022 +0000
+++ b/src/HOL/Transitive_Closure.thy Tue Jul 05 17:54:52 2022 +0200
@@ -312,6 +312,12 @@
subsection \<open>Transitive closure\<close>
+lemma totalp_on_tranclp: "totalp_on A R \<Longrightarrow> totalp_on A (tranclp R)"
+ by (auto intro: totalp_onI dest: totalp_onD)
+
+lemma total_on_trancl: "total_on A r \<Longrightarrow> total_on A (trancl r)"
+ by (rule totalp_on_tranclp[to_set])
+
lemma trancl_mono:
assumes "p \<in> r\<^sup>+" "r \<subseteq> s"
shows "p \<in> s\<^sup>+"