| author | wenzelm | 
| Thu, 20 Aug 2015 17:39:07 +0200 | |
| changeset 60986 | 077f663b6c24 | 
| parent 53374 | a14d2a854c02 | 
| permissions | -rw-r--r-- | 
| 53109 
186535065f5c
renamed theory Mrec to Legacy_Mrec, no longer included by default
 krauss parents: 
44890diff
changeset | 1 | theory Legacy_Mrec | 
| 37772 | 2 | imports Heap_Monad | 
| 3 | begin | |
| 4 | ||
| 5 | subsubsection {* A monadic combinator for simple recursive functions *}
 | |
| 6 | ||
| 53109 
186535065f5c
renamed theory Mrec to Legacy_Mrec, no longer included by default
 krauss parents: 
44890diff
changeset | 7 | text {*
 | 
| 
186535065f5c
renamed theory Mrec to Legacy_Mrec, no longer included by default
 krauss parents: 
44890diff
changeset | 8 | NOTE: The use of this obsolete combinator is discouraged. Instead, use the | 
| 
186535065f5c
renamed theory Mrec to Legacy_Mrec, no longer included by default
 krauss parents: 
44890diff
changeset | 9 |   @{text "partal_function (heap)"} command.
 | 
| 
186535065f5c
renamed theory Mrec to Legacy_Mrec, no longer included by default
 krauss parents: 
44890diff
changeset | 10 | *} | 
| 
186535065f5c
renamed theory Mrec to Legacy_Mrec, no longer included by default
 krauss parents: 
44890diff
changeset | 11 | |
| 37772 | 12 | text {* Using a locale to fix arguments f and g of MREC *}
 | 
| 13 | ||
| 14 | locale mrec = | |
| 15 |   fixes f :: "'a \<Rightarrow> ('b + 'a) Heap"
 | |
| 16 | and g :: "'a \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'b Heap" | |
| 17 | begin | |
| 18 | ||
| 19 | function (default "\<lambda>(x, h). None") mrec :: "'a \<Rightarrow> heap \<Rightarrow> ('b \<times> heap) option" where
 | |
| 20 | "mrec x h = (case execute (f x) h of | |
| 21 | Some (Inl r, h') \<Rightarrow> Some (r, h') | |
| 22 | | Some (Inr s, h') \<Rightarrow> (case mrec s h' of | |
| 23 | Some (z, h'') \<Rightarrow> execute (g x s z) h'' | |
| 24 | | None \<Rightarrow> None) | |
| 25 | | None \<Rightarrow> None)" | |
| 26 | by auto | |
| 27 | ||
| 28 | lemma graph_implies_dom: | |
| 29 | "mrec_graph x y \<Longrightarrow> mrec_dom x" | |
| 30 | apply (induct rule:mrec_graph.induct) | |
| 31 | apply (rule accpI) | |
| 32 | apply (erule mrec_rel.cases) | |
| 33 | by simp | |
| 34 | ||
| 35 | lemma mrec_default: "\<not> mrec_dom (x, h) \<Longrightarrow> mrec x h = None" | |
| 36 | unfolding mrec_def | |
| 37 | by (rule fundef_default_value[OF mrec_sumC_def graph_implies_dom, of _ _ "(x, h)", simplified]) | |
| 38 | ||
| 39 | lemma mrec_di_reverse: | |
| 40 | assumes "\<not> mrec_dom (x, h)" | |
| 41 | shows " | |
| 42 | (case execute (f x) h of | |
| 43 | Some (Inl r, h') \<Rightarrow> False | |
| 44 | | Some (Inr s, h') \<Rightarrow> \<not> mrec_dom (s, h') | |
| 45 | | None \<Rightarrow> False | |
| 46 | )" | |
| 47 | using assms apply (auto split: option.split sum.split) | |
| 48 | apply (rule ccontr) | |
| 49 | apply (erule notE, rule accpI, elim mrec_rel.cases, auto)+ | |
| 50 | done | |
| 51 | ||
| 52 | lemma mrec_rule: | |
| 53 | "mrec x h = | |
| 54 | (case execute (f x) h of | |
| 55 | Some (Inl r, h') \<Rightarrow> Some (r, h') | |
| 56 | | Some (Inr s, h') \<Rightarrow> | |
| 57 | (case mrec s h' of | |
| 58 | Some (z, h'') \<Rightarrow> execute (g x s z) h'' | |
| 59 | | None \<Rightarrow> None) | |
| 60 | | None \<Rightarrow> None | |
| 61 | )" | |
| 39754 | 62 | apply (cases "mrec_dom (x,h)", simp add: mrec.psimps) | 
| 37772 | 63 | apply (frule mrec_default) | 
| 64 | apply (frule mrec_di_reverse, simp) | |
| 65 | by (auto split: sum.split option.split simp: mrec_default) | |
| 66 | ||
| 67 | definition | |
| 68 | "MREC x = Heap_Monad.Heap (mrec x)" | |
| 69 | ||
| 70 | lemma MREC_rule: | |
| 71 | "MREC x = | |
| 37792 | 72 |   do { y \<leftarrow> f x;
 | 
| 37772 | 73 | (case y of | 
| 74 | Inl r \<Rightarrow> return r | |
| 75 | | Inr s \<Rightarrow> | |
| 37792 | 76 |                 do { z \<leftarrow> MREC s ;
 | 
| 77 | g x s z })}" | |
| 37772 | 78 | unfolding MREC_def | 
| 79 | unfolding bind_def return_def | |
| 80 | apply simp | |
| 81 | apply (rule ext) | |
| 82 | apply (unfold mrec_rule[of x]) | |
| 37787 
30dc3abf4a58
theorem collections do not contain default rules any longer
 haftmann parents: 
37772diff
changeset | 83 | by (auto simp add: execute_simps split: option.splits prod.splits sum.splits) | 
| 37772 | 84 | |
| 85 | lemma MREC_pinduct: | |
| 86 | assumes "execute (MREC x) h = Some (r, h')" | |
| 87 | assumes non_rec_case: "\<And> x h h' r. execute (f x) h = Some (Inl r, h') \<Longrightarrow> P x h h' r" | |
| 88 | assumes rec_case: "\<And> x h h1 h2 h' s z r. execute (f x) h = Some (Inr s, h1) \<Longrightarrow> execute (MREC s) h1 = Some (z, h2) \<Longrightarrow> P s h1 h2 z | |
| 89 | \<Longrightarrow> execute (g x s z) h2 = Some (r, h') \<Longrightarrow> P x h h' r" | |
| 90 | shows "P x h h' r" | |
| 91 | proof - | |
| 92 | from assms(1) have mrec: "mrec x h = Some (r, h')" | |
| 93 | unfolding MREC_def execute.simps . | |
| 94 | from mrec have dom: "mrec_dom (x, h)" | |
| 95 | apply - | |
| 96 | apply (rule ccontr) | |
| 97 | apply (drule mrec_default) by auto | |
| 98 | from mrec have h'_r: "h' = snd (the (mrec x h))" "r = fst (the (mrec x h))" | |
| 99 | by auto | |
| 100 | from mrec have "P x h (snd (the (mrec x h))) (fst (the (mrec x h)))" | |
| 101 | proof (induct arbitrary: r h' rule: mrec.pinduct[OF dom]) | |
| 102 | case (1 x h) | |
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
40671diff
changeset | 103 | obtain rr h' where "the (mrec x h) = (rr, h')" by fastforce | 
| 37772 | 104 | show ?case | 
| 105 | proof (cases "execute (f x) h") | |
| 106 | case (Some result) | |
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
40671diff
changeset | 107 | then obtain a h1 where exec_f: "execute (f x) h = Some (a, h1)" by fastforce | 
| 37772 | 108 | note Inl' = this | 
| 109 | show ?thesis | |
| 110 | proof (cases a) | |
| 111 | case (Inl aa) | |
| 112 | from this Inl' 1(1) exec_f mrec non_rec_case show ?thesis | |
| 39754 | 113 | by (auto simp: mrec.psimps) | 
| 37772 | 114 | next | 
| 115 | case (Inr b) | |
| 116 | note Inr' = this | |
| 117 | show ?thesis | |
| 118 | proof (cases "mrec b h1") | |
| 119 | case (Some result) | |
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
40671diff
changeset | 120 | then obtain aaa h2 where mrec_rec: "mrec b h1 = Some (aaa, h2)" by fastforce | 
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
53109diff
changeset | 121 | moreover from mrec_rec have "P b h1 (snd (the (mrec b h1))) (fst (the (mrec b h1)))" | 
| 37772 | 122 | apply (intro 1(2)) | 
| 123 | apply (auto simp add: Inr Inl') | |
| 124 | done | |
| 125 | moreover note mrec mrec_rec exec_f Inl' Inr' 1(1) 1(3) | |
| 126 | ultimately show ?thesis | |
| 127 | apply auto | |
| 128 | apply (rule rec_case) | |
| 129 | apply auto | |
| 39754 | 130 | unfolding MREC_def by (auto simp: mrec.psimps) | 
| 37772 | 131 | next | 
| 132 | case None | |
| 39754 | 133 | from this 1(1) exec_f mrec Inr' 1(3) show ?thesis by (auto simp: mrec.psimps) | 
| 37772 | 134 | qed | 
| 135 | qed | |
| 136 | next | |
| 137 | case None | |
| 39754 | 138 | from this 1(1) mrec 1(3) show ?thesis by (simp add: mrec.psimps) | 
| 37772 | 139 | qed | 
| 140 | qed | |
| 141 | from this h'_r show ?thesis by simp | |
| 142 | qed | |
| 143 | ||
| 144 | end | |
| 145 | ||
| 146 | text {* Providing global versions of the constant and the theorems *}
 | |
| 147 | ||
| 148 | abbreviation "MREC == mrec.MREC" | |
| 149 | lemmas MREC_rule = mrec.MREC_rule | |
| 150 | lemmas MREC_pinduct = mrec.MREC_pinduct | |
| 151 | ||
| 152 | lemma MREC_induct: | |
| 40671 | 153 | assumes "effect (MREC f g x) h h' r" | 
| 154 | assumes "\<And> x h h' r. effect (f x) h h' (Inl r) \<Longrightarrow> P x h h' r" | |
| 155 | assumes "\<And> x h h1 h2 h' s z r. effect (f x) h h1 (Inr s) \<Longrightarrow> effect (MREC f g s) h1 h2 z \<Longrightarrow> P s h1 h2 z | |
| 156 | \<Longrightarrow> effect (g x s z) h2 h' r \<Longrightarrow> P x h h' r" | |
| 37772 | 157 | shows "P x h h' r" | 
| 40671 | 158 | proof (rule MREC_pinduct[OF assms(1) [unfolded effect_def]]) | 
| 37772 | 159 | fix x h h1 h2 h' s z r | 
| 160 | assume "Heap_Monad.execute (f x) h = Some (Inr s, h1)" | |
| 161 | "Heap_Monad.execute (MREC f g s) h1 = Some (z, h2)" | |
| 162 | "P s h1 h2 z" | |
| 163 | "Heap_Monad.execute (g x s z) h2 = Some (r, h')" | |
| 40671 | 164 | from assms(3) [unfolded effect_def, OF this(1) this(2) this(3) this(4)] | 
| 37772 | 165 | show "P x h h' r" . | 
| 166 | next | |
| 40671 | 167 | qed (auto simp add: assms(2)[unfolded effect_def]) | 
| 37772 | 168 | |
| 169 | end |