| author | wenzelm | 
| Sat, 16 Aug 2014 18:08:55 +0200 | |
| changeset 57956 | 3ab5d15fac6b | 
| parent 57437 | 0baf08c075b9 | 
| child 58249 | 180f1b3508ed | 
| permissions | -rw-r--r-- | 
| 
37787
 
30dc3abf4a58
theorem collections do not contain default rules any longer
 
haftmann 
parents: 
37772 
diff
changeset
 | 
1  | 
(* Title: HOL/Imperative_HOL/Heap_Monad.thy  | 
| 26170 | 2  | 
Author: John Matthews, Galois Connections; Alexander Krauss, Lukas Bulwahn & Florian Haftmann, TU Muenchen  | 
3  | 
*)  | 
|
4  | 
||
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
5  | 
header {* A monad with a polymorphic heap and primitive reasoning infrastructure *}
 | 
| 26170 | 6  | 
|
7  | 
theory Heap_Monad  | 
|
| 
41413
 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 
wenzelm 
parents: 
40671 
diff
changeset
 | 
8  | 
imports  | 
| 
 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 
wenzelm 
parents: 
40671 
diff
changeset
 | 
9  | 
Heap  | 
| 
 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 
wenzelm 
parents: 
40671 
diff
changeset
 | 
10  | 
"~~/src/HOL/Library/Monad_Syntax"  | 
| 26170 | 11  | 
begin  | 
12  | 
||
13  | 
subsection {* The monad *}
 | 
|
14  | 
||
| 37758 | 15  | 
subsubsection {* Monad construction *}
 | 
| 26170 | 16  | 
|
17  | 
text {* Monadic heap actions either produce values
 | 
|
18  | 
and transform the heap, or fail *}  | 
|
| 37709 | 19  | 
datatype 'a Heap = Heap "heap \<Rightarrow> ('a \<times> heap) option"
 | 
| 26170 | 20  | 
|
| 
40266
 
d72f1f734e5a
remove term_of equations for Heap type explicitly
 
haftmann 
parents: 
40173 
diff
changeset
 | 
21  | 
lemma [code, code del]:  | 
| 
 
d72f1f734e5a
remove term_of equations for Heap type explicitly
 
haftmann 
parents: 
40173 
diff
changeset
 | 
22  | 
"(Code_Evaluation.term_of :: 'a::typerep Heap \<Rightarrow> Code_Evaluation.term) = Code_Evaluation.term_of"  | 
| 
 
d72f1f734e5a
remove term_of equations for Heap type explicitly
 
haftmann 
parents: 
40173 
diff
changeset
 | 
23  | 
..  | 
| 
 
d72f1f734e5a
remove term_of equations for Heap type explicitly
 
haftmann 
parents: 
40173 
diff
changeset
 | 
24  | 
|
| 37709 | 25  | 
primrec execute :: "'a Heap \<Rightarrow> heap \<Rightarrow> ('a \<times> heap) option" where
 | 
26  | 
[code del]: "execute (Heap f) = f"  | 
|
| 26170 | 27  | 
|
| 37758 | 28  | 
lemma Heap_cases [case_names succeed fail]:  | 
29  | 
fixes f and h  | 
|
30  | 
assumes succeed: "\<And>x h'. execute f h = Some (x, h') \<Longrightarrow> P"  | 
|
31  | 
assumes fail: "execute f h = None \<Longrightarrow> P"  | 
|
32  | 
shows P  | 
|
33  | 
using assms by (cases "execute f h") auto  | 
|
34  | 
||
| 26170 | 35  | 
lemma Heap_execute [simp]:  | 
36  | 
"Heap (execute f) = f" by (cases f) simp_all  | 
|
37  | 
||
38  | 
lemma Heap_eqI:  | 
|
39  | 
"(\<And>h. execute f h = execute g h) \<Longrightarrow> f = g"  | 
|
| 
39302
 
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
 
nipkow 
parents: 
39250 
diff
changeset
 | 
40  | 
by (cases f, cases g) (auto simp: fun_eq_iff)  | 
| 26170 | 41  | 
|
| 57956 | 42  | 
named_theorems execute_simps "simplification rules for execute"  | 
| 37758 | 43  | 
|
| 
37787
 
30dc3abf4a58
theorem collections do not contain default rules any longer
 
haftmann 
parents: 
37772 
diff
changeset
 | 
44  | 
lemma execute_Let [execute_simps]:  | 
| 37758 | 45  | 
"execute (let x = t in f x) = (let x = t in execute (f x))"  | 
46  | 
by (simp add: Let_def)  | 
|
47  | 
||
48  | 
||
49  | 
subsubsection {* Specialised lifters *}
 | 
|
50  | 
||
51  | 
definition tap :: "(heap \<Rightarrow> 'a) \<Rightarrow> 'a Heap" where  | 
|
52  | 
[code del]: "tap f = Heap (\<lambda>h. Some (f h, h))"  | 
|
53  | 
||
| 
37787
 
30dc3abf4a58
theorem collections do not contain default rules any longer
 
haftmann 
parents: 
37772 
diff
changeset
 | 
54  | 
lemma execute_tap [execute_simps]:  | 
| 37758 | 55  | 
"execute (tap f) h = Some (f h, h)"  | 
56  | 
by (simp add: tap_def)  | 
|
| 26170 | 57  | 
|
| 37709 | 58  | 
definition heap :: "(heap \<Rightarrow> 'a \<times> heap) \<Rightarrow> 'a Heap" where  | 
59  | 
[code del]: "heap f = Heap (Some \<circ> f)"  | 
|
| 26170 | 60  | 
|
| 
37787
 
30dc3abf4a58
theorem collections do not contain default rules any longer
 
haftmann 
parents: 
37772 
diff
changeset
 | 
61  | 
lemma execute_heap [execute_simps]:  | 
| 37709 | 62  | 
"execute (heap f) = Some \<circ> f"  | 
| 26170 | 63  | 
by (simp add: heap_def)  | 
64  | 
||
| 37754 | 65  | 
definition guard :: "(heap \<Rightarrow> bool) \<Rightarrow> (heap \<Rightarrow> 'a \<times> heap) \<Rightarrow> 'a Heap" where  | 
66  | 
[code del]: "guard P f = Heap (\<lambda>h. if P h then Some (f h) else None)"  | 
|
67  | 
||
| 37758 | 68  | 
lemma execute_guard [execute_simps]:  | 
| 37754 | 69  | 
"\<not> P h \<Longrightarrow> execute (guard P f) h = None"  | 
70  | 
"P h \<Longrightarrow> execute (guard P f) h = Some (f h)"  | 
|
71  | 
by (simp_all add: guard_def)  | 
|
72  | 
||
| 37758 | 73  | 
|
74  | 
subsubsection {* Predicate classifying successful computations *}
 | 
|
75  | 
||
76  | 
definition success :: "'a Heap \<Rightarrow> heap \<Rightarrow> bool" where  | 
|
77  | 
"success f h \<longleftrightarrow> execute f h \<noteq> None"  | 
|
78  | 
||
79  | 
lemma successI:  | 
|
80  | 
"execute f h \<noteq> None \<Longrightarrow> success f h"  | 
|
81  | 
by (simp add: success_def)  | 
|
82  | 
||
83  | 
lemma successE:  | 
|
84  | 
assumes "success f h"  | 
|
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
85  | 
obtains r h' where "r = fst (the (execute c h))"  | 
| 
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
86  | 
and "h' = snd (the (execute c h))"  | 
| 
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
87  | 
and "execute f h \<noteq> None"  | 
| 
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
88  | 
using assms by (simp add: success_def)  | 
| 37758 | 89  | 
|
| 57956 | 90  | 
named_theorems success_intros "introduction rules for success"  | 
| 37758 | 91  | 
|
| 
37787
 
30dc3abf4a58
theorem collections do not contain default rules any longer
 
haftmann 
parents: 
37772 
diff
changeset
 | 
92  | 
lemma success_tapI [success_intros]:  | 
| 37758 | 93  | 
"success (tap f) h"  | 
| 
37787
 
30dc3abf4a58
theorem collections do not contain default rules any longer
 
haftmann 
parents: 
37772 
diff
changeset
 | 
94  | 
by (rule successI) (simp add: execute_simps)  | 
| 37758 | 95  | 
|
| 
37787
 
30dc3abf4a58
theorem collections do not contain default rules any longer
 
haftmann 
parents: 
37772 
diff
changeset
 | 
96  | 
lemma success_heapI [success_intros]:  | 
| 37758 | 97  | 
"success (heap f) h"  | 
| 
37787
 
30dc3abf4a58
theorem collections do not contain default rules any longer
 
haftmann 
parents: 
37772 
diff
changeset
 | 
98  | 
by (rule successI) (simp add: execute_simps)  | 
| 37758 | 99  | 
|
100  | 
lemma success_guardI [success_intros]:  | 
|
101  | 
"P h \<Longrightarrow> success (guard P f) h"  | 
|
102  | 
by (rule successI) (simp add: execute_guard)  | 
|
103  | 
||
104  | 
lemma success_LetI [success_intros]:  | 
|
105  | 
"x = t \<Longrightarrow> success (f x) h \<Longrightarrow> success (let x = t in f x) h"  | 
|
106  | 
by (simp add: Let_def)  | 
|
107  | 
||
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
108  | 
lemma success_ifI:  | 
| 
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
109  | 
"(c \<Longrightarrow> success t h) \<Longrightarrow> (\<not> c \<Longrightarrow> success e h) \<Longrightarrow>  | 
| 
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
110  | 
success (if c then t else e) h"  | 
| 
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
111  | 
by (simp add: success_def)  | 
| 
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
112  | 
|
| 
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
113  | 
|
| 
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
114  | 
subsubsection {* Predicate for a simple relational calculus *}
 | 
| 
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
115  | 
|
| 
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
116  | 
text {*
 | 
| 40671 | 117  | 
  The @{text effect} predicate states that when a computation @{text c}
 | 
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
118  | 
  runs with the heap @{text h} will result in return value @{text r}
 | 
| 
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
119  | 
  and a heap @{text "h'"}, i.e.~no exception occurs.
 | 
| 
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
120  | 
*}  | 
| 
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
121  | 
|
| 40671 | 122  | 
definition effect :: "'a Heap \<Rightarrow> heap \<Rightarrow> heap \<Rightarrow> 'a \<Rightarrow> bool" where  | 
123  | 
effect_def: "effect c h h' r \<longleftrightarrow> execute c h = Some (r, h')"  | 
|
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
124  | 
|
| 40671 | 125  | 
lemma effectI:  | 
126  | 
"execute c h = Some (r, h') \<Longrightarrow> effect c h h' r"  | 
|
127  | 
by (simp add: effect_def)  | 
|
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
128  | 
|
| 40671 | 129  | 
lemma effectE:  | 
130  | 
assumes "effect c h h' r"  | 
|
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
131  | 
obtains "r = fst (the (execute c h))"  | 
| 
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
132  | 
and "h' = snd (the (execute c h))"  | 
| 
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
133  | 
and "success c h"  | 
| 
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
134  | 
proof (rule that)  | 
| 40671 | 135  | 
from assms have *: "execute c h = Some (r, h')" by (simp add: effect_def)  | 
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
136  | 
then show "success c h" by (simp add: success_def)  | 
| 
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
137  | 
from * have "fst (the (execute c h)) = r" and "snd (the (execute c h)) = h'"  | 
| 
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
138  | 
by simp_all  | 
| 
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
139  | 
then show "r = fst (the (execute c h))"  | 
| 
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
140  | 
and "h' = snd (the (execute c h))" by simp_all  | 
| 
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
141  | 
qed  | 
| 
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
142  | 
|
| 40671 | 143  | 
lemma effect_success:  | 
144  | 
"effect c h h' r \<Longrightarrow> success c h"  | 
|
145  | 
by (simp add: effect_def success_def)  | 
|
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
146  | 
|
| 40671 | 147  | 
lemma success_effectE:  | 
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
148  | 
assumes "success c h"  | 
| 40671 | 149  | 
obtains r h' where "effect c h h' r"  | 
150  | 
using assms by (auto simp add: effect_def success_def)  | 
|
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
151  | 
|
| 40671 | 152  | 
lemma effect_deterministic:  | 
153  | 
assumes "effect f h h' a"  | 
|
154  | 
and "effect f h h'' b"  | 
|
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
155  | 
shows "a = b" and "h' = h''"  | 
| 40671 | 156  | 
using assms unfolding effect_def by auto  | 
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
157  | 
|
| 57956 | 158  | 
named_theorems effect_intros "introduction rules for effect"  | 
159  | 
named_theorems effect_elims "elimination rules for effect"  | 
|
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
160  | 
|
| 40671 | 161  | 
lemma effect_LetI [effect_intros]:  | 
162  | 
assumes "x = t" "effect (f x) h h' r"  | 
|
163  | 
shows "effect (let x = t in f x) h h' r"  | 
|
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
164  | 
using assms by simp  | 
| 
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
165  | 
|
| 40671 | 166  | 
lemma effect_LetE [effect_elims]:  | 
167  | 
assumes "effect (let x = t in f x) h h' r"  | 
|
168  | 
obtains "effect (f t) h h' r"  | 
|
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
169  | 
using assms by simp  | 
| 
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
170  | 
|
| 40671 | 171  | 
lemma effect_ifI:  | 
172  | 
assumes "c \<Longrightarrow> effect t h h' r"  | 
|
173  | 
and "\<not> c \<Longrightarrow> effect e h h' r"  | 
|
174  | 
shows "effect (if c then t else e) h h' r"  | 
|
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
175  | 
by (cases c) (simp_all add: assms)  | 
| 
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
176  | 
|
| 40671 | 177  | 
lemma effect_ifE:  | 
178  | 
assumes "effect (if c then t else e) h h' r"  | 
|
179  | 
obtains "c" "effect t h h' r"  | 
|
180  | 
| "\<not> c" "effect e h h' r"  | 
|
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
181  | 
using assms by (cases c) simp_all  | 
| 
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
182  | 
|
| 40671 | 183  | 
lemma effect_tapI [effect_intros]:  | 
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
184  | 
assumes "h' = h" "r = f h"  | 
| 40671 | 185  | 
shows "effect (tap f) h h' r"  | 
186  | 
by (rule effectI) (simp add: assms execute_simps)  | 
|
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
187  | 
|
| 40671 | 188  | 
lemma effect_tapE [effect_elims]:  | 
189  | 
assumes "effect (tap f) h h' r"  | 
|
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
190  | 
obtains "h' = h" and "r = f h"  | 
| 40671 | 191  | 
using assms by (rule effectE) (auto simp add: execute_simps)  | 
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
192  | 
|
| 40671 | 193  | 
lemma effect_heapI [effect_intros]:  | 
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
194  | 
assumes "h' = snd (f h)" "r = fst (f h)"  | 
| 40671 | 195  | 
shows "effect (heap f) h h' r"  | 
196  | 
by (rule effectI) (simp add: assms execute_simps)  | 
|
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
197  | 
|
| 40671 | 198  | 
lemma effect_heapE [effect_elims]:  | 
199  | 
assumes "effect (heap f) h h' r"  | 
|
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
200  | 
obtains "h' = snd (f h)" and "r = fst (f h)"  | 
| 40671 | 201  | 
using assms by (rule effectE) (simp add: execute_simps)  | 
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
202  | 
|
| 40671 | 203  | 
lemma effect_guardI [effect_intros]:  | 
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
204  | 
assumes "P h" "h' = snd (f h)" "r = fst (f h)"  | 
| 40671 | 205  | 
shows "effect (guard P f) h h' r"  | 
206  | 
by (rule effectI) (simp add: assms execute_simps)  | 
|
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
207  | 
|
| 40671 | 208  | 
lemma effect_guardE [effect_elims]:  | 
209  | 
assumes "effect (guard P f) h h' r"  | 
|
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
210  | 
obtains "h' = snd (f h)" "r = fst (f h)" "P h"  | 
| 40671 | 211  | 
using assms by (rule effectE)  | 
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
212  | 
(auto simp add: execute_simps elim!: successE, cases "P h", auto simp add: execute_simps)  | 
| 
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
213  | 
|
| 37758 | 214  | 
|
215  | 
subsubsection {* Monad combinators *}
 | 
|
| 26170 | 216  | 
|
| 37709 | 217  | 
definition return :: "'a \<Rightarrow> 'a Heap" where  | 
| 26170 | 218  | 
[code del]: "return x = heap (Pair x)"  | 
219  | 
||
| 
37787
 
30dc3abf4a58
theorem collections do not contain default rules any longer
 
haftmann 
parents: 
37772 
diff
changeset
 | 
220  | 
lemma execute_return [execute_simps]:  | 
| 37709 | 221  | 
"execute (return x) = Some \<circ> Pair x"  | 
| 
37787
 
30dc3abf4a58
theorem collections do not contain default rules any longer
 
haftmann 
parents: 
37772 
diff
changeset
 | 
222  | 
by (simp add: return_def execute_simps)  | 
| 26170 | 223  | 
|
| 
37787
 
30dc3abf4a58
theorem collections do not contain default rules any longer
 
haftmann 
parents: 
37772 
diff
changeset
 | 
224  | 
lemma success_returnI [success_intros]:  | 
| 37758 | 225  | 
"success (return x) h"  | 
| 
37787
 
30dc3abf4a58
theorem collections do not contain default rules any longer
 
haftmann 
parents: 
37772 
diff
changeset
 | 
226  | 
by (rule successI) (simp add: execute_simps)  | 
| 37758 | 227  | 
|
| 40671 | 228  | 
lemma effect_returnI [effect_intros]:  | 
229  | 
"h = h' \<Longrightarrow> effect (return x) h h' x"  | 
|
230  | 
by (rule effectI) (simp add: execute_simps)  | 
|
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
231  | 
|
| 40671 | 232  | 
lemma effect_returnE [effect_elims]:  | 
233  | 
assumes "effect (return x) h h' r"  | 
|
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
234  | 
obtains "r = x" "h' = h"  | 
| 40671 | 235  | 
using assms by (rule effectE) (simp add: execute_simps)  | 
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
236  | 
|
| 37709 | 237  | 
definition raise :: "string \<Rightarrow> 'a Heap" where -- {* the string is just decoration *}
 | 
238  | 
[code del]: "raise s = Heap (\<lambda>_. None)"  | 
|
| 26170 | 239  | 
|
| 
37787
 
30dc3abf4a58
theorem collections do not contain default rules any longer
 
haftmann 
parents: 
37772 
diff
changeset
 | 
240  | 
lemma execute_raise [execute_simps]:  | 
| 37709 | 241  | 
"execute (raise s) = (\<lambda>_. None)"  | 
| 26170 | 242  | 
by (simp add: raise_def)  | 
243  | 
||
| 40671 | 244  | 
lemma effect_raiseE [effect_elims]:  | 
245  | 
assumes "effect (raise x) h h' r"  | 
|
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
246  | 
obtains "False"  | 
| 40671 | 247  | 
using assms by (rule effectE) (simp add: success_def execute_simps)  | 
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
248  | 
|
| 37792 | 249  | 
definition bind :: "'a Heap \<Rightarrow> ('a \<Rightarrow> 'b Heap) \<Rightarrow> 'b Heap" where
 | 
250  | 
[code del]: "bind f g = Heap (\<lambda>h. case execute f h of  | 
|
| 37709 | 251  | 
Some (x, h') \<Rightarrow> execute (g x) h'  | 
252  | 
| None \<Rightarrow> None)"  | 
|
253  | 
||
| 
52622
 
e0ff1625e96d
localized and modernized adhoc-overloading (patch by Christian Sternagel);
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
254  | 
adhoc_overloading  | 
| 
 
e0ff1625e96d
localized and modernized adhoc-overloading (patch by Christian Sternagel);
 
wenzelm 
parents: 
52435 
diff
changeset
 | 
255  | 
Monad_Syntax.bind Heap_Monad.bind  | 
| 37792 | 256  | 
|
| 37758 | 257  | 
lemma execute_bind [execute_simps]:  | 
| 37709 | 258  | 
"execute f h = Some (x, h') \<Longrightarrow> execute (f \<guillemotright>= g) h = execute (g x) h'"  | 
259  | 
"execute f h = None \<Longrightarrow> execute (f \<guillemotright>= g) h = None"  | 
|
| 
37756
 
59caa6180fff
avoid slightly odd "M" suffix; rename mapM to fold_map (fold_map_abort would be more correct, though)
 
haftmann 
parents: 
37754 
diff
changeset
 | 
260  | 
by (simp_all add: bind_def)  | 
| 37709 | 261  | 
|
| 38409 | 262  | 
lemma execute_bind_case:  | 
263  | 
"execute (f \<guillemotright>= g) h = (case (execute f h) of  | 
|
264  | 
Some (x, h') \<Rightarrow> execute (g x) h' | None \<Rightarrow> None)"  | 
|
265  | 
by (simp add: bind_def)  | 
|
266  | 
||
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
267  | 
lemma execute_bind_success:  | 
| 
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
268  | 
"success f h \<Longrightarrow> execute (f \<guillemotright>= g) h = execute (g (fst (the (execute f h)))) (snd (the (execute f h)))"  | 
| 
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
269  | 
by (cases f h rule: Heap_cases) (auto elim!: successE simp add: bind_def)  | 
| 
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
270  | 
|
| 
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
271  | 
lemma success_bind_executeI:  | 
| 
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
272  | 
"execute f h = Some (x, h') \<Longrightarrow> success (g x) h' \<Longrightarrow> success (f \<guillemotright>= g) h"  | 
| 37758 | 273  | 
by (auto intro!: successI elim!: successE simp add: bind_def)  | 
274  | 
||
| 40671 | 275  | 
lemma success_bind_effectI [success_intros]:  | 
276  | 
"effect f h h' x \<Longrightarrow> success (g x) h' \<Longrightarrow> success (f \<guillemotright>= g) h"  | 
|
277  | 
by (auto simp add: effect_def success_def bind_def)  | 
|
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
278  | 
|
| 40671 | 279  | 
lemma effect_bindI [effect_intros]:  | 
280  | 
assumes "effect f h h' r" "effect (g r) h' h'' r'"  | 
|
281  | 
shows "effect (f \<guillemotright>= g) h h'' r'"  | 
|
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
282  | 
using assms  | 
| 40671 | 283  | 
apply (auto intro!: effectI elim!: effectE successE)  | 
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
284  | 
apply (subst execute_bind, simp_all)  | 
| 
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
285  | 
done  | 
| 
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
286  | 
|
| 40671 | 287  | 
lemma effect_bindE [effect_elims]:  | 
288  | 
assumes "effect (f \<guillemotright>= g) h h'' r'"  | 
|
289  | 
obtains h' r where "effect f h h' r" "effect (g r) h' h'' r'"  | 
|
290  | 
using assms by (auto simp add: effect_def bind_def split: option.split_asm)  | 
|
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
291  | 
|
| 
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
292  | 
lemma execute_bind_eq_SomeI:  | 
| 37878 | 293  | 
assumes "execute f h = Some (x, h')"  | 
294  | 
and "execute (g x) h' = Some (y, h'')"  | 
|
295  | 
shows "execute (f \<guillemotright>= g) h = Some (y, h'')"  | 
|
| 
37756
 
59caa6180fff
avoid slightly odd "M" suffix; rename mapM to fold_map (fold_map_abort would be more correct, though)
 
haftmann 
parents: 
37754 
diff
changeset
 | 
296  | 
using assms by (simp add: bind_def)  | 
| 37754 | 297  | 
|
| 37709 | 298  | 
lemma return_bind [simp]: "return x \<guillemotright>= f = f x"  | 
| 
51485
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
299  | 
by (rule Heap_eqI) (simp add: execute_simps)  | 
| 37709 | 300  | 
|
301  | 
lemma bind_return [simp]: "f \<guillemotright>= return = f"  | 
|
| 
37787
 
30dc3abf4a58
theorem collections do not contain default rules any longer
 
haftmann 
parents: 
37772 
diff
changeset
 | 
302  | 
by (rule Heap_eqI) (simp add: bind_def execute_simps split: option.splits)  | 
| 37709 | 303  | 
|
| 37828 | 304  | 
lemma bind_bind [simp]: "(f \<guillemotright>= g) \<guillemotright>= k = (f :: 'a Heap) \<guillemotright>= (\<lambda>x. g x \<guillemotright>= k)"  | 
| 
37787
 
30dc3abf4a58
theorem collections do not contain default rules any longer
 
haftmann 
parents: 
37772 
diff
changeset
 | 
305  | 
by (rule Heap_eqI) (simp add: bind_def execute_simps split: option.splits)  | 
| 37709 | 306  | 
|
307  | 
lemma raise_bind [simp]: "raise e \<guillemotright>= f = raise e"  | 
|
| 
37787
 
30dc3abf4a58
theorem collections do not contain default rules any longer
 
haftmann 
parents: 
37772 
diff
changeset
 | 
308  | 
by (rule Heap_eqI) (simp add: execute_simps)  | 
| 37709 | 309  | 
|
| 26170 | 310  | 
|
| 37758 | 311  | 
subsection {* Generic combinators *}
 | 
| 26170 | 312  | 
|
| 37758 | 313  | 
subsubsection {* Assertions *}
 | 
| 26170 | 314  | 
|
| 37709 | 315  | 
definition assert :: "('a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a Heap" where
 | 
316  | 
"assert P x = (if P x then return x else raise ''assert'')"  | 
|
| 28742 | 317  | 
|
| 37758 | 318  | 
lemma execute_assert [execute_simps]:  | 
| 37754 | 319  | 
"P x \<Longrightarrow> execute (assert P x) h = Some (x, h)"  | 
320  | 
"\<not> P x \<Longrightarrow> execute (assert P x) h = None"  | 
|
| 
37787
 
30dc3abf4a58
theorem collections do not contain default rules any longer
 
haftmann 
parents: 
37772 
diff
changeset
 | 
321  | 
by (simp_all add: assert_def execute_simps)  | 
| 37754 | 322  | 
|
| 37758 | 323  | 
lemma success_assertI [success_intros]:  | 
324  | 
"P x \<Longrightarrow> success (assert P x) h"  | 
|
325  | 
by (rule successI) (simp add: execute_assert)  | 
|
326  | 
||
| 40671 | 327  | 
lemma effect_assertI [effect_intros]:  | 
328  | 
"P x \<Longrightarrow> h' = h \<Longrightarrow> r = x \<Longrightarrow> effect (assert P x) h h' r"  | 
|
329  | 
by (rule effectI) (simp add: execute_assert)  | 
|
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
330  | 
|
| 40671 | 331  | 
lemma effect_assertE [effect_elims]:  | 
332  | 
assumes "effect (assert P x) h h' r"  | 
|
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
333  | 
obtains "P x" "r = x" "h' = h"  | 
| 40671 | 334  | 
using assms by (rule effectE) (cases "P x", simp_all add: execute_assert success_def)  | 
| 
37771
 
1bec64044b5e
spelt out relational framework in a consistent way
 
haftmann 
parents: 
37758 
diff
changeset
 | 
335  | 
|
| 28742 | 336  | 
lemma assert_cong [fundef_cong]:  | 
337  | 
assumes "P = P'"  | 
|
338  | 
assumes "\<And>x. P' x \<Longrightarrow> f x = f' x"  | 
|
339  | 
shows "(assert P x >>= f) = (assert P' x >>= f')"  | 
|
| 37754 | 340  | 
by (rule Heap_eqI) (insert assms, simp add: assert_def)  | 
| 28742 | 341  | 
|
| 37758 | 342  | 
|
343  | 
subsubsection {* Plain lifting *}
 | 
|
344  | 
||
| 37754 | 345  | 
definition lift :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a \<Rightarrow> 'b Heap" where
 | 
346  | 
"lift f = return o f"  | 
|
| 37709 | 347  | 
|
| 37754 | 348  | 
lemma lift_collapse [simp]:  | 
349  | 
"lift f x = return (f x)"  | 
|
350  | 
by (simp add: lift_def)  | 
|
| 37709 | 351  | 
|
| 37754 | 352  | 
lemma bind_lift:  | 
353  | 
"(f \<guillemotright>= lift g) = (f \<guillemotright>= (\<lambda>x. return (g x)))"  | 
|
354  | 
by (simp add: lift_def comp_def)  | 
|
| 37709 | 355  | 
|
| 37758 | 356  | 
|
357  | 
subsubsection {* Iteration -- warning: this is rarely useful! *}
 | 
|
358  | 
||
| 
37756
 
59caa6180fff
avoid slightly odd "M" suffix; rename mapM to fold_map (fold_map_abort would be more correct, though)
 
haftmann 
parents: 
37754 
diff
changeset
 | 
359  | 
primrec fold_map :: "('a \<Rightarrow> 'b Heap) \<Rightarrow> 'a list \<Rightarrow> 'b list Heap" where
 | 
| 
 
59caa6180fff
avoid slightly odd "M" suffix; rename mapM to fold_map (fold_map_abort would be more correct, though)
 
haftmann 
parents: 
37754 
diff
changeset
 | 
360  | 
"fold_map f [] = return []"  | 
| 37792 | 361  | 
| "fold_map f (x # xs) = do {
 | 
| 37709 | 362  | 
y \<leftarrow> f x;  | 
| 
37756
 
59caa6180fff
avoid slightly odd "M" suffix; rename mapM to fold_map (fold_map_abort would be more correct, though)
 
haftmann 
parents: 
37754 
diff
changeset
 | 
363  | 
ys \<leftarrow> fold_map f xs;  | 
| 37709 | 364  | 
return (y # ys)  | 
| 37792 | 365  | 
}"  | 
| 37709 | 366  | 
|
| 
37756
 
59caa6180fff
avoid slightly odd "M" suffix; rename mapM to fold_map (fold_map_abort would be more correct, though)
 
haftmann 
parents: 
37754 
diff
changeset
 | 
367  | 
lemma fold_map_append:  | 
| 
 
59caa6180fff
avoid slightly odd "M" suffix; rename mapM to fold_map (fold_map_abort would be more correct, though)
 
haftmann 
parents: 
37754 
diff
changeset
 | 
368  | 
"fold_map f (xs @ ys) = fold_map f xs \<guillemotright>= (\<lambda>xs. fold_map f ys \<guillemotright>= (\<lambda>ys. return (xs @ ys)))"  | 
| 37754 | 369  | 
by (induct xs) simp_all  | 
370  | 
||
| 37758 | 371  | 
lemma execute_fold_map_unchanged_heap [execute_simps]:  | 
| 37754 | 372  | 
assumes "\<And>x. x \<in> set xs \<Longrightarrow> \<exists>y. execute (f x) h = Some (y, h)"  | 
| 
37756
 
59caa6180fff
avoid slightly odd "M" suffix; rename mapM to fold_map (fold_map_abort would be more correct, though)
 
haftmann 
parents: 
37754 
diff
changeset
 | 
373  | 
shows "execute (fold_map f xs) h =  | 
| 37754 | 374  | 
Some (List.map (\<lambda>x. fst (the (execute (f x) h))) xs, h)"  | 
375  | 
using assms proof (induct xs)  | 
|
| 
37787
 
30dc3abf4a58
theorem collections do not contain default rules any longer
 
haftmann 
parents: 
37772 
diff
changeset
 | 
376  | 
case Nil show ?case by (simp add: execute_simps)  | 
| 37754 | 377  | 
next  | 
378  | 
case (Cons x xs)  | 
|
379  | 
from Cons.prems obtain y  | 
|
380  | 
where y: "execute (f x) h = Some (y, h)" by auto  | 
|
| 
37756
 
59caa6180fff
avoid slightly odd "M" suffix; rename mapM to fold_map (fold_map_abort would be more correct, though)
 
haftmann 
parents: 
37754 
diff
changeset
 | 
381  | 
moreover from Cons.prems Cons.hyps have "execute (fold_map f xs) h =  | 
| 37754 | 382  | 
Some (map (\<lambda>x. fst (the (execute (f x) h))) xs, h)" by auto  | 
| 
37787
 
30dc3abf4a58
theorem collections do not contain default rules any longer
 
haftmann 
parents: 
37772 
diff
changeset
 | 
383  | 
ultimately show ?case by (simp, simp only: execute_bind(1), simp add: execute_simps)  | 
| 37754 | 384  | 
qed  | 
385  | 
||
| 40267 | 386  | 
|
387  | 
subsection {* Partial function definition setup *}
 | 
|
388  | 
||
389  | 
definition Heap_ord :: "'a Heap \<Rightarrow> 'a Heap \<Rightarrow> bool" where  | 
|
390  | 
"Heap_ord = img_ord execute (fun_ord option_ord)"  | 
|
391  | 
||
| 
44174
 
d1d79f0e1ea6
make more HOL theories work with separate set type
 
huffman 
parents: 
43324 
diff
changeset
 | 
392  | 
definition Heap_lub :: "'a Heap set \<Rightarrow> 'a Heap" where  | 
| 40267 | 393  | 
"Heap_lub = img_lub execute Heap (fun_lub (flat_lub None))"  | 
394  | 
||
| 
54630
 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 
Andreas Lochbihler 
parents: 
53361 
diff
changeset
 | 
395  | 
lemma Heap_lub_empty: "Heap_lub {} = Heap Map.empty"
 | 
| 
 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 
Andreas Lochbihler 
parents: 
53361 
diff
changeset
 | 
396  | 
by(simp add: Heap_lub_def img_lub_def fun_lub_def flat_lub_def)  | 
| 
 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 
Andreas Lochbihler 
parents: 
53361 
diff
changeset
 | 
397  | 
|
| 
51485
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
398  | 
lemma heap_interpretation: "partial_function_definitions Heap_ord Heap_lub"  | 
| 40267 | 399  | 
proof -  | 
400  | 
have "partial_function_definitions (fun_ord option_ord) (fun_lub (flat_lub None))"  | 
|
401  | 
by (rule partial_function_lift) (rule flat_interpretation)  | 
|
402  | 
then have "partial_function_definitions (img_ord execute (fun_ord option_ord))  | 
|
403  | 
(img_lub execute Heap (fun_lub (flat_lub None)))"  | 
|
404  | 
by (rule partial_function_image) (auto intro: Heap_eqI)  | 
|
405  | 
then show "partial_function_definitions Heap_ord Heap_lub"  | 
|
406  | 
by (simp only: Heap_ord_def Heap_lub_def)  | 
|
407  | 
qed  | 
|
408  | 
||
| 
51485
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
409  | 
interpretation heap!: partial_function_definitions Heap_ord Heap_lub  | 
| 
54630
 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 
Andreas Lochbihler 
parents: 
53361 
diff
changeset
 | 
410  | 
  where "Heap_lub {} \<equiv> Heap Map.empty"
 | 
| 
 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 
Andreas Lochbihler 
parents: 
53361 
diff
changeset
 | 
411  | 
by (fact heap_interpretation)(simp add: Heap_lub_empty)  | 
| 
51485
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
412  | 
|
| 
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
413  | 
lemma heap_step_admissible:  | 
| 
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
414  | 
"option.admissible  | 
| 
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
415  | 
      (\<lambda>f:: 'a => ('b * 'c) option. \<forall>h h' r. f h = Some (r, h') \<longrightarrow> P x h h' r)"
 | 
| 
53361
 
1cb7d3c0cf31
move admissible out of class ccpo to avoid unnecessary class predicate in foundational theorems
 
Andreas Lochbihler 
parents: 
52728 
diff
changeset
 | 
416  | 
proof (rule ccpo.admissibleI)  | 
| 
51485
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
417  | 
  fix A :: "('a \<Rightarrow> ('b * 'c) option) set"
 | 
| 
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
418  | 
assume ch: "Complete_Partial_Order.chain option.le_fun A"  | 
| 
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
419  | 
and IH: "\<forall>f\<in>A. \<forall>h h' r. f h = Some (r, h') \<longrightarrow> P x h h' r"  | 
| 
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
420  | 
  from ch have ch': "\<And>x. Complete_Partial_Order.chain option_ord {y. \<exists>f\<in>A. y = f x}" by (rule chain_fun)
 | 
| 
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
421  | 
show "\<forall>h h' r. option.lub_fun A h = Some (r, h') \<longrightarrow> P x h h' r"  | 
| 
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
422  | 
proof (intro allI impI)  | 
| 
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
423  | 
fix h h' r assume "option.lub_fun A h = Some (r, h')"  | 
| 
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
424  | 
from flat_lub_in_chain[OF ch' this[unfolded fun_lub_def]]  | 
| 
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
425  | 
    have "Some (r, h') \<in> {y. \<exists>f\<in>A. y = f h}" by simp
 | 
| 
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
426  | 
then have "\<exists>f\<in>A. f h = Some (r, h')" by auto  | 
| 
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
427  | 
with IH show "P x h h' r" by auto  | 
| 
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
428  | 
qed  | 
| 
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
429  | 
qed  | 
| 
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
430  | 
|
| 
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
431  | 
lemma admissible_heap:  | 
| 
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
432  | 
"heap.admissible (\<lambda>f. \<forall>x h h' r. effect (f x) h h' r \<longrightarrow> P x h h' r)"  | 
| 
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
433  | 
proof (rule admissible_fun[OF heap_interpretation])  | 
| 
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
434  | 
fix x  | 
| 
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
435  | 
show "ccpo.admissible Heap_lub Heap_ord (\<lambda>a. \<forall>h h' r. effect a h h' r \<longrightarrow> P x h h' r)"  | 
| 
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
436  | 
unfolding Heap_ord_def Heap_lub_def  | 
| 
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
437  | 
proof (intro admissible_image partial_function_lift flat_interpretation)  | 
| 
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
438  | 
show "option.admissible ((\<lambda>a. \<forall>h h' r. effect a h h' r \<longrightarrow> P x h h' r) \<circ> Heap)"  | 
| 
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
439  | 
unfolding comp_def effect_def execute.simps  | 
| 
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
440  | 
by (rule heap_step_admissible)  | 
| 
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
441  | 
qed (auto simp add: Heap_eqI)  | 
| 
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
442  | 
qed  | 
| 
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
443  | 
|
| 
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
444  | 
lemma fixp_induct_heap:  | 
| 
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
445  | 
fixes F :: "'c \<Rightarrow> 'c" and  | 
| 
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
446  | 
U :: "'c \<Rightarrow> 'b \<Rightarrow> 'a Heap" and  | 
| 
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
447  | 
    C :: "('b \<Rightarrow> 'a Heap) \<Rightarrow> 'c" and
 | 
| 
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
448  | 
P :: "'b \<Rightarrow> heap \<Rightarrow> heap \<Rightarrow> 'a \<Rightarrow> bool"  | 
| 
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
449  | 
assumes mono: "\<And>x. monotone (fun_ord Heap_ord) Heap_ord (\<lambda>f. U (F (C f)) x)"  | 
| 
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
450  | 
assumes eq: "f \<equiv> C (ccpo.fixp (fun_lub Heap_lub) (fun_ord Heap_ord) (\<lambda>f. U (F (C f))))"  | 
| 
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
451  | 
assumes inverse2: "\<And>f. U (C f) = f"  | 
| 
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
452  | 
assumes step: "\<And>f x h h' r. (\<And>x h h' r. effect (U f x) h h' r \<Longrightarrow> P x h h' r)  | 
| 
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
453  | 
\<Longrightarrow> effect (U (F f) x) h h' r \<Longrightarrow> P x h h' r"  | 
| 
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
454  | 
assumes defined: "effect (U f x) h h' r"  | 
| 
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
455  | 
shows "P x h h' r"  | 
| 
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
456  | 
using step defined heap.fixp_induct_uc[of U F C, OF mono eq inverse2 admissible_heap, of P]  | 
| 
54630
 
9061af4d5ebc
restrict admissibility to non-empty chains to allow more syntax-directed proof rules
 
Andreas Lochbihler 
parents: 
53361 
diff
changeset
 | 
457  | 
unfolding effect_def execute.simps  | 
| 
51485
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
458  | 
by blast  | 
| 
 
637aa1649ac7
added rudimentary induction rule for partial_function (heap)
 
krauss 
parents: 
51143 
diff
changeset
 | 
459  | 
|
| 
42949
 
618adb3584e5
separate initializations for different modes of partial_function -- generation of induction rules will be non-uniform
 
krauss 
parents: 
41413 
diff
changeset
 | 
460  | 
declaration {* Partial_Function.init "heap" @{term heap.fixp_fun}
 | 
| 
52728
 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 
krauss 
parents: 
52622 
diff
changeset
 | 
461  | 
  @{term heap.mono_body} @{thm heap.fixp_rule_uc} @{thm heap.fixp_induct_uc}
 | 
| 
 
470b579f35d2
derive specialized version of full fixpoint induction (with admissibility)
 
krauss 
parents: 
52622 
diff
changeset
 | 
462  | 
  (SOME @{thm fixp_induct_heap}) *}
 | 
| 
42949
 
618adb3584e5
separate initializations for different modes of partial_function -- generation of induction rules will be non-uniform
 
krauss 
parents: 
41413 
diff
changeset
 | 
463  | 
|
| 
 
618adb3584e5
separate initializations for different modes of partial_function -- generation of induction rules will be non-uniform
 
krauss 
parents: 
41413 
diff
changeset
 | 
464  | 
|
| 40267 | 465  | 
abbreviation "mono_Heap \<equiv> monotone (fun_ord Heap_ord) Heap_ord"  | 
466  | 
||
467  | 
lemma Heap_ordI:  | 
|
468  | 
assumes "\<And>h. execute x h = None \<or> execute x h = execute y h"  | 
|
469  | 
shows "Heap_ord x y"  | 
|
470  | 
using assms unfolding Heap_ord_def img_ord_def fun_ord_def flat_ord_def  | 
|
471  | 
by blast  | 
|
472  | 
||
473  | 
lemma Heap_ordE:  | 
|
474  | 
assumes "Heap_ord x y"  | 
|
475  | 
obtains "execute x h = None" | "execute x h = execute y h"  | 
|
476  | 
using assms unfolding Heap_ord_def img_ord_def fun_ord_def flat_ord_def  | 
|
477  | 
by atomize_elim blast  | 
|
478  | 
||
| 
46029
 
4a19e3d147c3
attribute code_abbrev superseedes code_unfold_post; tuned names and spacing
 
haftmann 
parents: 
45294 
diff
changeset
 | 
479  | 
lemma bind_mono [partial_function_mono]:  | 
| 40267 | 480  | 
assumes mf: "mono_Heap B" and mg: "\<And>y. mono_Heap (\<lambda>f. C y f)"  | 
481  | 
shows "mono_Heap (\<lambda>f. B f \<guillemotright>= (\<lambda>y. C y f))"  | 
|
482  | 
proof (rule monotoneI)  | 
|
483  | 
fix f g :: "'a \<Rightarrow> 'b Heap" assume fg: "fun_ord Heap_ord f g"  | 
|
484  | 
from mf  | 
|
485  | 
have 1: "Heap_ord (B f) (B g)" by (rule monotoneD) (rule fg)  | 
|
486  | 
from mg  | 
|
487  | 
have 2: "\<And>y'. Heap_ord (C y' f) (C y' g)" by (rule monotoneD) (rule fg)  | 
|
488  | 
||
489  | 
have "Heap_ord (B f \<guillemotright>= (\<lambda>y. C y f)) (B g \<guillemotright>= (\<lambda>y. C y f))"  | 
|
490  | 
(is "Heap_ord ?L ?R")  | 
|
491  | 
proof (rule Heap_ordI)  | 
|
492  | 
fix h  | 
|
493  | 
from 1 show "execute ?L h = None \<or> execute ?L h = execute ?R h"  | 
|
494  | 
by (rule Heap_ordE[where h = h]) (auto simp: execute_bind_case)  | 
|
495  | 
qed  | 
|
496  | 
also  | 
|
497  | 
have "Heap_ord (B g \<guillemotright>= (\<lambda>y'. C y' f)) (B g \<guillemotright>= (\<lambda>y'. C y' g))"  | 
|
498  | 
(is "Heap_ord ?L ?R")  | 
|
499  | 
proof (rule Heap_ordI)  | 
|
500  | 
fix h  | 
|
501  | 
show "execute ?L h = None \<or> execute ?L h = execute ?R h"  | 
|
502  | 
proof (cases "execute (B g) h")  | 
|
503  | 
case None  | 
|
504  | 
then have "execute ?L h = None" by (auto simp: execute_bind_case)  | 
|
505  | 
thus ?thesis ..  | 
|
506  | 
next  | 
|
507  | 
case Some  | 
|
508  | 
then obtain r h' where "execute (B g) h = Some (r, h')"  | 
|
509  | 
by (metis surjective_pairing)  | 
|
510  | 
then have "execute ?L h = execute (C r f) h'"  | 
|
511  | 
"execute ?R h = execute (C r g) h'"  | 
|
512  | 
by (auto simp: execute_bind_case)  | 
|
513  | 
with 2[of r] show ?thesis by (auto elim: Heap_ordE)  | 
|
514  | 
qed  | 
|
515  | 
qed  | 
|
516  | 
finally (heap.leq_trans)  | 
|
517  | 
show "Heap_ord (B f \<guillemotright>= (\<lambda>y. C y f)) (B g \<guillemotright>= (\<lambda>y'. C y' g))" .  | 
|
518  | 
qed  | 
|
519  | 
||
520  | 
||
| 26182 | 521  | 
subsection {* Code generator setup *}
 | 
522  | 
||
523  | 
subsubsection {* Logical intermediate layer *}
 | 
|
524  | 
||
| 
39250
 
548a3e5521ab
changing String.literal to a type instead of a datatype
 
bulwahn 
parents: 
39198 
diff
changeset
 | 
525  | 
definition raise' :: "String.literal \<Rightarrow> 'a Heap" where  | 
| 57437 | 526  | 
[code del]: "raise' s = raise (String.explode s)"  | 
| 
39250
 
548a3e5521ab
changing String.literal to a type instead of a datatype
 
bulwahn 
parents: 
39198 
diff
changeset
 | 
527  | 
|
| 
46029
 
4a19e3d147c3
attribute code_abbrev superseedes code_unfold_post; tuned names and spacing
 
haftmann 
parents: 
45294 
diff
changeset
 | 
528  | 
lemma [code_abbrev]: "raise' (STR s) = raise s"  | 
| 
 
4a19e3d147c3
attribute code_abbrev superseedes code_unfold_post; tuned names and spacing
 
haftmann 
parents: 
45294 
diff
changeset
 | 
529  | 
unfolding raise'_def by (simp add: STR_inverse)  | 
| 26182 | 530  | 
|
| 
46029
 
4a19e3d147c3
attribute code_abbrev superseedes code_unfold_post; tuned names and spacing
 
haftmann 
parents: 
45294 
diff
changeset
 | 
531  | 
lemma raise_raise': (* FIXME delete candidate *)  | 
| 37709 | 532  | 
"raise s = raise' (STR s)"  | 
| 
39250
 
548a3e5521ab
changing String.literal to a type instead of a datatype
 
bulwahn 
parents: 
39198 
diff
changeset
 | 
533  | 
unfolding raise'_def by (simp add: STR_inverse)  | 
| 26182 | 534  | 
|
| 37709 | 535  | 
code_datatype raise' -- {* avoid @{const "Heap"} formally *}
 | 
| 26182 | 536  | 
|
537  | 
||
| 27707 | 538  | 
subsubsection {* SML and OCaml *}
 | 
| 26182 | 539  | 
|
| 
52435
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
52388 
diff
changeset
 | 
540  | 
code_printing type_constructor Heap \<rightharpoonup> (SML) "(unit/ ->/ _)"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
52388 
diff
changeset
 | 
541  | 
code_printing constant bind \<rightharpoonup> (SML) "!(fn/ f'_/ =>/ fn/ ()/ =>/ f'_/ (_/ ())/ ())"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
52388 
diff
changeset
 | 
542  | 
code_printing constant return \<rightharpoonup> (SML) "!(fn/ ()/ =>/ _)"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
52388 
diff
changeset
 | 
543  | 
code_printing constant Heap_Monad.raise' \<rightharpoonup> (SML) "!(raise/ Fail/ _)"  | 
| 26182 | 544  | 
|
| 
52435
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
52388 
diff
changeset
 | 
545  | 
code_printing type_constructor Heap \<rightharpoonup> (OCaml) "(unit/ ->/ _)"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
52388 
diff
changeset
 | 
546  | 
code_printing constant bind \<rightharpoonup> (OCaml) "!(fun/ f'_/ ()/ ->/ f'_/ (_/ ())/ ())"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
52388 
diff
changeset
 | 
547  | 
code_printing constant return \<rightharpoonup> (OCaml) "!(fun/ ()/ ->/ _)"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
52388 
diff
changeset
 | 
548  | 
code_printing constant Heap_Monad.raise' \<rightharpoonup> (OCaml) "failwith"  | 
| 27707 | 549  | 
|
| 37838 | 550  | 
|
551  | 
subsubsection {* Haskell *}
 | 
|
552  | 
||
553  | 
text {* Adaption layer *}
 | 
|
554  | 
||
| 55372 | 555  | 
code_printing code_module "Heap" \<rightharpoonup> (Haskell)  | 
| 37838 | 556  | 
{*import qualified Control.Monad;
 | 
557  | 
import qualified Control.Monad.ST;  | 
|
558  | 
import qualified Data.STRef;  | 
|
559  | 
import qualified Data.Array.ST;  | 
|
560  | 
||
561  | 
type RealWorld = Control.Monad.ST.RealWorld;  | 
|
562  | 
type ST s a = Control.Monad.ST.ST s a;  | 
|
563  | 
type STRef s a = Data.STRef.STRef s a;  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
48073 
diff
changeset
 | 
564  | 
type STArray s a = Data.Array.ST.STArray s Integer a;  | 
| 37838 | 565  | 
|
566  | 
newSTRef = Data.STRef.newSTRef;  | 
|
567  | 
readSTRef = Data.STRef.readSTRef;  | 
|
568  | 
writeSTRef = Data.STRef.writeSTRef;  | 
|
569  | 
||
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
48073 
diff
changeset
 | 
570  | 
newArray :: Integer -> a -> ST s (STArray s a);  | 
| 37838 | 571  | 
newArray k = Data.Array.ST.newArray (0, k);  | 
572  | 
||
573  | 
newListArray :: [a] -> ST s (STArray s a);  | 
|
| 37964 | 574  | 
newListArray xs = Data.Array.ST.newListArray (0, (fromInteger . toInteger . length) xs) xs;  | 
| 37838 | 575  | 
|
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
48073 
diff
changeset
 | 
576  | 
newFunArray :: Integer -> (Integer -> a) -> ST s (STArray s a);  | 
| 37838 | 577  | 
newFunArray k f = Data.Array.ST.newListArray (0, k) (map f [0..k-1]);  | 
578  | 
||
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
48073 
diff
changeset
 | 
579  | 
lengthArray :: STArray s a -> ST s Integer;  | 
| 37838 | 580  | 
lengthArray a = Control.Monad.liftM snd (Data.Array.ST.getBounds a);  | 
581  | 
||
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
48073 
diff
changeset
 | 
582  | 
readArray :: STArray s a -> Integer -> ST s a;  | 
| 37838 | 583  | 
readArray = Data.Array.ST.readArray;  | 
584  | 
||
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
48073 
diff
changeset
 | 
585  | 
writeArray :: STArray s a -> Integer -> a -> ST s ();  | 
| 37838 | 586  | 
writeArray = Data.Array.ST.writeArray;*}  | 
587  | 
||
588  | 
code_reserved Haskell Heap  | 
|
589  | 
||
590  | 
text {* Monad *}
 | 
|
591  | 
||
| 
52435
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
52388 
diff
changeset
 | 
592  | 
code_printing type_constructor Heap \<rightharpoonup> (Haskell) "Heap.ST/ Heap.RealWorld/ _"  | 
| 37838 | 593  | 
code_monad bind Haskell  | 
| 
52435
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
52388 
diff
changeset
 | 
594  | 
code_printing constant return \<rightharpoonup> (Haskell) "return"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
52388 
diff
changeset
 | 
595  | 
code_printing constant Heap_Monad.raise' \<rightharpoonup> (Haskell) "error"  | 
| 37838 | 596  | 
|
597  | 
||
598  | 
subsubsection {* Scala *}
 | 
|
599  | 
||
| 55372 | 600  | 
code_printing code_module "Heap" \<rightharpoonup> (Scala)  | 
| 
38968
 
e55deaa22fff
do not print object frame around Scala includes -- this is in the responsibility of the user
 
haftmann 
parents: 
38773 
diff
changeset
 | 
601  | 
{*object Heap {
 | 
| 
 
e55deaa22fff
do not print object frame around Scala includes -- this is in the responsibility of the user
 
haftmann 
parents: 
38773 
diff
changeset
 | 
602  | 
def bind[A, B](f: Unit => A, g: A => Unit => B): Unit => B = (_: Unit) => g (f ()) ()  | 
| 
 
e55deaa22fff
do not print object frame around Scala includes -- this is in the responsibility of the user
 
haftmann 
parents: 
38773 
diff
changeset
 | 
603  | 
}  | 
| 37842 | 604  | 
|
605  | 
class Ref[A](x: A) {
 | 
|
606  | 
var value = x  | 
|
607  | 
}  | 
|
608  | 
||
609  | 
object Ref {
 | 
|
| 38771 | 610  | 
def apply[A](x: A): Ref[A] =  | 
611  | 
new Ref[A](x)  | 
|
612  | 
def lookup[A](r: Ref[A]): A =  | 
|
613  | 
r.value  | 
|
614  | 
def update[A](r: Ref[A], x: A): Unit =  | 
|
615  | 
    { r.value = x }
 | 
|
| 37842 | 616  | 
}  | 
617  | 
||
| 37964 | 618  | 
object Array {
 | 
| 
38968
 
e55deaa22fff
do not print object frame around Scala includes -- this is in the responsibility of the user
 
haftmann 
parents: 
38773 
diff
changeset
 | 
619  | 
import collection.mutable.ArraySeq  | 
| 
51143
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
48073 
diff
changeset
 | 
620  | 
def alloc[A](n: BigInt)(x: A): ArraySeq[A] =  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
48073 
diff
changeset
 | 
621  | 
ArraySeq.fill(n.toInt)(x)  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
48073 
diff
changeset
 | 
622  | 
def make[A](n: BigInt)(f: BigInt => A): ArraySeq[A] =  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
48073 
diff
changeset
 | 
623  | 
ArraySeq.tabulate(n.toInt)((k: Int) => f(BigInt(k)))  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
48073 
diff
changeset
 | 
624  | 
def len[A](a: ArraySeq[A]): BigInt =  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
48073 
diff
changeset
 | 
625  | 
BigInt(a.length)  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
48073 
diff
changeset
 | 
626  | 
def nth[A](a: ArraySeq[A], n: BigInt): A =  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
48073 
diff
changeset
 | 
627  | 
a(n.toInt)  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
48073 
diff
changeset
 | 
628  | 
def upd[A](a: ArraySeq[A], n: BigInt, x: A): Unit =  | 
| 
 
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
 
haftmann 
parents: 
48073 
diff
changeset
 | 
629  | 
a.update(n.toInt, x)  | 
| 38771 | 630  | 
def freeze[A](a: ArraySeq[A]): List[A] =  | 
631  | 
a.toList  | 
|
| 
38968
 
e55deaa22fff
do not print object frame around Scala includes -- this is in the responsibility of the user
 
haftmann 
parents: 
38773 
diff
changeset
 | 
632  | 
}  | 
| 
 
e55deaa22fff
do not print object frame around Scala includes -- this is in the responsibility of the user
 
haftmann 
parents: 
38773 
diff
changeset
 | 
633  | 
*}  | 
| 37842 | 634  | 
|
| 
38968
 
