| author | wenzelm | 
| Mon, 27 Nov 2017 15:59:24 +0100 | |
| changeset 67097 | d1b8464654c5 | 
| parent 58310 | 91ea607a34d8 | 
| child 67406 | 23307fd33906 | 
| permissions | -rw-r--r-- | 
| 43158 | 1  | 
theory OO imports Main begin  | 
2  | 
||
3  | 
subsection "Towards an OO Language: A Language of Records"  | 
|
4  | 
||
5  | 
(* FIXME: move to HOL/Fun *)  | 
|
6  | 
abbreviation fun_upd2 :: "('a \<Rightarrow> 'b \<Rightarrow> 'c) \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'c \<Rightarrow> 'a \<Rightarrow> 'b \<Rightarrow> 'c"
 | 
|
7  | 
  ("_/'((2_,_ :=/ _)')" [1000,0,0,0] 900)
 | 
|
8  | 
where "f(x,y := z) == f(x := (f x)(y := z))"  | 
|
9  | 
||
10  | 
type_synonym addr = nat  | 
|
| 58310 | 11  | 
datatype ref = null | Ref addr  | 
| 43158 | 12  | 
|
13  | 
type_synonym obj = "string \<Rightarrow> ref"  | 
|
14  | 
type_synonym venv = "string \<Rightarrow> ref"  | 
|
15  | 
type_synonym store = "addr \<Rightarrow> obj"  | 
|
16  | 
||
| 58310 | 17  | 
datatype exp =  | 
| 43158 | 18  | 
Null |  | 
19  | 
New |  | 
|
20  | 
V string |  | 
|
21  | 
  Faccess exp string       ("_\<bullet>/_" [63,1000] 63) |
 | 
|
22  | 
  Vassign string exp       ("(_ ::=/ _)" [1000,61] 62) |
 | 
|
23  | 
  Fassign exp string exp   ("(_\<bullet>_ ::=/ _)" [63,0,62] 62) |
 | 
|
24  | 
  Mcall exp string exp     ("(_\<bullet>/_<_>)" [63,0,0] 63) |
 | 
|
| 47818 | 25  | 
  Seq exp exp              ("_;/ _" [61,60] 60) |
 | 
| 43158 | 26  | 
  If bexp exp exp          ("IF _/ THEN (2_)/ ELSE (2_)" [0,0,61] 61)
 | 
27  | 
and bexp = B bool | Not bexp | And bexp bexp | Eq exp exp  | 
|
28  | 
||
29  | 
type_synonym menv = "string \<Rightarrow> exp"  | 
|
30  | 
type_synonym config = "venv \<times> store \<times> addr"  | 
|
31  | 
||
32  | 
inductive  | 
|
33  | 
big_step :: "menv \<Rightarrow> exp \<times> config \<Rightarrow> ref \<times> config \<Rightarrow> bool"  | 
|
34  | 
    ("(_ \<turnstile>/ (_/ \<Rightarrow> _))" [60,0,60] 55) and
 | 
|
35  | 
bval :: "menv \<Rightarrow> bexp \<times> config \<Rightarrow> bool \<times> config \<Rightarrow> bool"  | 
|
36  | 
    ("_ \<turnstile> _ \<rightarrow> _" [60,0,60] 55)
 | 
|
37  | 
where  | 
|
38  | 
Null:  | 
|
39  | 
"me \<turnstile> (Null,c) \<Rightarrow> (null,c)" |  | 
|
40  | 
New:  | 
|
41  | 
"me \<turnstile> (New,ve,s,n) \<Rightarrow> (Ref n,ve,s(n := (\<lambda>f. null)),n+1)" |  | 
|
42  | 
Vaccess:  | 
|
43  | 
"me \<turnstile> (V x,ve,sn) \<Rightarrow> (ve x,ve,sn)" |  | 
|
44  | 
Faccess:  | 
|
45  | 
"me \<turnstile> (e,c) \<Rightarrow> (Ref a,ve',s',n') \<Longrightarrow>  | 
|
46  | 
me \<turnstile> (e\<bullet>f,c) \<Rightarrow> (s' a f,ve',s',n')" |  | 
|
47  | 
Vassign:  | 
|
48  | 
"me \<turnstile> (e,c) \<Rightarrow> (r,ve',sn') \<Longrightarrow>  | 
|
49  | 
me \<turnstile> (x ::= e,c) \<Rightarrow> (r,ve'(x:=r),sn')" |  | 
|
50  | 
Fassign:  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
47818 
diff
changeset
 | 
