merged
authordesharna
Mon, 10 Jan 2022 21:34:09 +0100
changeset 74977 e4fd3989554d
parent 74974 7733c794cfea (diff)
parent 74976 70aab133dc8d (current diff)
child 74978 1f30aa91f496
child 74980 8dd527e97ddb
merged
--- a/src/HOL/Library/While_Combinator.thy	Mon Jan 10 14:13:23 2022 +0100
+++ b/src/HOL/Library/While_Combinator.thy	Mon Jan 10 21:34:09 2022 +0100
@@ -239,6 +239,15 @@
   apply blast
   done
 
+text \<open>Combine invariant preservation and variant decrease in one goal:\<close>
+theorem while_rule2:
+  "[| P s;
+      !!s. [| P s; b s  |] ==> P (c s) \<and> (c s, s) \<in> r;
+      !!s. [| P s; \<not> b s  |] ==> Q s;
+      wf r |] ==>
+   Q (while b c s)"
+using while_rule[of P] by metis
+
 text\<open>Proving termination:\<close>
 
 theorem wf_while_option_Some: