| author | haftmann | 
| Thu, 10 May 2007 10:21:48 +0200 | |
| changeset 22919 | 3de2d0b5b89a | 
| parent 20523 | 36a59e5d0039 | 
| child 25758 | 89c7c22e64b4 | 
| permissions | -rw-r--r-- | 
| 16202 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 1 | (* Title: HOLCF/FunCpo.thy | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 2 | ID: $Id$ | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 3 | Author: Franz Regensburger | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 4 | |
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 5 | Definition of the partial ordering for the type of all functions => (fun) | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 6 | |
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 7 | Class instance of => (fun) for class pcpo. | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 8 | *) | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 9 | |
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 10 | header {* Class instances for the full function space *}
 | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 11 | |
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 12 | theory Ffun | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 13 | imports Pcpo | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 14 | begin | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 15 | |
| 18291 | 16 | subsection {* Full function space is a partial order *}
 | 
| 16202 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 17 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
18291diff
changeset | 18 | instance "fun" :: (type, sq_ord) sq_ord .. | 
| 16202 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 19 | |
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 20 | defs (overloaded) | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 21 | less_fun_def: "(op \<sqsubseteq>) \<equiv> (\<lambda>f g. \<forall>x. f x \<sqsubseteq> g x)" | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 22 | |
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 23 | lemma refl_less_fun: "(f::'a::type \<Rightarrow> 'b::po) \<sqsubseteq> f" | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 24 | by (simp add: less_fun_def) | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 25 | |
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 26 | lemma antisym_less_fun: | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 27 | "\<lbrakk>(f1::'a::type \<Rightarrow> 'b::po) \<sqsubseteq> f2; f2 \<sqsubseteq> f1\<rbrakk> \<Longrightarrow> f1 = f2" | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 28 | by (simp add: less_fun_def expand_fun_eq antisym_less) | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 29 | |
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 30 | lemma trans_less_fun: | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 31 | "\<lbrakk>(f1::'a::type \<Rightarrow> 'b::po) \<sqsubseteq> f2; f2 \<sqsubseteq> f3\<rbrakk> \<Longrightarrow> f1 \<sqsubseteq> f3" | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 32 | apply (unfold less_fun_def) | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 33 | apply clarify | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 34 | apply (rule trans_less) | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 35 | apply (erule spec) | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 36 | apply (erule spec) | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 37 | done | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 38 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
18291diff
changeset | 39 | instance "fun" :: (type, po) po | 
| 16202 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 40 | by intro_classes | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 41 | (assumption | rule refl_less_fun antisym_less_fun trans_less_fun)+ | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 42 | |
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 43 | text {* make the symbol @{text "<<"} accessible for type fun *}
 | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 44 | |
| 17831 | 45 | lemma expand_fun_less: "(f \<sqsubseteq> g) = (\<forall>x. f x \<sqsubseteq> g x)" | 
| 16202 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 46 | by (simp add: less_fun_def) | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 47 | |
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 48 | lemma less_fun_ext: "(\<And>x. f x \<sqsubseteq> g x) \<Longrightarrow> f \<sqsubseteq> g" | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 49 | by (simp add: less_fun_def) | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 50 | |
| 18291 | 51 | subsection {* Full function space is chain complete *}
 | 
| 16202 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 52 | |
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 53 | text {* chains of functions yield chains in the po range *}
 | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 54 | |
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 55 | lemma ch2ch_fun: "chain S \<Longrightarrow> chain (\<lambda>i. S i x)" | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 56 | by (simp add: chain_def less_fun_def) | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 57 | |
| 18092 
2c5d5da79a1e
renamed and added ch2ch, cont2cont, mono2mono theorems ending in _fun, _lambda, _LAM
 huffman parents: 
