| author | wenzelm | 
| Fri, 03 Nov 2023 10:13:41 +0100 | |
| changeset 78886 | b82c2f801f2c | 
| parent 67613 | ce654b0e6d69 | 
| permissions | -rw-r--r-- | 
| 42151 | 1 | (* Title: HOL/HOLCF/ex/Loop.thy | 
| 1479 | 2 | Author: Franz Regensburger | 
| 244 | 3 | *) | 
| 4 | ||
| 62175 | 5 | section \<open>Theory for a loop primitive like while\<close> | 
| 244 | 6 | |
| 17291 | 7 | theory Loop | 
| 8 | imports HOLCF | |
| 9 | begin | |
| 244 | 10 | |
| 19763 | 11 | definition | 
| 63549 | 12 |   step  :: "('a \<rightarrow> tr) \<rightarrow> ('a \<rightarrow> 'a) \<rightarrow> 'a \<rightarrow> 'a" where
 | 
| 13 | "step = (LAM b g x. If b\<cdot>x then g\<cdot>x else x)" | |
| 19742 | 14 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
19763diff
changeset | 15 | definition | 
| 63549 | 16 |   while :: "('a \<rightarrow> tr) \<rightarrow> ('a \<rightarrow> 'a) \<rightarrow> 'a \<rightarrow> 'a" where
 | 
