| author | wenzelm | 
| Wed, 11 Oct 2023 12:37:11 +0200 | |
| changeset 78761 | 6b3f13d39612 | 
| parent 75669 | 43f5dfb7fa35 | 
| permissions | -rw-r--r-- | 
| 40107 | 1 | (* Title: HOL/Partial_Function.thy | 
| 2 | Author: Alexander Krauss, TU Muenchen | |
| 3 | *) | |
| 4 | ||
| 60758 | 5 | section \<open>Partial Function Definitions\<close> | 
| 40107 | 6 | |
| 7 | theory Partial_Function | |
| 66364 | 8 | imports Complete_Partial_Order Option | 
| 69913 | 9 | keywords "partial_function" :: thy_defn | 
| 40107 | 10 | begin | 
| 11 | ||
| 57959 | 12 | named_theorems partial_function_mono "monotonicity rules for partial function definitions" | 
| 69605 | 13 | ML_file \<open>Tools/Function/partial_function.ML\<close> | 
| 57959 | 14 | |
| 60062 | 15 | lemma (in ccpo) in_chain_finite: | 
| 67399 | 16 |   assumes "Complete_Partial_Order.chain (\<le>) A" "finite A" "A \<noteq> {}"
 | 
| 60062 | 17 | shows "\<Squnion>A \<in> A" | 
| 18 | using assms(2,1,3) | |
| 19 | proof induction | |
| 20 | case empty thus ?case by simp | |
| 21 | next | |
| 22 | case (insert x A) | |
| 67399 | 23 | note chain = \<open>Complete_Partial_Order.chain (\<le>) (insert x A)\<close> | 
| 60062 | 24 | show ?case | 
| 25 |   proof(cases "A = {}")
 | |