17831diff
changeset | 58 | lemma ch2ch_lambda: "(\<And>x. chain (\<lambda>i. S i x)) \<Longrightarrow> chain S" | 
| 16202 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 59 | by (simp add: chain_def less_fun_def) | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 60 | |
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 61 | |
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 62 | text {* upper bounds of function chains yield upper bound in the po range *}
 | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 63 | |
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 64 | lemma ub2ub_fun: | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 65 | "range (S::nat \<Rightarrow> 'a \<Rightarrow> 'b::po) <| u \<Longrightarrow> range (\<lambda>i. S i x) <| u x" | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 66 | by (auto simp add: is_ub_def less_fun_def) | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 67 | |
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 68 | text {* Type @{typ "'a::type => 'b::cpo"} is chain complete *}
 | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 69 | |
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 70 | lemma lub_fun: | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 71 | "chain (S::nat \<Rightarrow> 'a::type \<Rightarrow> 'b::cpo) | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 72 | \<Longrightarrow> range S <<| (\<lambda>x. \<Squnion>i. S i x)" | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 73 | apply (rule is_lubI) | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 74 | apply (rule ub_rangeI) | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 75 | apply (rule less_fun_ext) | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 76 | apply (rule is_ub_thelub) | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 77 | apply (erule ch2ch_fun) | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 78 | apply (rule less_fun_ext) | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 79 | apply (rule is_lub_thelub) | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 80 | apply (erule ch2ch_fun) | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 81 | apply (erule ub2ub_fun) | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 82 | done | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 83 | |
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 84 | lemma thelub_fun: | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 85 | "chain (S::nat \<Rightarrow> 'a::type \<Rightarrow> 'b::cpo) | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 86 | \<Longrightarrow> lub (range S) = (\<lambda>x. \<Squnion>i. S i x)" | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 87 | by (rule lub_fun [THEN thelubI]) | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 88 | |
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 89 | lemma cpo_fun: | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 90 | "chain (S::nat \<Rightarrow> 'a::type \<Rightarrow> 'b::cpo) \<Longrightarrow> \<exists>x. range S <<| x" | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 91 | by (rule exI, erule lub_fun) | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 92 | |
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
18291diff
changeset | 93 | instance "fun" :: (type, cpo) cpo | 
| 16202 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 94 | by intro_classes (rule cpo_fun) | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 95 | |
| 18291 | 96 | subsection {* Full function space is pointed *}
 | 
| 17831 | 97 | |
| 98 | lemma minimal_fun: "(\<lambda>x. \<bottom>) \<sqsubseteq> f" | |
| 99 | by (simp add: less_fun_def) | |
| 100 | ||
| 101 | lemma least_fun: "\<exists>x::'a \<Rightarrow> 'b::pcpo. \<forall>y. x \<sqsubseteq> y" | |
| 102 | apply (rule_tac x = "\<lambda>x. \<bottom>" in exI) | |
| 103 | apply (rule minimal_fun [THEN allI]) | |
| 104 | done | |
| 105 | ||
| 20523 
36a59e5d0039
Major update to function package, including new syntax and the (only theoretical)
 krauss parents: 
18291diff
changeset | 106 | instance "fun" :: (type, pcpo) pcpo | 
| 16202 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 107 | by intro_classes (rule least_fun) | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 108 | |
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 109 | text {* for compatibility with old HOLCF-Version *}
 | 
| 17831 | 110 | lemma inst_fun_pcpo: "\<bottom> = (\<lambda>x. \<bottom>)" | 
| 16202 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 111 | by (rule minimal_fun [THEN UU_I, symmetric]) | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 112 | |
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 113 | text {* function application is strict in the left argument *}
 | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 114 | lemma app_strict [simp]: "\<bottom> x = \<bottom>" | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 115 | by (simp add: inst_fun_pcpo) | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 116 | |
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 117 | end | 
| 
61811f31ce5a
renamed FunCpo theory to Ffun; added theorems ch2ch_fun_rev and app_strict
 huffman parents: diff
changeset | 118 |