e55deaa22fff
do not print object frame around Scala includes -- this is in the responsibility of the user
 
haftmann 
parents: 
38773 
diff
changeset
 | 
635  | 
code_reserved Scala Heap Ref Array  | 
| 37838 | 636  | 
|
| 
52435
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
52388 
diff
changeset
 | 
637  | 
code_printing type_constructor Heap \<rightharpoonup> (Scala) "(Unit/ =>/ _)"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
52388 
diff
changeset
 | 
638  | 
code_printing constant bind \<rightharpoonup> (Scala) "Heap.bind"  | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
52388 
diff
changeset
 | 
639  | 
code_printing constant return \<rightharpoonup> (Scala) "('_: Unit)/ =>/ _"
 | 
| 
 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 
haftmann 
parents: 
52388 
diff
changeset
 | 
640  | 
code_printing constant Heap_Monad.raise' \<rightharpoonup> (Scala) "!sys.error((_))"  | 
| 37838 | 641  | 
|
642  | 
||
643  | 
subsubsection {* Target variants with less units *}
 | 
|
644  | 
||
| 31871 | 645  | 
setup {*
 | 
646  | 
||
647  | 
let  | 
|
| 27707 | 648  | 
|
| 31871 | 649  | 
open Code_Thingol;  | 
650  | 
||
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
54630 
diff
changeset
 | 
651  | 
val imp_program =  | 
| 31871 | 652  | 
let  | 
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
54630 
diff
changeset
 | 
653  | 
    val is_bind = curry (op =) @{const_name bind};
 | 
| 
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
54630 
diff
changeset
 | 
654  | 
    val is_return = curry (op =) @{const_name return};
 | 
| 31893 | 655  | 
val dummy_name = "";  | 
656  | 
val dummy_case_term = IVar NONE;  | 
|
| 31871 | 657  | 
(*assumption: dummy values are not relevant for serialization*)  | 
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
54630 
diff
changeset
 | 
658  | 
    val unitT = @{type_name unit} `%% [];
 | 
| 
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
54630 
diff
changeset
 | 
659  | 
val unitt =  | 
| 
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
54630 
diff
changeset
 | 
660  | 
      IConst { sym = Code_Symbol.Constant @{const_name Unity}, typargs = [], dicts = [], dom = [],
 | 
| 
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
54630 
diff
changeset
 | 
661  | 
range = unitT, annotate = false };  | 
| 31871 | 662  | 
fun dest_abs ((v, ty) `|=> t, _) = ((v, ty), t)  | 
663  | 
| dest_abs (t, ty) =  | 
|
664  | 
let  | 
|
665  | 
val vs = fold_varnames cons t [];  | 
|
| 
43324
 
2b47822868e4
discontinued Name.variant to emphasize that this is old-style / indirect;
 
wenzelm 
parents: 
43080 
diff
changeset
 | 
666  | 
val v = singleton (Name.variant_list vs) "x";  | 
| 31871 | 667  | 
val ty' = (hd o fst o unfold_fun) ty;  | 
| 31893 | 668  | 
in ((SOME v, ty'), t `$ IVar (SOME v)) end;  | 
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
54630 
diff
changeset
 | 
669  | 
    fun force (t as IConst { sym = Code_Symbol.Constant c, ... } `$ t') = if is_return c
 | 
| 31871 | 670  | 
then t' else t `$ unitt  | 
671  | 
| force t = t `$ unitt;  | 
|
| 38385 | 672  | 
fun tr_bind'' [(t1, _), (t2, ty2)] =  | 
| 31871 | 673  | 
let  | 
674  | 
val ((v, ty), t) = dest_abs (t2, ty2);  | 
|
| 
48072
 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 
haftmann 
parents: 
46029 
diff
changeset
 | 
675  | 
      in ICase { term = force t1, typ = ty, clauses = [(IVar v, tr_bind' t)], primitive = dummy_case_term } end
 | 
| 38385 | 676  | 
and tr_bind' t = case unfold_app t  | 
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
54630 
diff
changeset
 | 
677  | 
     of (IConst { sym = Code_Symbol.Constant c, dom = ty1 :: ty2 :: _, ... }, [x1, x2]) => if is_bind c
 | 
| 38386 | 678  | 
then tr_bind'' [(x1, ty1), (x2, ty2)]  | 
679  | 
else force t  | 
|
680  | 
| _ => force t;  | 
|
| 
48072
 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 
haftmann 
parents: 
46029 
diff
changeset
 | 
681  | 
fun imp_monad_bind'' ts = (SOME dummy_name, unitT) `|=>  | 
| 
 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 
haftmann 
parents: 
46029 
diff
changeset
 | 
682  | 
      ICase { term = IVar (SOME dummy_name), typ = unitT, clauses = [(unitt, tr_bind'' ts)], primitive = dummy_case_term }
 | 
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
54630 
diff
changeset
 | 
683  | 
    fun imp_monad_bind' (const as { sym = Code_Symbol.Constant c, dom = dom, ... }) ts = if is_bind c then case (ts, dom)
 | 
| 31871 | 684  | 
of ([t1, t2], ty1 :: ty2 :: _) => imp_monad_bind'' [(t1, ty1), (t2, ty2)]  | 
685  | 
| ([t1, t2, t3], ty1 :: ty2 :: _) => imp_monad_bind'' [(t1, ty1), (t2, ty2)] `$ t3  | 
|
686  | 
| (ts, _) => imp_monad_bind (eta_expand 2 (const, ts))  | 
|
687  | 
else IConst const `$$ map imp_monad_bind ts  | 
|
688  | 
and imp_monad_bind (IConst const) = imp_monad_bind' const []  | 
|
689  | 
| imp_monad_bind (t as IVar _) = t  | 
|
690  | 
| imp_monad_bind (t as _ `$ _) = (case unfold_app t  | 
|
691  | 
of (IConst const, ts) => imp_monad_bind' const ts  | 
|
692  | 
| (t, ts) => imp_monad_bind t `$$ map imp_monad_bind ts)  | 
|
693  | 
| imp_monad_bind (v_ty `|=> t) = v_ty `|=> imp_monad_bind t  | 
|
| 
48072
 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 
haftmann 
parents: 
46029 
diff
changeset
 | 
694  | 
      | imp_monad_bind (ICase { term = t, typ = ty, clauses = clauses, primitive = t0 }) =
 | 
| 
 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 
haftmann 
parents: 
46029 
diff
changeset
 | 
695  | 
          ICase { term = imp_monad_bind t, typ = ty,
 | 
| 
 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 
haftmann 
parents: 
46029 
diff
changeset
 | 
696  | 
clauses = (map o pairself) imp_monad_bind clauses, primitive = imp_monad_bind t0 };  | 
| 
28663
 
bd8438543bf2
code identifier namings are no longer imperative
 
haftmann 
parents: 
28562 
diff
changeset
 | 
697  | 
|
| 
55147
 
bce3dbc11f95
prefer explicit code symbol type over ad-hoc name mangling
 
haftmann 
parents: 
54630 
diff
changeset
 | 
698  | 
in (Code_Symbol.Graph.map o K o map_terms_stmt) imp_monad_bind end;  | 
| 27707 | 699  | 
|
700  | 
in  | 
|
701  | 
||
| 31871 | 702  | 
Code_Target.extend_target ("SML_imp", ("SML", imp_program))
 | 
703  | 
#> Code_Target.extend_target ("OCaml_imp", ("OCaml", imp_program))
 | 
|
| 37838 | 704  | 
#> Code_Target.extend_target ("Scala_imp", ("Scala", imp_program))
 | 
| 27707 | 705  | 
|
706  | 
end  | 
|
| 31871 | 707  | 
|
| 27707 | 708  | 
*}  | 
709  | 
||
| 37758 | 710  | 
hide_const (open) Heap heap guard raise' fold_map  | 
| 37724 | 711  | 
|
| 26170 | 712  | 
end  | 
| 
48072
 
ace701efe203
prefer records with speaking labels over deeply nested tuples
 
haftmann 
parents: 
46029 
diff
changeset
 | 
713  |