src/HOL/Complete_Partial_Order.thy
changeset 63612 7195acc2fe93
parent 62093 bd73a2279fcd
child 63810 67b091896158
--- a/src/HOL/Complete_Partial_Order.thy	Fri Aug 05 16:36:03 2016 +0200
+++ b/src/HOL/Complete_Partial_Order.thy	Fri Aug 05 18:14:28 2016 +0200
@@ -1,12 +1,12 @@
-(* Title:    HOL/Complete_Partial_Order.thy
-   Author:   Brian Huffman, Portland State University
-   Author:   Alexander Krauss, TU Muenchen
+(*  Title:      HOL/Complete_Partial_Order.thy
+    Author:     Brian Huffman, Portland State University
+    Author:     Alexander Krauss, TU Muenchen
 *)
 
 section \<open>Chain-complete partial orders and their fixpoints\<close>
 
 theory Complete_Partial_Order
-imports Product_Type
+  imports Product_Type
 begin
 
 subsection \<open>Monotone functions\<close>
@@ -14,131 +14,139 @@
 text \<open>Dictionary-passing version of @{const Orderings.mono}.\<close>
 
 definition monotone :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('b \<Rightarrow> 'b \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> bool"
-where "monotone orda ordb f \<longleftrightarrow> (\<forall>x y. orda x y \<longrightarrow> ordb (f x) (f y))"
+  where "monotone orda ordb f \<longleftrightarrow> (\<forall>x y. orda x y \<longrightarrow> ordb (f x) (f y))"
 
-lemma monotoneI[intro?]: "(\<And>x y. orda x y \<Longrightarrow> ordb (f x) (f y))
- \<Longrightarrow> monotone orda ordb f"
-unfolding monotone_def by iprover
+lemma monotoneI[intro?]: "(\<And>x y. orda x y \<Longrightarrow> ordb (f x) (f y)) \<Longrightarrow> monotone orda ordb f"
+  unfolding monotone_def by iprover
 
 lemma monotoneD[dest?]: "monotone orda ordb f \<Longrightarrow> orda x y \<Longrightarrow> ordb (f x) (f y)"
-unfolding monotone_def by iprover
+  unfolding monotone_def by iprover
 
 
 subsection \<open>Chains\<close>
 
-text \<open>A chain is a totally-ordered set. Chains are parameterized over
+text \<open>
+  A chain is a totally-ordered set. Chains are parameterized over
   the order for maximal flexibility, since type classes are not enough.
 \<close>
 
-definition
-  chain :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> bool"
-where
-  "chain ord S \<longleftrightarrow> (\<forall>x\<in>S. \<forall>y\<in>S. ord x y \<or> ord y x)"
+definition chain :: "('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a set \<Rightarrow> bool"
+  where "chain ord S \<longleftrightarrow> (\<forall>x\<in>S. \<forall>y\<in>S. ord x y \<or> ord y x)"
 
 lemma chainI:
   assumes "\<And>x y. x \<in> S \<Longrightarrow> y \<in> S \<Longrightarrow> ord x y \<or> ord y x"
   shows "chain ord S"
-using assms unfolding chain_def by fast
+  using assms unfolding chain_def by fast
 
 lemma chainD:
   assumes "chain ord S" and "x \<in> S" and "y \<in> S"
   shows "ord x y \<or> ord y x"
-using assms unfolding chain_def by fast
+  using assms unfolding chain_def by fast
 
 lemma chainE:
   assumes "chain ord S" and "x \<in> S" and "y \<in> S"
   obtains "ord x y" | "ord y x"
-using assms unfolding chain_def by fast
+  using assms unfolding chain_def by fast
 
 lemma chain_empty: "chain ord {}"
-by(simp add: chain_def)
+  by (simp add: chain_def)
 
 lemma chain_equality: "chain op = A \<longleftrightarrow> (\<forall>x\<in>A. \<forall>y\<in>A. x = y)"
-by(auto simp add: chain_def)
+  by (auto simp add: chain_def)
+
+lemma chain_subset: "chain ord A \<Longrightarrow> B \<subseteq> A \<Longrightarrow> chain ord B"
+  by (rule chainI) (blast dest: chainD)
 
-lemma chain_subset:
-  "\<lbrakk> chain ord A; B \<subseteq> A \<rbrakk>
-  \<Longrightarrow> chain ord B"
-by(rule chainI)(blast dest: chainD)
+lemma chain_imageI:
+  assumes chain: "chain le_a Y"
+    and mono: "\<And>x y. x \<in> Y \<Longrightarrow> y \<in> Y \<Longrightarrow> le_a x y \<Longrightarrow> le_b (f x) (f y)"
+  shows "chain le_b (f ` Y)"
+  by (blast intro: chainI dest: chainD[OF chain] mono)
 
-lemma chain_imageI: 
-  assumes chain: "chain le_a Y"
-  and mono: "\<And>x y. \<lbrakk> x \<in> Y; y \<in> Y; le_a x y \<rbrakk> \<Longrightarrow> le_b (f x) (f y)"
-  shows "chain le_b (f ` Y)"
-by(blast intro: chainI dest: chainD[OF chain] mono)
 
 subsection \<open>Chain-complete partial orders\<close>
 
 text \<open>
-  A ccpo has a least upper bound for any chain.  In particular, the
-  empty set is a chain, so every ccpo must have a bottom element.
+  A \<open>ccpo\<close> has a least upper bound for any chain.  In particular, the
+  empty set is a chain, so every \<open>ccpo\<close> must have a bottom element.
 \<close>
 
 class ccpo = order + Sup +
-  assumes ccpo_Sup_upper: "\<lbrakk>chain (op \<le>) A; x \<in> A\<rbrakk> \<Longrightarrow> x \<le> Sup A"
-  assumes ccpo_Sup_least: "\<lbrakk>chain (op \<le>) A; \<And>x. x \<in> A \<Longrightarrow> x \<le> z\<rbrakk> \<Longrightarrow> Sup A \<le> z"
+  assumes ccpo_Sup_upper: "chain (op \<le>) A \<Longrightarrow> x \<in> A \<Longrightarrow> x \<le> Sup A"
+  assumes ccpo_Sup_least: "chain (op \<le>) A \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> x \<le> z) \<Longrightarrow> Sup A \<le> z"
 begin
 
 lemma chain_singleton: "Complete_Partial_Order.chain op \<le> {x}"
-by(rule chainI) simp
+  by (rule chainI) simp
 
 lemma ccpo_Sup_singleton [simp]: "\<Squnion>{x} = x"
-by(rule antisym)(auto intro: ccpo_Sup_least ccpo_Sup_upper simp add: chain_singleton)
+  by (rule antisym)(auto intro: ccpo_Sup_least ccpo_Sup_upper simp add: chain_singleton)
+
 
 subsection \<open>Transfinite iteration of a function\<close>
 
-context notes [[inductive_internals]] begin
+context notes [[inductive_internals]]
+begin
 
 inductive_set iterates :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a set"
-for f :: "'a \<Rightarrow> 'a"
-where
-  step: "x \<in> iterates f \<Longrightarrow> f x \<in> iterates f"
-| Sup: "chain (op \<le>) M \<Longrightarrow> \<forall>x\<in>M. x \<in> iterates f \<Longrightarrow> Sup M \<in> iterates f"
+  for f :: "'a \<Rightarrow> 'a"
+  where
+    step: "x \<in> iterates f \<Longrightarrow> f x \<in> iterates f"
+  | Sup: "chain (op \<le>) M \<Longrightarrow> \<forall>x\<in>M. x \<in> iterates f \<Longrightarrow> Sup M \<in> iterates f"
 
 end
 
-lemma iterates_le_f:
-  "x \<in> iterates f \<Longrightarrow> monotone (op \<le>) (op \<le>) f \<Longrightarrow> x \<le> f x"
-by (induct x rule: iterates.induct)
-  (force dest: monotoneD intro!: ccpo_Sup_upper ccpo_Sup_least)+
+lemma iterates_le_f: "x \<in> iterates f \<Longrightarrow> monotone (op \<le>) (op \<le>) f \<Longrightarrow> x \<le> f x"
+  by (induct x rule: iterates.induct)
+    (force dest: monotoneD intro!: ccpo_Sup_upper ccpo_Sup_least)+
 
 lemma chain_iterates:
   assumes f: "monotone (op \<le>) (op \<le>) f"
   shows "chain (op \<le>) (iterates f)" (is "chain _ ?C")
 proof (rule chainI)
-  fix x y assume "x \<in> ?C" "y \<in> ?C"
+  fix x y
+  assume "x \<in> ?C" "y \<in> ?C"
   then show "x \<le> y \<or> y \<le> x"
   proof (induct x arbitrary: y rule: iterates.induct)
-    fix x y assume y: "y \<in> ?C"
-    and IH: "\<And>z. z \<in> ?C \<Longrightarrow> x \<le> z \<or> z \<le> x"
+    fix x y
+    assume y: "y \<in> ?C"
+      and IH: "\<And>z. z \<in> ?C \<Longrightarrow> x \<le> z \<or> z \<le> x"
     from y show "f x \<le> y \<or> y \<le> f x"
     proof (induct y rule: iterates.induct)
-      case (step y) with IH f show ?case by (auto dest: monotoneD)
+      case (step y)
+      with IH f show ?case by (auto dest: monotoneD)
     next
       case (Sup M)
       then have chM: "chain (op \<le>) M"
         and IH': "\<And>z. z \<in> M \<Longrightarrow> f x \<le> z \<or> z \<le> f x" by auto
       show "f x \<le> Sup M \<or> Sup M \<le> f x"
       proof (cases "\<exists>z\<in>M. f x \<le> z")
-        case True then have "f x \<le> Sup M"
+        case True
+        then have "f x \<le> Sup M"
           apply rule
           apply (erule order_trans)
-          by (rule ccpo_Sup_upper[OF chM])
-        thus ?thesis ..
+          apply (rule ccpo_Sup_upper[OF chM])
+          apply assumption
+          done
+        then show ?thesis ..
       next
-        case False with IH'
-        show ?thesis by (auto intro: ccpo_Sup_least[OF chM])
+        case False
+        with IH' show ?thesis
+          by (auto intro: ccpo_Sup_least[OF chM])
       qed
     qed
   next
     case (Sup M y)
     show ?case
     proof (cases "\<exists>x\<in>M. y \<le> x")
-      case True then have "y \<le> Sup M"
+      case True
+      then have "y \<le> Sup M"
         apply rule
         apply (erule order_trans)
-        by (rule ccpo_Sup_upper[OF Sup(1)])
-      thus ?thesis ..
+        apply (rule ccpo_Sup_upper[OF Sup(1)])
+        apply assumption
+        done
+      then show ?thesis ..
     next
       case False with Sup
       show ?thesis by (auto intro: ccpo_Sup_least)
@@ -147,19 +155,19 @@
 qed
 
 lemma bot_in_iterates: "Sup {} \<in> iterates f"
-by(auto intro: iterates.Sup simp add: chain_empty)
+  by (auto intro: iterates.Sup simp add: chain_empty)
+
 
 subsection \<open>Fixpoint combinator\<close>
 
-definition
-  fixp :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a"
-where
-  "fixp f = Sup (iterates f)"
+definition fixp :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a"
+  where "fixp f = Sup (iterates f)"
 
 lemma iterates_fixp:
-  assumes f: "monotone (op \<le>) (op \<le>) f" shows "fixp f \<in> iterates f"
-unfolding fixp_def
-by (simp add: iterates.Sup chain_iterates f)
+  assumes f: "monotone (op \<le>) (op \<le>) f"
+  shows "fixp f \<in> iterates f"
+  unfolding fixp_def
+  by (simp add: iterates.Sup chain_iterates f)
 
 lemma fixp_unfold:
   assumes f: "monotone (op \<le>) (op \<le>) f"
@@ -169,35 +177,45 @@
     by (intro iterates_le_f iterates_fixp f)
   have "f (fixp f) \<le> Sup (iterates f)"
     by (intro ccpo_Sup_upper chain_iterates f iterates.step iterates_fixp)
-  thus "f (fixp f) \<le> fixp f"
-    unfolding fixp_def .
+  then show "f (fixp f) \<le> fixp f"
+    by (simp only: fixp_def)
 qed
 
 lemma fixp_lowerbound:
-  assumes f: "monotone (op \<le>) (op \<le>) f" and z: "f z \<le> z" shows "fixp f \<le> z"
-unfolding fixp_def
+  assumes f: "monotone (op \<le>) (op \<le>) f"
+    and z: "f z \<le> z"
+  shows "fixp f \<le> z"
+  unfolding fixp_def
 proof (rule ccpo_Sup_least[OF chain_iterates[OF f]])
-  fix x assume "x \<in> iterates f"
-  thus "x \<le> z"
+  fix x
+  assume "x \<in> iterates f"
+  then show "x \<le> z"
   proof (induct x rule: iterates.induct)
-    fix x assume "x \<le> z" with f have "f x \<le> f z" by (rule monotoneD)
-    also note z finally show "f x \<le> z" .
-  qed (auto intro: ccpo_Sup_least)
+    case (step x)
+    from f \<open>x \<le> z\<close> have "f x \<le> f z" by (rule monotoneD)
+    also note z
+    finally show "f x \<le> z" .
+  next
+    case (Sup M)
+    then show ?case
+      by (auto intro: ccpo_Sup_least)
+  qed
 qed
 
 end
 
+
 subsection \<open>Fixpoint induction\<close>
 
 setup \<open>Sign.map_naming (Name_Space.mandatory_path "ccpo")\<close>
 
 definition admissible :: "('a set \<Rightarrow> 'a) \<Rightarrow> ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool"
-where "admissible lub ord P = (\<forall>A. chain ord A \<longrightarrow> (A \<noteq> {}) \<longrightarrow> (\<forall>x\<in>A. P x) \<longrightarrow> P (lub A))"
+  where "admissible lub ord P \<longleftrightarrow> (\<forall>A. chain ord A \<longrightarrow> A \<noteq> {} \<longrightarrow> (\<forall>x\<in>A. P x) \<longrightarrow> P (lub A))"
 
 lemma admissibleI:
   assumes "\<And>A. chain ord A \<Longrightarrow> A \<noteq> {} \<Longrightarrow> \<forall>x\<in>A. P x \<Longrightarrow> P (lub A)"
   shows "ccpo.admissible lub ord P"
-using assms unfolding ccpo.admissible_def by fast
+  using assms unfolding ccpo.admissible_def by fast
 
 lemma admissibleD:
   assumes "ccpo.admissible lub ord P"
@@ -205,7 +223,7 @@
   assumes "A \<noteq> {}"
   assumes "\<And>x. x \<in> A \<Longrightarrow> P x"
   shows "P (lub A)"
-using assms by (auto simp: ccpo.admissible_def)
+  using assms by (auto simp: ccpo.admissible_def)
 
 setup \<open>Sign.map_naming Name_Space.parent_path\<close>
 
@@ -215,44 +233,54 @@
   assumes bot: "P (Sup {})"
   assumes step: "\<And>x. P x \<Longrightarrow> P (f x)"
   shows "P (fixp f)"
-unfolding fixp_def using adm chain_iterates[OF mono]
+  unfolding fixp_def
+  using adm chain_iterates[OF mono]
 proof (rule ccpo.admissibleD)
-  show "iterates f \<noteq> {}" using bot_in_iterates by auto
-  fix x assume "x \<in> iterates f"
-  thus "P x"
-    by (induct rule: iterates.induct)
-      (case_tac "M = {}", auto intro: step bot ccpo.admissibleD adm)
+  show "iterates f \<noteq> {}"
+    using bot_in_iterates by auto
+next
+  fix x
+  assume "x \<in> iterates f"
+  then show "P x"
+  proof (induct rule: iterates.induct)
+    case prems: (step x)
+    from this(2) show ?case by (rule step)
+  next
+    case (Sup M)
+    then show ?case by (cases "M = {}") (auto intro: step bot ccpo.admissibleD adm)
+  qed
 qed
 
 lemma admissible_True: "ccpo.admissible lub ord (\<lambda>x. True)"
-unfolding ccpo.admissible_def by simp
+  unfolding ccpo.admissible_def by simp
 
 (*lemma admissible_False: "\<not> ccpo.admissible lub ord (\<lambda>x. False)"
 unfolding ccpo.admissible_def chain_def by simp
 *)
 lemma admissible_const: "ccpo.admissible lub ord (\<lambda>x. t)"
-by(auto intro: ccpo.admissibleI)
+  by (auto intro: ccpo.admissibleI)
 
 lemma admissible_conj:
   assumes "ccpo.admissible lub ord (\<lambda>x. P x)"
   assumes "ccpo.admissible lub ord (\<lambda>x. Q x)"
   shows "ccpo.admissible lub ord (\<lambda>x. P x \<and> Q x)"
-using assms unfolding ccpo.admissible_def by simp
+  using assms unfolding ccpo.admissible_def by simp
 
 lemma admissible_all:
   assumes "\<And>y. ccpo.admissible lub ord (\<lambda>x. P x y)"
   shows "ccpo.admissible lub ord (\<lambda>x. \<forall>y. P x y)"
-using assms unfolding ccpo.admissible_def by fast
+  using assms unfolding ccpo.admissible_def by fast
 
 lemma admissible_ball:
   assumes "\<And>y. y \<in> A \<Longrightarrow> ccpo.admissible lub ord (\<lambda>x. P x y)"
   shows "ccpo.admissible lub ord (\<lambda>x. \<forall>y\<in>A. P x y)"
-using assms unfolding ccpo.admissible_def by fast
+  using assms unfolding ccpo.admissible_def by fast
 
 lemma chain_compr: "chain ord A \<Longrightarrow> chain ord {x \<in> A. P x}"
-unfolding chain_def by fast
+  unfolding chain_def by fast
 
-context ccpo begin
+context ccpo
+begin
 
 lemma admissible_disj_lemma:
   assumes A: "chain (op \<le>)A"
@@ -280,17 +308,24 @@
   assumes Q: "ccpo.admissible Sup (op \<le>) (\<lambda>x. Q x)"
   shows "ccpo.admissible Sup (op \<le>) (\<lambda>x. P x \<or> Q x)"
 proof (rule ccpo.admissibleI)
-  fix A :: "'a set" assume A: "chain (op \<le>) A"
-  assume "A \<noteq> {}"
-    and "\<forall>x\<in>A. P x \<or> Q x"
-  hence "(\<exists>x\<in>A. P x) \<and> (\<forall>x\<in>A. \<exists>y\<in>A. x \<le> y \<and> P y) \<or> (\<exists>x\<in>A. Q x) \<and> (\<forall>x\<in>A. \<exists>y\<in>A. x \<le> y \<and> Q y)"
-    using chainD[OF A] by blast
-  hence "(\<exists>x. x \<in> A \<and> P x) \<and> Sup A = Sup {x \<in> A. P x} \<or> (\<exists>x. x \<in> A \<and> Q x) \<and> Sup A = Sup {x \<in> A. Q x}"
+  fix A :: "'a set"
+  assume A: "chain (op \<le>) A"
+  assume "A \<noteq> {}" and "\<forall>x\<in>A. P x \<or> Q x"
+  then have "(\<exists>x\<in>A. P x) \<and> (\<forall>x\<in>A. \<exists>y\<in>A. x \<le> y \<and> P y) \<or> (\<exists>x\<in>A. Q x) \<and> (\<forall>x\<in>A. \<exists>y\<in>A. x \<le> y \<and> Q y)"
+    using chainD[OF A] by blast  (* slow *)
+  then have "(\<exists>x. x \<in> A \<and> P x) \<and> Sup A = Sup {x \<in> A. P x} \<or> (\<exists>x. x \<in> A \<and> Q x) \<and> Sup A = Sup {x \<in> A. Q x}"
     using admissible_disj_lemma [OF A] by blast
-  thus "P (Sup A) \<or> Q (Sup A)"
-    apply (rule disjE, simp_all)
-    apply (rule disjI1, rule ccpo.admissibleD [OF P chain_compr [OF A]], simp, simp)
-    apply (rule disjI2, rule ccpo.admissibleD [OF Q chain_compr [OF A]], simp, simp)
+  then show "P (Sup A) \<or> Q (Sup A)"
+    apply (rule disjE)
+     apply simp_all
+     apply (rule disjI1)
+     apply (rule ccpo.admissibleD [OF P chain_compr [OF A]])
+      apply simp
+     apply simp
+    apply (rule disjI2)
+    apply (rule ccpo.admissibleD [OF Q chain_compr [OF A]])
+     apply simp
+    apply simp
     done
 qed
 
@@ -300,7 +335,8 @@
   by standard (fast intro: Sup_upper Sup_least)+
 
 lemma lfp_eq_fixp:
-  assumes f: "mono f" shows "lfp f = fixp f"
+  assumes f: "mono f"
+  shows "lfp f = fixp f"
 proof (rule antisym)
   from f have f': "monotone (op \<le>) (op \<le>) f"
     unfolding mono_def monotone_def .