| 26 | case True thus ?thesis by simp | |
| 27 | next | |
| 28 | case False | |
| 67399 | 29 | from chain have chain': "Complete_Partial_Order.chain (\<le>) A" | 
| 60062 | 30 | by(rule chain_subset) blast | 
| 31 | hence "\<Squnion>A \<in> A" using False by(rule insert.IH) | |
| 32 | show ?thesis | |
| 33 | proof(cases "x \<le> \<Squnion>A") | |
| 34 | case True | |
| 69745 | 35 | have "\<Squnion>(insert x A) \<le> \<Squnion>A" using chain | 
| 60062 | 36 | by(rule ccpo_Sup_least)(auto simp add: True intro: ccpo_Sup_upper[OF chain']) | 
| 69745 | 37 | hence "\<Squnion>(insert x A) = \<Squnion>A" | 
| 73411 | 38 | by(rule order.antisym)(blast intro: ccpo_Sup_upper[OF chain] ccpo_Sup_least[OF chain']) | 
| 60758 | 39 | with \<open>\<Squnion>A \<in> A\<close> show ?thesis by simp | 
| 60062 | 40 | next | 
| 41 | case False | |
| 60758 | 42 | with chainD[OF chain, of x "\<Squnion>A"] \<open>\<Squnion>A \<in> A\<close> | 
| 69745 | 43 | have "\<Squnion>(insert x A) = x" | 
| 73411 | 44 | by(auto intro: order.antisym ccpo_Sup_least[OF chain] order_trans[OF ccpo_Sup_upper[OF chain']] ccpo_Sup_upper[OF chain]) | 
| 60062 | 45 | thus ?thesis by simp | 
| 46 | qed | |
| 47 | qed | |
| 48 | qed | |
| 40107 | 49 | |
| 63561 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
62390diff
changeset | 50 | lemma (in ccpo) admissible_chfin: | 
| 67399 | 51 | "(\<forall>S. Complete_Partial_Order.chain (\<le>) S \<longrightarrow> finite S) | 
| 52 | \<Longrightarrow> ccpo.admissible Sup (\<le>) P" | |
| 63561 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
62390diff
changeset | 53 | using in_chain_finite by (blast intro: ccpo.admissibleI) | 
| 
fba08009ff3e
add lemmas contributed by Peter Gammie
 Andreas Lochbihler parents: 
62390diff
changeset | 54 | |
| 60758 | 55 | subsection \<open>Axiomatic setup\<close> | 
| 40107 | 56 | |
| 60758 | 57 | text \<open>This techical locale constains the requirements for function | 
| 58 | definitions with ccpo fixed points.\<close> | |
| 40107 | 59 | |
| 60 | definition "fun_ord ord f g \<longleftrightarrow> (\<forall>x. ord (f x) (g x))" | |
| 61 | definition "fun_lub L A = (\<lambda>x. L {y. \<exists>f\<in>A. y = f x})"
 | |
| 62 | definition "img_ord f ord = (\<lambda>x y. ord (f x) (f y))" | |
| 63 | definition "img_lub f g Lub = (\<lambda>A. g (Lub (f ` A)))" | |
| 64 | ||
| 43081 | 65 | lemma chain_fun: | 
| 66 | assumes A: "chain (fun_ord ord) A" | |
| 67 |   shows "chain ord {y. \<exists>f\<in>A. y = f a}" (is "chain ord ?C")
 | |
| 68 | proof (rule chainI) | |
| 69 | fix x y assume "x \<in> ?C" "y \<in> ?C" | |
| 70 | then obtain f g where fg: "f \<in> A" "g \<in> A" | |
| 71 | and [simp]: "x = f a" "y = g a" by blast | |
| 72 | from chainD[OF A fg] | |
| 73 | show "ord x y \<or> ord y x" unfolding fun_ord_def by auto | |
| 74 | qed | |
| 75 | ||
| 40107 | 76 | lemma call_mono[partial_function_mono]: "monotone (fun_ord ord) ord (\<lambda>f. f t)" | 
| 77 | by (rule monotoneI) (auto simp: fun_ord_def) | |
| 78 | ||
| 40288 | 79 | lemma let_mono[partial_function_mono]: | 
| 80 | "(\<And>x. monotone orda ordb (\<lambda>f. b f x)) | |
| 81 | \<Longrightarrow> monotone orda ordb (\<lambda>f. Let t (b f))" | |
| 82 | by (simp add: Let_def) | |
| 83 | ||
| 40107 | 84 | lemma if_mono[partial_function_mono]: "monotone orda ordb F | 
| 85 | \<Longrightarrow> monotone orda ordb G | |
| 86 | \<Longrightarrow> monotone orda ordb (\<lambda>f. if c then F f else G f)" | |
| 87 | unfolding monotone_def by simp | |
| 88 | ||
| 89 | definition "mk_less R = (\<lambda>x y. R x y \<and> \<not> R y x)" | |
| 90 | ||
| 91 | locale partial_function_definitions = | |
| 92 | fixes leq :: "'a \<Rightarrow> 'a \<Rightarrow> bool" | |
| 93 | fixes lub :: "'a set \<Rightarrow> 'a" | |
| 94 | assumes leq_refl: "leq x x" | |
| 95 | assumes leq_trans: "leq x y \<Longrightarrow> leq y z \<Longrightarrow> leq x z" | |
| 96 | assumes leq_antisym: "leq x y \<Longrightarrow> leq y x \<Longrightarrow> x = y" | |
| 97 | assumes lub_upper: "chain leq A \<Longrightarrow> x \<in> A \<Longrightarrow> leq x (lub A)" | |
| 98 | assumes lub_least: "chain leq A \<Longrightarrow> (\<And>x. x \<in> A \<Longrightarrow> leq x z) \<Longrightarrow> leq (lub A) z" | |
| 99 | ||
| 100 | lemma partial_function_lift: | |
| 101 | assumes "partial_function_definitions ord lb" | |
| 102 | shows "partial_function_definitions (fun_ord ord) (fun_lub lb)" (is "partial_function_definitions ?ordf ?lubf") | |
| 103 | proof - | |
| 104 | interpret partial_function_definitions ord lb by fact | |
| 105 | ||
| 106 | show ?thesis | |
| 107 | proof | |
| 108 | fix x show "?ordf x x" | |
| 109 | unfolding fun_ord_def by (auto simp: leq_refl) | |
| 110 | next | |
| 111 | fix x y z assume "?ordf x y" "?ordf y z" | |
| 112 | thus "?ordf x z" unfolding fun_ord_def | |
| 113 | by (force dest: leq_trans) | |
| 114 | next | |
| 115 | fix x y assume "?ordf x y" "?ordf y x" | |
| 116 | thus "x = y" unfolding fun_ord_def | |
| 43081 | 117 | by (force intro!: dest: leq_antisym) | 
| 40107 | 118 | next | 
| 119 | fix A f assume f: "f \<in> A" and A: "chain ?ordf A" | |
| 120 | thus "?ordf f (?lubf A)" | |
| 121 | unfolding fun_lub_def fun_ord_def | |
| 122 | by (blast intro: lub_upper chain_fun[OF A] f) | |
| 123 | next | |
| 124 |     fix A :: "('b \<Rightarrow> 'a) set" and g :: "'b \<Rightarrow> 'a"
 | |
| 125 | assume A: "chain ?ordf A" and g: "\<And>f. f \<in> A \<Longrightarrow> ?ordf f g" | |
| 126 | show "?ordf (?lubf A) g" unfolding fun_lub_def fun_ord_def | |
| 127 | by (blast intro: lub_least chain_fun[OF A] dest: g[unfolded fun_ord_def]) | |
| 128 | qed | |
| 129 | qed | |
| 130 | ||
| 131 | lemma ccpo: assumes "partial_function_definitions ord lb" | |
| 46041 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
45297diff
changeset | 132 | shows "class.ccpo lb ord (mk_less ord)" | 
| 40107 | 133 | using assms unfolding partial_function_definitions_def mk_less_def | 
| 134 | by unfold_locales blast+ | |
| 135 | ||
| 136 | lemma partial_function_image: | |
| 137 | assumes "partial_function_definitions ord Lub" | |
| 138 | assumes inj: "\<And>x y. f x = f y \<Longrightarrow> x = y" | |
| 139 | assumes inv: "\<And>x. f (g x) = x" | |
| 140 | shows "partial_function_definitions (img_ord f ord) (img_lub f g Lub)" | |
| 141 | proof - | |
| 142 | let ?iord = "img_ord f ord" | |
| 143 | let ?ilub = "img_lub f g Lub" | |
| 144 | ||
| 145 | interpret partial_function_definitions ord Lub by fact | |
| 146 | show ?thesis | |
| 147 | proof | |
| 148 | fix A x assume "chain ?iord A" "x \<in> A" | |
| 149 | then have "chain ord (f ` A)" "f x \<in> f ` A" | |
| 150 | by (auto simp: img_ord_def intro: chainI dest: chainD) | |
| 151 | thus "?iord x (?ilub A)" | |
| 152 | unfolding inv img_lub_def img_ord_def by (rule lub_upper) | |
| 153 | next | |
| 154 | fix A x assume "chain ?iord A" | |
| 155 | and 1: "\<And>z. z \<in> A \<Longrightarrow> ?iord z x" | |
| 156 | then have "chain ord (f ` A)" | |
| 157 | by (auto simp: img_ord_def intro: chainI dest: chainD) | |
| 158 | thus "?iord (?ilub A) x" | |
| 159 | unfolding inv img_lub_def img_ord_def | |
| 160 | by (rule lub_least) (auto dest: 1[unfolded img_ord_def]) | |
| 161 | qed (auto simp: img_ord_def intro: leq_refl dest: leq_trans leq_antisym inj) | |
| 162 | qed | |
| 163 | ||
| 164 | context partial_function_definitions | |
| 165 | begin | |
| 166 | ||
| 167 | abbreviation "le_fun \<equiv> fun_ord leq" | |
| 168 | abbreviation "lub_fun \<equiv> fun_lub lub" | |
| 46041 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
45297diff
changeset | 169 | abbreviation "fixp_fun \<equiv> ccpo.fixp lub_fun le_fun" | 
| 40107 | 170 | abbreviation "mono_body \<equiv> monotone le_fun leq" | 
| 46041 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
45297diff
changeset | 171 | abbreviation "admissible \<equiv> ccpo.admissible lub_fun le_fun" | 
| 40107 | 172 | |
| 60758 | 173 | text \<open>Interpret manually, to avoid flooding everything with facts about | 
| 174 | orders\<close> | |
| 40107 | 175 | |
| 46041 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
45297diff
changeset | 176 | lemma ccpo: "class.ccpo lub_fun le_fun (mk_less le_fun)" | 
| 40107 | 177 | apply (rule ccpo) | 
| 178 | apply (rule partial_function_lift) | |
| 179 | apply (rule partial_function_definitions_axioms) | |
| 180 | done | |
| 181 | ||
| 60758 | 182 | text \<open>The crucial fixed-point theorem\<close> | 
| 40107 | 183 | |
| 184 | lemma mono_body_fixp: | |
| 185 | "(\<And>x. mono_body (\<lambda>f. F f x)) \<Longrightarrow> fixp_fun F = F (fixp_fun F)" | |
| 186 | by (rule ccpo.fixp_unfold[OF ccpo]) (auto simp: monotone_def fun_ord_def) | |
| 187 | ||
| 60758 | 188 | text \<open>Version with curry/uncurry combinators, to be used by package\<close> | 
| 40107 | 189 | |
| 190 | lemma fixp_rule_uc: | |
| 191 | fixes F :: "'c \<Rightarrow> 'c" and | |
| 192 | U :: "'c \<Rightarrow> 'b \<Rightarrow> 'a" and | |
| 193 |     C :: "('b \<Rightarrow> 'a) \<Rightarrow> 'c"
 | |
| 194 | assumes mono: "\<And>x. mono_body (\<lambda>f. U (F (C f)) x)" | |
| 195 | assumes eq: "f \<equiv> C (fixp_fun (\<lambda>f. U (F (C f))))" | |
| 196 | assumes inverse: "\<And>f. C (U f) = f" | |
| 197 | shows "f = F f" | |
| 198 | proof - | |
| 199 | have "f = C (fixp_fun (\<lambda>f. U (F (C f))))" by (simp add: eq) | |
| 200 | also have "... = C (U (F (C (fixp_fun (\<lambda>f. U (F (C f)))))))" | |
| 201 | by (subst mono_body_fixp[of "%f. U (F (C f))", OF mono]) (rule refl) | |
| 202 | also have "... = F (C (fixp_fun (\<lambda>f. U (F (C f)))))" by (rule inverse) | |
| 203 | also have "... = F f" by (simp add: eq) | |
| 204 | finally show "f = F f" . | |
| 205 | qed | |
| 206 | ||
| 60758 | 207 | text \<open>Fixpoint induction rule\<close> | 
| 43082 
8d0c44de9773
generic fixpoint induction (with explicit curry/uncurry predicates) and instance for option type
 krauss parents: 
43081diff
changeset | 208 | |
| 
8d0c44de9773
generic fixpoint induction (with explicit curry/uncurry predicates) and instance for option type
 krauss parents: 
43081diff
changeset | 209 | lemma fixp_induct_uc: | 
| 59647 
c6f413b660cf
clarified Drule.gen_all: observe context more carefully;
 wenzelm parents: 
59517diff
changeset | 210 | fixes F :: "'c \<Rightarrow> 'c" | 
| 
c6f413b660cf
clarified Drule.gen_all: observe context more carefully;
 wenzelm parents: 
59517diff
changeset | 211 | and U :: "'c \<Rightarrow> 'b \<Rightarrow> 'a" | 
| 
c6f413b660cf
clarified Drule.gen_all: observe context more carefully;
 wenzelm parents: 
59517diff
changeset | 212 |     and C :: "('b \<Rightarrow> 'a) \<Rightarrow> 'c"
 | 
| 
c6f413b660cf
clarified Drule.gen_all: observe context more carefully;
 wenzelm parents: 
59517diff
changeset | 213 |     and P :: "('b \<Rightarrow> 'a) \<Rightarrow> bool"
 | 
| 43082 
8d0c44de9773
generic fixpoint induction (with explicit curry/uncurry predicates) and instance for option type
 krauss parents: 
43081diff
changeset | 214 | assumes mono: "\<And>x. mono_body (\<lambda>f. U (F (C f)) x)" | 
| 59647 
c6f413b660cf
clarified Drule.gen_all: observe context more carefully;
 wenzelm parents: 
59517diff
changeset | 215 | and eq: "f \<equiv> C (fixp_fun (\<lambda>f. U (F (C f))))" | 
| 
c6f413b660cf
clarified Drule.gen_all: observe context more carefully;
 wenzelm parents: 
59517diff
changeset | 216 | and inverse: "\<And>f. U (C f) = f" | 
| 
c6f413b660cf
clarified Drule.gen_all: observe context more carefully;
 wenzelm parents: 
59517diff
changeset | 217 | and adm: "ccpo.admissible lub_fun le_fun P" | 
| 
c6f413b660cf
clarified Drule.gen_all: observe context more carefully;
 wenzelm parents: 
59517diff
changeset | 218 |     and bot: "P (\<lambda>_. lub {})"
 | 
| 
c6f413b660cf
clarified Drule.gen_all: observe context more carefully;
 wenzelm parents: 
59517diff
changeset | 219 | and step: "\<And>f. P (U f) \<Longrightarrow> P (U (F f))" | 
| 43082 
8d0c44de9773
generic fixpoint induction (with explicit curry/uncurry predicates) and instance for option type
 krauss parents: 
43081diff
changeset | 220 | shows "P (U f)" | 
| 
8d0c44de9773
generic fixpoint induction (with explicit curry/uncurry predicates) and instance for option type
 krauss parents: 
43081diff
changeset | 221 | unfolding eq inverse | 
| 75669 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
73411diff
changeset | 222 | proof (rule ccpo.fixp_induct[OF ccpo adm]) | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
73411diff
changeset | 223 | show "monotone le_fun le_fun (\<lambda>f. U (F (C f)))" | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
73411diff
changeset | 224 | using mono by (auto simp: monotone_def fun_ord_def) | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
73411diff
changeset | 225 | next | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
73411diff
changeset | 226 |   show "P (lub_fun {})"
 | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
73411diff
changeset | 227 | by (auto simp: bot fun_lub_def) | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
73411diff
changeset | 228 | next | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
73411diff
changeset | 229 | fix x | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
73411diff
changeset | 230 | assume "P x" | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
73411diff
changeset | 231 | then show "P (U (F (C x)))" | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
73411diff
changeset | 232 | using step[of "C x"] by (simp add: inverse) | 
| 
43f5dfb7fa35
tuned (some HOL lints, by Yecine Megdiche);
 Fabian Huch <huch@in.tum.de> parents: 
73411diff
changeset | 233 | qed | 
| 43082 
8d0c44de9773
generic fixpoint induction (with explicit curry/uncurry predicates) and instance for option type
 krauss parents: 
43081diff
changeset | 234 | |
| 
8d0c44de9773
generic fixpoint induction (with explicit curry/uncurry predicates) and instance for option type
 krauss parents: 
43081diff
changeset | 235 | |
| 69593 | 236 | text \<open>Rules for \<^term>\<open>mono_body\<close>:\<close> | 
| 40107 | 237 | |
| 238 | lemma const_mono[partial_function_mono]: "monotone ord leq (\<lambda>f. c)" | |
| 239 | by (rule monotoneI) (rule leq_refl) | |
| 240 | ||
| 241 | end | |
| 242 | ||
| 243 | ||
| 60758 | 244 | subsection \<open>Flat interpretation: tailrec and option\<close> | 
| 40107 | 245 | |
| 246 | definition | |
| 247 | "flat_ord b x y \<longleftrightarrow> x = b \<or> x = y" | |
| 248 | ||
| 249 | definition | |
| 250 |   "flat_lub b A = (if A \<subseteq> {b} then b else (THE x. x \<in> A - {b}))"
 | |
| 251 | ||
| 252 | lemma flat_interpretation: | |
| 253 | "partial_function_definitions (flat_ord b) (flat_lub b)" | |
| 254 | proof | |
| 255 | fix A x assume 1: "chain (flat_ord b) A" "x \<in> A" | |
| 256 | show "flat_ord b x (flat_lub b A)" | |
| 257 | proof cases | |
| 258 | assume "x = b" | |
| 259 | thus ?thesis by (simp add: flat_ord_def) | |
| 260 | next | |
| 261 | assume "x \<noteq> b" | |
| 262 |     with 1 have "A - {b} = {x}"
 | |
| 263 | by (auto elim: chainE simp: flat_ord_def) | |
| 264 | then have "flat_lub b A = x" | |
| 265 | by (auto simp: flat_lub_def) | |
| 266 | thus ?thesis by (auto simp: flat_ord_def) | |
| 267 | qed | |
| 268 | next | |
| 269 | fix A z assume A: "chain (flat_ord b) A" | |
| 270 | and z: "\<And>x. x \<in> A \<Longrightarrow> flat_ord b x z" | |
| 271 | show "flat_ord b (flat_lub b A) z" | |
| 272 | proof cases | |
| 273 |     assume "A \<subseteq> {b}"
 | |
| 274 | thus ?thesis | |
| 275 | by (auto simp: flat_lub_def flat_ord_def) | |
| 276 | next | |
| 277 |     assume nb: "\<not> A \<subseteq> {b}"
 | |
| 278 | then obtain y where y: "y \<in> A" "y \<noteq> b" by auto | |
| 279 |     with A have "A - {b} = {y}"
 | |
| 280 | by (auto elim: chainE simp: flat_ord_def) | |
| 281 | with nb have "flat_lub b A = y" | |
| 282 | by (auto simp: flat_lub_def) | |
| 283 | with z y show ?thesis by auto | |
| 284 | qed | |
| 285 | qed (auto simp: flat_ord_def) | |
| 286 | ||
| 59517 | 287 | lemma flat_ordI: "(x \<noteq> a \<Longrightarrow> x = y) \<Longrightarrow> flat_ord a x y" | 
| 288 | by(auto simp add: flat_ord_def) | |
| 289 | ||
| 290 | lemma flat_ord_antisym: "\<lbrakk> flat_ord a x y; flat_ord a y x \<rbrakk> \<Longrightarrow> x = y" | |
| 291 | by(auto simp add: flat_ord_def) | |
| 292 | ||
| 64634 | 293 | lemma antisymp_flat_ord: "antisymp (flat_ord a)" | 
| 294 | by(rule antisympI)(auto dest: flat_ord_antisym) | |
| 59517 | 295 | |
| 61605 | 296 | interpretation tailrec: | 
| 40107 | 297 | partial_function_definitions "flat_ord undefined" "flat_lub undefined" | 
| 61566 
c3d6e570ccef
Keyword 'rewrites' identifies rewrite morphisms.
 ballarin parents: 
60758diff
changeset | 298 |   rewrites "flat_lub undefined {} \<equiv> undefined"
 | 
| 54630 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
53949diff
changeset | 299 | by (rule flat_interpretation)(simp add: flat_lub_def) | 
| 40107 | 300 | |
| 61605 | 301 | interpretation option: | 
| 40107 | 302 | partial_function_definitions "flat_ord None" "flat_lub None" | 
| 61566 
c3d6e570ccef
Keyword 'rewrites' identifies rewrite morphisms.
 ballarin parents: 
60758diff
changeset | 303 |   rewrites "flat_lub None {} \<equiv> None"
 | 
| 54630 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
53949diff
changeset | 304 | by (rule flat_interpretation)(simp add: flat_lub_def) | 
| 40107 | 305 | |
| 42949 
618adb3584e5
separate initializations for different modes of partial_function -- generation of induction rules will be non-uniform
 krauss parents: 
40288diff
changeset | 306 | |
| 51459 
bc3651180a09
add induction rule for partial_function (tailrec)
 Andreas Lochbihler parents: 
48891diff
changeset | 307 | abbreviation "tailrec_ord \<equiv> flat_ord undefined" | 
| 
bc3651180a09
add induction rule for partial_function (tailrec)
 Andreas Lochbihler parents: 
48891diff
changeset | 308 | abbreviation "mono_tailrec \<equiv> monotone (fun_ord tailrec_ord) tailrec_ord" | 
| 
bc3651180a09
add induction rule for partial_function (tailrec)
 Andreas Lochbihler parents: 
48891diff
changeset | 309 | |
| 
bc3651180a09
add induction rule for partial_function (tailrec)
 Andreas Lochbihler parents: 
48891diff
changeset | 310 | lemma tailrec_admissible: | 
| 53949 | 311 | "ccpo.admissible (fun_lub (flat_lub c)) (fun_ord (flat_ord c)) | 
| 312 | (\<lambda>a. \<forall>x. a x \<noteq> c \<longrightarrow> P x (a x))" | |
| 53361 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
52728diff
changeset | 313 | proof(intro ccpo.admissibleI strip) | 
| 51459 
bc3651180a09
add induction rule for partial_function (tailrec)
 Andreas Lochbihler parents: 
48891diff
changeset | 314 | fix A x | 
| 53949 | 315 | assume chain: "Complete_Partial_Order.chain (fun_ord (flat_ord c)) A" | 
| 316 | and P [rule_format]: "\<forall>f\<in>A. \<forall>x. f x \<noteq> c \<longrightarrow> P x (f x)" | |
| 317 | and defined: "fun_lub (flat_lub c) A x \<noteq> c" | |
| 318 | from defined obtain f where f: "f \<in> A" "f x \<noteq> c" | |
| 62390 | 319 | by(auto simp add: fun_lub_def flat_lub_def split: if_split_asm) | 
| 51459 
bc3651180a09
add induction rule for partial_function (tailrec)
 Andreas Lochbihler parents: 
48891diff
changeset | 320 | hence "P x (f x)" by(rule P) | 
| 53949 | 321 | moreover from chain f have "\<forall>f' \<in> A. f' x = c \<or> f' x = f x" | 
| 51459 
bc3651180a09
add induction rule for partial_function (tailrec)
 Andreas Lochbihler parents: 
48891diff
changeset | 322 | by(auto 4 4 simp add: Complete_Partial_Order.chain_def flat_ord_def fun_ord_def) | 
| 53949 | 323 | hence "fun_lub (flat_lub c) A x = f x" | 
| 51459 
bc3651180a09
add induction rule for partial_function (tailrec)
 Andreas Lochbihler parents: 
48891diff
changeset | 324 | using f by(auto simp add: fun_lub_def flat_lub_def) | 
| 53949 | 325 | ultimately show "P x (fun_lub (flat_lub c) A x)" by simp | 
| 51459 
bc3651180a09
add induction rule for partial_function (tailrec)
 Andreas Lochbihler parents: 
48891diff
changeset | 326 | qed | 
| 
bc3651180a09
add induction rule for partial_function (tailrec)
 Andreas Lochbihler parents: 
48891diff
changeset | 327 | |
| 
bc3651180a09
add induction rule for partial_function (tailrec)
 Andreas Lochbihler parents: 
48891diff
changeset | 328 | lemma fixp_induct_tailrec: | 
| 
bc3651180a09
add induction rule for partial_function (tailrec)
 Andreas Lochbihler parents: 
48891diff
changeset | 329 | fixes F :: "'c \<Rightarrow> 'c" and | 
| 
bc3651180a09
add induction rule for partial_function (tailrec)
 Andreas Lochbihler parents: 
48891diff
changeset | 330 | U :: "'c \<Rightarrow> 'b \<Rightarrow> 'a" and | 
| 
bc3651180a09
add induction rule for partial_function (tailrec)
 Andreas Lochbihler parents: 
48891diff
changeset | 331 |     C :: "('b \<Rightarrow> 'a) \<Rightarrow> 'c" and
 | 
| 
bc3651180a09
add induction rule for partial_function (tailrec)
 Andreas Lochbihler parents: 
48891diff
changeset | 332 | P :: "'b \<Rightarrow> 'a \<Rightarrow> bool" and | 
| 
bc3651180a09
add induction rule for partial_function (tailrec)
 Andreas Lochbihler parents: 
48891diff
changeset | 333 | x :: "'b" | 
| 53949 | 334 | assumes mono: "\<And>x. monotone (fun_ord (flat_ord c)) (flat_ord c) (\<lambda>f. U (F (C f)) x)" | 
| 335 | assumes eq: "f \<equiv> C (ccpo.fixp (fun_lub (flat_lub c)) (fun_ord (flat_ord c)) (\<lambda>f. U (F (C f))))" | |
| 51459 
bc3651180a09
add induction rule for partial_function (tailrec)
 Andreas Lochbihler parents: 
48891diff
changeset | 336 | assumes inverse2: "\<And>f. U (C f) = f" | 
| 53949 | 337 | assumes step: "\<And>f x y. (\<And>x y. U f x = y \<Longrightarrow> y \<noteq> c \<Longrightarrow> P x y) \<Longrightarrow> U (F f) x = y \<Longrightarrow> y \<noteq> c \<Longrightarrow> P x y" | 
| 51459 
bc3651180a09
add induction rule for partial_function (tailrec)
 Andreas Lochbihler parents: 
48891diff
changeset | 338 | assumes result: "U f x = y" | 
| 53949 | 339 | assumes defined: "y \<noteq> c" | 
| 51459 
bc3651180a09
add induction rule for partial_function (tailrec)
 Andreas Lochbihler parents: 
48891diff
changeset | 340 | shows "P x y" | 
| 
bc3651180a09
add induction rule for partial_function (tailrec)
 Andreas Lochbihler parents: 
48891diff
changeset | 341 | proof - | 
| 53949 | 342 | have "\<forall>x y. U f x = y \<longrightarrow> y \<noteq> c \<longrightarrow> P x y" | 
| 343 | by(rule partial_function_definitions.fixp_induct_uc[OF flat_interpretation, of _ U F C, OF mono eq inverse2]) | |
| 54630 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
53949diff
changeset | 344 | (auto intro: step tailrec_admissible simp add: fun_lub_def flat_lub_def) | 
| 51459 
bc3651180a09
add induction rule for partial_function (tailrec)
 Andreas Lochbihler parents: 
48891diff
changeset | 345 | thus ?thesis using result defined by blast | 
| 
bc3651180a09
add induction rule for partial_function (tailrec)
 Andreas Lochbihler parents: 
48891diff
changeset | 346 | qed | 
| 
bc3651180a09
add induction rule for partial_function (tailrec)
 Andreas Lochbihler parents: 
48891diff
changeset | 347 | |
| 51485 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 krauss parents: 
51459diff
changeset | 348 | lemma admissible_image: | 
| 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 krauss parents: 
51459diff
changeset | 349 | assumes pfun: "partial_function_definitions le lub" | 
| 67091 | 350 | assumes adm: "ccpo.admissible lub le (P \<circ> g)" | 
| 51485 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 krauss parents: 
51459diff
changeset | 351 | assumes inj: "\<And>x y. f x = f y \<Longrightarrow> x = y" | 
| 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 krauss parents: 
51459diff
changeset | 352 | assumes inv: "\<And>x. f (g x) = x" | 
| 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 krauss parents: 
51459diff
changeset | 353 | shows "ccpo.admissible (img_lub f g lub) (img_ord f le) P" | 
| 53361 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
52728diff
changeset | 354 | proof (rule ccpo.admissibleI) | 
| 51485 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 krauss parents: 
51459diff
changeset | 355 | fix A assume "chain (img_ord f le) A" | 
| 54630 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
53949diff
changeset | 356 | then have ch': "chain le (f ` A)" | 
| 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
53949diff
changeset | 357 | by (auto simp: img_ord_def intro: chainI dest: chainD) | 
| 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
53949diff
changeset | 358 |   assume "A \<noteq> {}"
 | 
| 51485 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 krauss parents: 
51459diff
changeset | 359 | assume P_A: "\<forall>x\<in>A. P x" | 
| 67091 | 360 | have "(P \<circ> g) (lub (f ` A))" using adm ch' | 
| 53361 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
52728diff
changeset | 361 | proof (rule ccpo.admissibleD) | 
| 51485 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 krauss parents: 
51459diff
changeset | 362 | fix x assume "x \<in> f ` A" | 
| 67091 | 363 | with P_A show "(P \<circ> g) x" by (auto simp: inj[OF inv]) | 
| 60758 | 364 |   qed(simp add: \<open>A \<noteq> {}\<close>)
 | 
| 51485 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 krauss parents: 
51459diff
changeset | 365 | thus "P (img_lub f g lub A)" unfolding img_lub_def by simp | 
| 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 krauss parents: 
51459diff
changeset | 366 | qed | 
| 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 krauss parents: 
51459diff
changeset | 367 | |
| 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 krauss parents: 
51459diff
changeset | 368 | lemma admissible_fun: | 
| 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 krauss parents: 
51459diff
changeset | 369 | assumes pfun: "partial_function_definitions le lub" | 
| 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 krauss parents: 
51459diff
changeset | 370 | assumes adm: "\<And>x. ccpo.admissible lub le (Q x)" | 
| 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 krauss parents: 
51459diff
changeset | 371 | shows "ccpo.admissible (fun_lub lub) (fun_ord le) (\<lambda>f. \<forall>x. Q x (f x))" | 
| 53361 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
52728diff
changeset | 372 | proof (rule ccpo.admissibleI) | 
| 51485 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 krauss parents: 
51459diff
changeset | 373 |   fix A :: "('b \<Rightarrow> 'a) set"
 | 
| 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 krauss parents: 
51459diff
changeset | 374 | assume Q: "\<forall>f\<in>A. \<forall>x. Q x (f x)" | 
| 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 krauss parents: 
51459diff
changeset | 375 | assume ch: "chain (fun_ord le) A" | 
| 54630 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
53949diff
changeset | 376 |   assume "A \<noteq> {}"
 | 
| 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
53949diff
changeset | 377 |   hence non_empty: "\<And>a. {y. \<exists>f\<in>A. y = f a} \<noteq> {}" by auto
 | 
| 51485 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 krauss parents: 
51459diff
changeset | 378 | show "\<forall>x. Q x (fun_lub lub A x)" | 
| 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 krauss parents: 
51459diff
changeset | 379 | unfolding fun_lub_def | 
| 54630 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
53949diff
changeset | 380 | by (rule allI, rule ccpo.admissibleD[OF adm chain_fun[OF ch] non_empty]) | 
| 51485 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 krauss parents: 
51459diff
changeset | 381 | (auto simp: Q) | 
| 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 krauss parents: 
51459diff
changeset | 382 | qed | 
| 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 krauss parents: 
51459diff
changeset | 383 | |
| 51459 
bc3651180a09
add induction rule for partial_function (tailrec)
 Andreas Lochbihler parents: 
48891diff
changeset | 384 | |
| 40107 | 385 | abbreviation "option_ord \<equiv> flat_ord None" | 
| 386 | abbreviation "mono_option \<equiv> monotone (fun_ord option_ord) option_ord" | |
| 387 | ||
| 388 | lemma bind_mono[partial_function_mono]: | |
| 389 | assumes mf: "mono_option B" and mg: "\<And>y. mono_option (\<lambda>f. C y f)" | |
| 390 | shows "mono_option (\<lambda>f. Option.bind (B f) (\<lambda>y. C y f))" | |
| 391 | proof (rule monotoneI) | |
| 392 | fix f g :: "'a \<Rightarrow> 'b option" assume fg: "fun_ord option_ord f g" | |
| 393 | with mf | |
| 394 | have "option_ord (B f) (B g)" by (rule monotoneD[of _ _ _ f g]) | |
| 395 | then have "option_ord (Option.bind (B f) (\<lambda>y. C y f)) (Option.bind (B g) (\<lambda>y. C y f))" | |
| 396 | unfolding flat_ord_def by auto | |
| 397 | also from mg | |
| 398 | have "\<And>y'. option_ord (C y' f) (C y' g)" | |
| 399 | by (rule monotoneD) (rule fg) | |
| 400 | then have "option_ord (Option.bind (B g) (\<lambda>y'. C y' f)) (Option.bind (B g) (\<lambda>y'. C y' g))" | |
| 401 | unfolding flat_ord_def by (cases "B g") auto | |
| 402 | finally (option.leq_trans) | |
| 403 | show "option_ord (Option.bind (B f) (\<lambda>y. C y f)) (Option.bind (B g) (\<lambda>y'. C y' g))" . | |
| 404 | qed | |
| 405 | ||
| 43081 | 406 | lemma flat_lub_in_chain: | 
| 407 | assumes ch: "chain (flat_ord b) A " | |
| 408 | assumes lub: "flat_lub b A = a" | |
| 409 | shows "a = b \<or> a \<in> A" | |
| 410 | proof (cases "A \<subseteq> {b}")
 | |
| 411 | case True | |
| 412 | then have "flat_lub b A = b" unfolding flat_lub_def by simp | |
| 413 | with lub show ?thesis by simp | |
| 414 | next | |
| 415 | case False | |
| 416 | then obtain c where "c \<in> A" and "c \<noteq> b" by auto | |
| 417 |   { fix z assume "z \<in> A"
 | |
| 60758 | 418 | from chainD[OF ch \<open>c \<in> A\<close> this] have "z = c \<or> z = b" | 
| 419 | unfolding flat_ord_def using \<open>c \<noteq> b\<close> by auto } | |
| 43081 | 420 |   with False have "A - {b} = {c}" by auto
 | 
| 421 | with False have "flat_lub b A = c" by (auto simp: flat_lub_def) | |
| 60758 | 422 | with \<open>c \<in> A\<close> lub show ?thesis by simp | 
| 43081 | 423 | qed | 
| 424 | ||
| 425 | lemma option_admissible: "option.admissible (%(f::'a \<Rightarrow> 'b option). | |
| 426 | (\<forall>x y. f x = Some y \<longrightarrow> P x y))" | |
| 53361 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 Andreas Lochbihler parents: 
52728diff
changeset | 427 | proof (rule ccpo.admissibleI) | 
| 43081 | 428 |   fix A :: "('a \<Rightarrow> 'b option) set"
 | 
| 429 | assume ch: "chain option.le_fun A" | |
| 430 | and IH: "\<forall>f\<in>A. \<forall>x y. f x = Some y \<longrightarrow> P x y" | |
| 431 |   from ch have ch': "\<And>x. chain option_ord {y. \<exists>f\<in>A. y = f x}" by (rule chain_fun)
 | |
| 432 | show "\<forall>x y. option.lub_fun A x = Some y \<longrightarrow> P x y" | |
| 433 | proof (intro allI impI) | |
| 434 | fix x y assume "option.lub_fun A x = Some y" | |
| 435 | from flat_lub_in_chain[OF ch' this[unfolded fun_lub_def]] | |
| 436 |     have "Some y \<in> {y. \<exists>f\<in>A. y = f x}" by simp
 | |
| 437 | then have "\<exists>f\<in>A. f x = Some y" by auto | |
| 438 | with IH show "P x y" by auto | |
| 439 | qed | |
| 440 | qed | |
| 441 | ||
| 43082 
8d0c44de9773
generic fixpoint induction (with explicit curry/uncurry predicates) and instance for option type
 krauss parents: 
43081diff
changeset | 442 | lemma fixp_induct_option: | 
| 
8d0c44de9773
generic fixpoint induction (with explicit curry/uncurry predicates) and instance for option type
 krauss parents: 
43081diff
changeset | 443 | fixes F :: "'c \<Rightarrow> 'c" and | 
| 
8d0c44de9773
generic fixpoint induction (with explicit curry/uncurry predicates) and instance for option type
 krauss parents: 
43081diff
changeset | 444 | U :: "'c \<Rightarrow> 'b \<Rightarrow> 'a option" and | 
| 
8d0c44de9773
generic fixpoint induction (with explicit curry/uncurry predicates) and instance for option type
 krauss parents: 
43081diff
changeset | 445 |     C :: "('b \<Rightarrow> 'a option) \<Rightarrow> 'c" and
 | 
| 
8d0c44de9773
generic fixpoint induction (with explicit curry/uncurry predicates) and instance for option type
 krauss parents: 
43081diff
changeset | 446 | P :: "'b \<Rightarrow> 'a \<Rightarrow> bool" | 
| 
8d0c44de9773
generic fixpoint induction (with explicit curry/uncurry predicates) and instance for option type
 krauss parents: 
43081diff
changeset | 447 | assumes mono: "\<And>x. mono_option (\<lambda>f. U (F (C f)) x)" | 
| 46041 
1e3ff542e83e
remove constant 'ccpo.lub', re-use constant 'Sup' instead
 huffman parents: 
45297diff
changeset | 448 | assumes eq: "f \<equiv> C (ccpo.fixp (fun_lub (flat_lub None)) (fun_ord option_ord) (\<lambda>f. U (F (C f))))" | 
| 43082 
8d0c44de9773
generic fixpoint induction (with explicit curry/uncurry predicates) and instance for option type
 krauss parents: 
43081diff
changeset | 449 | assumes inverse2: "\<And>f. U (C f) = f" | 
| 
8d0c44de9773
generic fixpoint induction (with explicit curry/uncurry predicates) and instance for option type
 krauss parents: 
43081diff
changeset | 450 | assumes step: "\<And>f x y. (\<And>x y. U f x = Some y \<Longrightarrow> P x y) \<Longrightarrow> U (F f) x = Some y \<Longrightarrow> P x y" | 
| 
8d0c44de9773
generic fixpoint induction (with explicit curry/uncurry predicates) and instance for option type
 krauss parents: 
43081diff
changeset | 451 | assumes defined: "U f x = Some y" | 
| 
8d0c44de9773
generic fixpoint induction (with explicit curry/uncurry predicates) and instance for option type
 krauss parents: 
43081diff
changeset | 452 | shows "P x y" | 
| 
8d0c44de9773
generic fixpoint induction (with explicit curry/uncurry predicates) and instance for option type
 krauss parents: 
43081diff
changeset | 453 | using step defined option.fixp_induct_uc[of U F C, OF mono eq inverse2 option_admissible] | 
| 54630 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 Andreas Lochbihler parents: 
53949diff
changeset | 454 | unfolding fun_lub_def flat_lub_def by(auto 9 2) | 
| 43082 
8d0c44de9773
generic fixpoint induction (with explicit curry/uncurry predicates) and instance for option type
 krauss parents: 
43081diff
changeset | 455 | |
| 69593 | 456 | declaration \<open>Partial_Function.init "tailrec" \<^term>\<open>tailrec.fixp_fun\<close> | 
| 457 |   \<^term>\<open>tailrec.mono_body\<close> @{thm tailrec.fixp_rule_uc} @{thm tailrec.fixp_induct_uc}
 | |
| 61841 
4d3527b94f2a
more general types Proof.method / context_tactic;
 wenzelm parents: 
61605diff
changeset | 458 |   (SOME @{thm fixp_induct_tailrec[where c = undefined]})\<close>
 | 
| 43082 
8d0c44de9773
generic fixpoint induction (with explicit curry/uncurry predicates) and instance for option type
 krauss parents: 
43081diff
changeset | 459 | |
| 69593 | 460 | declaration \<open>Partial_Function.init "option" \<^term>\<open>option.fixp_fun\<close> | 
| 461 |   \<^term>\<open>option.mono_body\<close> @{thm option.fixp_rule_uc} @{thm option.fixp_induct_uc}
 | |
| 60758 | 462 |   (SOME @{thm fixp_induct_option})\<close>
 | 
| 43082 
8d0c44de9773
generic fixpoint induction (with explicit curry/uncurry predicates) and instance for option type
 krauss parents: 
43081diff
changeset | 463 | |
| 40252 
029400b6c893
hide_const various constants, in particular to avoid ugly qualifiers in HOLCF
 krauss parents: 
40107diff
changeset | 464 | hide_const (open) chain | 
| 40107 | 465 | |
| 466 | end |