| 17 | "while = (LAM b g. fix\<cdot>(LAM f x. If b\<cdot>x then f\<cdot>(g\<cdot>x) else x))" | |
| 244 | 18 | |
| 19742 | 19 | (* ------------------------------------------------------------------------- *) | 
| 20 | (* access to definitions *) | |
| 21 | (* ------------------------------------------------------------------------- *) | |
| 22 | ||
| 23 | ||
| 63549 | 24 | lemma step_def2: "step\<cdot>b\<cdot>g\<cdot>x = If b\<cdot>x then g\<cdot>x else x" | 
| 19742 | 25 | apply (unfold step_def) | 
| 26 | apply simp | |
| 27 | done | |
| 28 | ||
| 63549 | 29 | lemma while_def2: "while\<cdot>b\<cdot>g = fix\<cdot>(LAM f x. If b\<cdot>x then f\<cdot>(g\<cdot>x) else x)" | 
| 19742 | 30 | apply (unfold while_def) | 
| 31 | apply simp | |
| 32 | done | |
| 33 | ||
| 34 | ||
| 35 | (* ------------------------------------------------------------------------- *) | |
| 36 | (* rekursive properties of while *) | |
| 37 | (* ------------------------------------------------------------------------- *) | |
| 38 | ||
| 63549 | 39 | lemma while_unfold: "while\<cdot>b\<cdot>g\<cdot>x = If b\<cdot>x then while\<cdot>b\<cdot>g\<cdot>(g\<cdot>x) else x" | 
| 19742 | 40 | apply (rule trans) | 
| 41 | apply (rule while_def2 [THEN fix_eq5]) | |
| 42 | apply simp | |
| 43 | done | |
| 44 | ||
| 63549 | 45 | lemma while_unfold2: "\<forall>x. while\<cdot>b\<cdot>g\<cdot>x = while\<cdot>b\<cdot>g\<cdot>(iterate k\<cdot>(step\<cdot>b\<cdot>g)\<cdot>x)" | 
| 19742 | 46 | apply (induct_tac k) | 
| 47 | apply simp | |
| 48 | apply (rule allI) | |
| 49 | apply (rule trans) | |
| 40429 | 50 | apply (rule while_unfold) | 
| 19742 | 51 | apply (subst iterate_Suc2) | 
| 52 | apply (rule trans) | |
| 53 | apply (erule_tac [2] spec) | |
| 54 | apply (subst step_def2) | |
| 63549 | 55 | apply (rule_tac p = "b\<cdot>x" in trE) | 
| 19742 | 56 | apply simp | 
| 57 | apply (subst while_unfold) | |
| 63549 | 58 | apply (rule_tac s = "UU" and t = "b\<cdot>UU" in ssubst) | 
| 40429 | 59 | apply (erule strictI) | 
| 19742 | 60 | apply simp | 
| 61 | apply simp | |
| 62 | apply simp | |
| 63 | apply (subst while_unfold) | |
| 64 | apply simp | |
| 65 | done | |
| 66 | ||
| 63549 | 67 | lemma while_unfold3: "while\<cdot>b\<cdot>g\<cdot>x = while\<cdot>b\<cdot>g\<cdot>(step\<cdot>b\<cdot>g\<cdot>x)" | 
| 68 | apply (rule_tac s = "while\<cdot>b\<cdot>g\<cdot>(iterate (Suc 0)\<cdot>(step\<cdot>b\<cdot>g)\<cdot>x)" in trans) | |
| 19742 | 69 | apply (rule while_unfold2 [THEN spec]) | 
| 70 | apply simp | |
| 71 | done | |
| 72 | ||
| 73 | ||
| 74 | (* ------------------------------------------------------------------------- *) | |
| 75 | (* properties of while and iterations *) | |
| 76 | (* ------------------------------------------------------------------------- *) | |
| 77 | ||
| 67613 | 78 | lemma loop_lemma1: "\<lbrakk>\<exists>y. b\<cdot>y = FF; iterate k\<cdot>(step\<cdot>b\<cdot>g)\<cdot>x = UU\<rbrakk> | 
| 63549 | 79 | \<Longrightarrow> iterate(Suc k)\<cdot>(step\<cdot>b\<cdot>g)\<cdot>x = UU" | 
| 19742 | 80 | apply (simp (no_asm)) | 
| 81 | apply (rule trans) | |
| 82 | apply (rule step_def2) | |
| 83 | apply simp | |
| 84 | apply (erule exE) | |
| 85 | apply (erule flat_codom [THEN disjE]) | |
| 86 | apply simp_all | |
| 87 | done | |
| 88 | ||
| 63549 | 89 | lemma loop_lemma2: "\<lbrakk>\<exists>y. b\<cdot>y = FF; iterate (Suc k)\<cdot>(step\<cdot>b\<cdot>g)\<cdot>x \<noteq> UU\<rbrakk> \<Longrightarrow> | 
| 90 | iterate k\<cdot>(step\<cdot>b\<cdot>g)\<cdot>x \<noteq> UU" | |
| 19742 | 91 | apply (blast intro: loop_lemma1) | 
| 92 | done | |
| 93 | ||
| 94 | lemma loop_lemma3 [rule_format (no_asm)]: | |
| 63549 | 95 | "\<lbrakk>\<forall>x. INV x \<and> b\<cdot>x = TT \<and> g\<cdot>x \<noteq> UU \<longrightarrow> INV (g\<cdot>x); | 
| 96 | \<exists>y. b\<cdot>y = FF; INV x\<rbrakk> | |
| 97 | \<Longrightarrow> iterate k\<cdot>(step\<cdot>b\<cdot>g)\<cdot>x \<noteq> UU \<longrightarrow> INV (iterate k\<cdot>(step\<cdot>b\<cdot>g)\<cdot>x)" | |
| 19742 | 98 | apply (induct_tac "k") | 
| 99 | apply (simp (no_asm_simp)) | |
| 100 | apply (intro strip) | |
| 101 | apply (simp (no_asm) add: step_def2) | |
| 63549 | 102 | apply (rule_tac p = "b\<cdot>(iterate n\<cdot>(step\<cdot>b\<cdot>g)\<cdot>x)" in trE) | 
| 19742 | 103 | apply (erule notE) | 
| 40028 | 104 | apply (simp add: step_def2) | 
| 105 | apply (simp (no_asm_simp)) | |
| 19742 | 106 | apply (rule mp) | 
| 107 | apply (erule spec) | |
| 40028 | 108 | apply (simp (no_asm_simp) del: iterate_Suc add: loop_lemma2) | 
| 63549 | 109 | apply (rule_tac s = "iterate (Suc n)\<cdot>(step\<cdot>b\<cdot>g)\<cdot>x" | 
| 110 | and t = "g\<cdot>(iterate n\<cdot>(step\<cdot>b\<cdot>g)\<cdot>x)" in ssubst) | |
| 19742 | 111 | prefer 2 apply (assumption) | 
| 112 | apply (simp add: step_def2) | |
| 35170 | 113 | apply (drule (1) loop_lemma2, simp) | 
| 19742 | 114 | done | 
| 115 | ||
| 116 | lemma loop_lemma4 [rule_format]: | |
| 63549 | 117 | "\<forall>x. b\<cdot>(iterate k\<cdot>(step\<cdot>b\<cdot>g)\<cdot>x) = FF \<longrightarrow> while\<cdot>b\<cdot>g\<cdot>x = iterate k\<cdot>(step\<cdot>b\<cdot>g)\<cdot>x" | 
| 19742 | 118 | apply (induct_tac k) | 
| 119 | apply (simp (no_asm)) | |
| 120 | apply (intro strip) | |
| 121 | apply (simplesubst while_unfold) | |
| 122 | apply simp | |
| 123 | apply (rule allI) | |
| 124 | apply (simplesubst iterate_Suc2) | |
| 125 | apply (intro strip) | |
| 126 | apply (rule trans) | |
| 127 | apply (rule while_unfold3) | |
| 128 | apply simp | |
| 129 | done | |
| 130 | ||
| 131 | lemma loop_lemma5 [rule_format (no_asm)]: | |
| 63549 | 132 | "\<forall>k. b\<cdot>(iterate k\<cdot>(step\<cdot>b\<cdot>g)\<cdot>x) \<noteq> FF \<Longrightarrow> | 
| 133 | \<forall>m. while\<cdot>b\<cdot>g\<cdot>(iterate m\<cdot>(step\<cdot>b\<cdot>g)\<cdot>x) = UU" | |
| 19742 | 134 | apply (simplesubst while_def2) | 
| 135 | apply (rule fix_ind) | |
| 35948 | 136 | apply simp | 
| 137 | apply simp | |
| 19742 | 138 | apply (rule allI) | 
| 139 | apply (simp (no_asm)) | |
| 63549 | 140 | apply (rule_tac p = "b\<cdot>(iterate m\<cdot>(step\<cdot>b\<cdot>g)\<cdot>x)" in trE) | 
| 19742 | 141 | apply (simp (no_asm_simp)) | 
| 142 | apply (simp (no_asm_simp)) | |
| 63549 | 143 | apply (rule_tac s = "xa\<cdot>(iterate (Suc m)\<cdot>(step\<cdot>b\<cdot>g)\<cdot>x)" in trans) | 
| 19742 | 144 | apply (erule_tac [2] spec) | 
| 145 | apply (rule cfun_arg_cong) | |
| 146 | apply (rule trans) | |
| 147 | apply (rule_tac [2] iterate_Suc [symmetric]) | |
| 148 | apply (simp add: step_def2) | |
| 149 | apply blast | |
| 150 | done | |
| 151 | ||
| 63549 | 152 | lemma loop_lemma6: "\<forall>k. b\<cdot>(iterate k\<cdot>(step\<cdot>b\<cdot>g)\<cdot>x) \<noteq> FF \<Longrightarrow> while\<cdot>b\<cdot>g\<cdot>x = UU" | 
| 19742 | 153 | apply (rule_tac t = "x" in iterate_0 [THEN subst]) | 
| 154 | apply (erule loop_lemma5) | |
| 155 | done | |
| 156 | ||
| 63549 | 157 | lemma loop_lemma7: "while\<cdot>b\<cdot>g\<cdot>x \<noteq> UU \<Longrightarrow> \<exists>k. b\<cdot>(iterate k\<cdot>(step\<cdot>b\<cdot>g)\<cdot>x) = FF" | 
| 19742 | 158 | apply (blast intro: loop_lemma6) | 
| 159 | done | |
| 160 | ||
| 161 | ||
| 162 | (* ------------------------------------------------------------------------- *) | |
| 163 | (* an invariant rule for loops *) | |
| 164 | (* ------------------------------------------------------------------------- *) | |
| 165 | ||
| 166 | lemma loop_inv2: | |
| 63549 | 167 | "\<lbrakk>(\<forall>y. INV y \<and> b\<cdot>y = TT \<and> g\<cdot>y \<noteq> UU \<longrightarrow> INV (g\<cdot>y)); | 
| 168 | (\<forall>y. INV y \<and> b\<cdot>y = FF \<longrightarrow> Q y); | |
| 169 | INV x; while\<cdot>b\<cdot>g\<cdot>x \<noteq> UU\<rbrakk> \<Longrightarrow> Q (while\<cdot>b\<cdot>g\<cdot>x)" | |
| 170 | apply (rule_tac P = "\<lambda>k. b\<cdot>(iterate k\<cdot>(step\<cdot>b\<cdot>g)\<cdot>x) = FF" in exE) | |
| 19742 | 171 | apply (erule loop_lemma7) | 
| 172 | apply (simplesubst loop_lemma4) | |
| 173 | apply assumption | |
| 174 | apply (drule spec, erule mp) | |
| 175 | apply (rule conjI) | |
| 176 | prefer 2 apply (assumption) | |
| 177 | apply (rule loop_lemma3) | |
| 178 | apply assumption | |
| 179 | apply (blast intro: loop_lemma6) | |
| 180 | apply assumption | |
| 181 | apply (rotate_tac -1) | |
| 182 | apply (simp add: loop_lemma4) | |
| 183 | done | |
| 184 | ||
| 185 | lemma loop_inv: | |
| 186 | assumes premP: "P(x)" | |
| 63549 | 187 | and premI: "\<And>y. P y \<Longrightarrow> INV y" | 
| 188 | and premTT: "\<And>y. \<lbrakk>INV y; b\<cdot>y = TT; g\<cdot>y \<noteq> UU\<rbrakk> \<Longrightarrow> INV (g\<cdot>y)" | |
| 189 | and premFF: "\<And>y. \<lbrakk>INV y; b\<cdot>y = FF\<rbrakk> \<Longrightarrow> Q y" | |
| 190 | and premW: "while\<cdot>b\<cdot>g\<cdot>x \<noteq> UU" | |
| 191 | shows "Q (while\<cdot>b\<cdot>g\<cdot>x)" | |
| 19742 | 192 | apply (rule loop_inv2) | 
| 193 | apply (rule_tac [3] premP [THEN premI]) | |
| 194 | apply (rule_tac [3] premW) | |
| 195 | apply (blast intro: premTT) | |
| 196 | apply (blast intro: premFF) | |
| 197 | done | |
| 244 | 198 | |
| 199 | end |