51  | 
"\<lbrakk> me \<turnstile> (oe,c\<^sub>1) \<Rightarrow> (Ref a,c\<^sub>2); me \<turnstile> (e,c\<^sub>2) \<Rightarrow> (r,ve\<^sub>3,s\<^sub>3,n\<^sub>3) \<rbrakk> \<Longrightarrow>  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
47818 
diff
changeset
 | 
52  | 
me \<turnstile> (oe\<bullet>f ::= e,c\<^sub>1) \<Rightarrow> (r,ve\<^sub>3,s\<^sub>3(a,f := r),n\<^sub>3)" |  | 
| 43158 | 53  | 
Mcall:  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
47818 
diff
changeset
 | 
54  | 
"\<lbrakk> me \<turnstile> (oe,c\<^sub>1) \<Rightarrow> (or,c\<^sub>2); me \<turnstile> (pe,c\<^sub>2) \<Rightarrow> (pr,ve\<^sub>3,sn\<^sub>3);  | 
| 43158 | 55  | 
   ve = (\<lambda>x. null)(''this'' := or, ''param'' := pr);
 | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
47818 
diff
changeset
 | 
56  | 
me \<turnstile> (me m,ve,sn\<^sub>3) \<Rightarrow> (r,ve',sn\<^sub>4) \<rbrakk>  | 
| 43158 | 57  | 
\<Longrightarrow>  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
47818 
diff
changeset
 | 
58  | 
me \<turnstile> (oe\<bullet>m<pe>,c\<^sub>1) \<Rightarrow> (r,ve\<^sub>3,sn\<^sub>4)" |  | 
| 47818 | 59  | 
Seq:  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
47818 
diff
changeset
 | 
60  | 
"\<lbrakk> me \<turnstile> (e\<^sub>1,c\<^sub>1) \<Rightarrow> (r,c\<^sub>2); me \<turnstile> (e\<^sub>2,c\<^sub>2) \<Rightarrow> c\<^sub>3 \<rbrakk> \<Longrightarrow>  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
47818 
diff
changeset
 | 
61  | 
me \<turnstile> (e\<^sub>1; e\<^sub>2,c\<^sub>1) \<Rightarrow> c\<^sub>3" |  | 
| 43158 | 62  | 
IfTrue:  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
47818 
diff
changeset
 | 
63  | 
"\<lbrakk> me \<turnstile> (b,c\<^sub>1) \<rightarrow> (True,c\<^sub>2); me \<turnstile> (e\<^sub>1,c\<^sub>2) \<Rightarrow> c\<^sub>3 \<rbrakk> \<Longrightarrow>  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
47818 
diff
changeset
 | 
64  | 
me \<turnstile> (IF b THEN e\<^sub>1 ELSE e\<^sub>2,c\<^sub>1) \<Rightarrow> c\<^sub>3" |  | 
| 43158 | 65  | 
IfFalse:  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
47818 
diff
changeset
 | 
66  | 
"\<lbrakk> me \<turnstile> (b,c\<^sub>1) \<rightarrow> (False,c\<^sub>2); me \<turnstile> (e\<^sub>2,c\<^sub>2) \<Rightarrow> c\<^sub>3 \<rbrakk> \<Longrightarrow>  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
47818 
diff
changeset
 | 
67  | 
me \<turnstile> (IF b THEN e\<^sub>1 ELSE e\<^sub>2,c\<^sub>1) \<Rightarrow> c\<^sub>3" |  | 
| 43158 | 68  | 
|
69  | 
"me \<turnstile> (B bv,c) \<rightarrow> (bv,c)" |  | 
|
70  | 
||
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
47818 
diff
changeset
 | 
71  | 
"me \<turnstile> (b,c\<^sub>1) \<rightarrow> (bv,c\<^sub>2) \<Longrightarrow> me \<turnstile> (Not b,c\<^sub>1) \<rightarrow> (\<not>bv,c\<^sub>2)" |  | 
| 43158 | 72  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
47818 
diff
changeset
 | 
73  | 
"\<lbrakk> me \<turnstile> (b\<^sub>1,c\<^sub>1) \<rightarrow> (bv\<^sub>1,c\<^sub>2); me \<turnstile> (b\<^sub>2,c\<^sub>2) \<rightarrow> (bv\<^sub>2,c\<^sub>3) \<rbrakk> \<Longrightarrow>  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
47818 
diff
changeset
 | 
74  | 
me \<turnstile> (And b\<^sub>1 b\<^sub>2,c\<^sub>1) \<rightarrow> (bv\<^sub>1\<and>bv\<^sub>2,c\<^sub>3)" |  | 
| 43158 | 75  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
47818 
diff
changeset
 | 
76  | 
"\<lbrakk> me \<turnstile> (e\<^sub>1,c\<^sub>1) \<Rightarrow> (r\<^sub>1,c\<^sub>2); me \<turnstile> (e\<^sub>2,c\<^sub>2) \<Rightarrow> (r\<^sub>2,c\<^sub>3) \<rbrakk> \<Longrightarrow>  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
47818 
diff
changeset
 | 
77  | 
me \<turnstile> (Eq e\<^sub>1 e\<^sub>2,c\<^sub>1) \<rightarrow> (r\<^sub>1=r\<^sub>2,c\<^sub>3)"  | 
| 43158 | 78  | 
|
79  | 
||
80  | 
code_pred (modes: i => i => o => bool) big_step .  | 
|
81  | 
||
82  | 
text{* Example: natural numbers encoded as objects with a predecessor
 | 
|
83  | 
field. Null is zero. Method succ adds an object in front, method add  | 
|
84  | 
adds as many objects in front as the parameter specifies.  | 
|
85  | 
||
86  | 
First, the method bodies: *}  | 
|
87  | 
||
88  | 
definition  | 
|
89  | 
"m_succ  =  (''s'' ::= New)\<bullet>''pred'' ::= V ''this''; V ''s''"
 | 
|
90  | 
||
91  | 
definition "m_add =  | 
|
92  | 
IF Eq (V ''param'') Null  | 
|
93  | 
THEN V ''this''  | 
|
94  | 
ELSE V ''this''\<bullet>''succ''<Null>\<bullet>''add''<V ''param''\<bullet>''pred''>"  | 
|
95  | 
||
96  | 
text{* The method environment: *}
 | 
|
97  | 
definition  | 
|
98  | 
"menv = (\<lambda>m. Null)(''succ'' := m_succ, ''add'' := m_add)"
 | 
|
99  | 
||
100  | 
text{* The main code, adding 1 and 2: *}
 | 
|
101  | 
definition "main =  | 
|
102  | 
''1'' ::= Null\<bullet>''succ''<Null>;  | 
|
103  | 
''2'' ::= V ''1''\<bullet>''succ''<Null>;  | 
|
104  | 
V ''2'' \<bullet> ''add'' <V ''1''>"  | 
|
105  | 
||
106  | 
text{* Execution of semantics. The final variable environment and store are
 | 
|
107  | 
converted into lists of references based on given lists of variable and field  | 
|
108  | 
names to extract. *}  | 
|
109  | 
||
110  | 
values  | 
|
111  | 
 "{(r, map ve' [''1'',''2''], map (\<lambda>n. map (s' n)[''pred'']) [0..<n])|
 | 
|
112  | 
r ve' s' n. menv \<turnstile> (main, \<lambda>x. null, nth[], 0) \<Rightarrow> (r,ve',s',n)}"  | 
|
113  | 
||
114  | 
end  |