| author | wenzelm | 
| Sun, 07 Feb 2016 19:43:40 +0100 | |
| changeset 62271 | 4cfe65cfd369 | 
| parent 62042 | 6c6ccf573479 | 
| child 67443 | 3abf6a722518 | 
| permissions | -rw-r--r-- | 
| 12857 | 1  | 
(* Title: HOL/Bali/Type.thy  | 
| 12854 | 2  | 
Author: David von Oheimb  | 
3  | 
*)  | 
|
4  | 
||
| 62042 | 5  | 
subsection \<open>Java types\<close>  | 
| 12854 | 6  | 
|
| 16417 | 7  | 
theory Type imports Name begin  | 
| 12854 | 8  | 
|
| 62042 | 9  | 
text \<open>  | 
| 12854 | 10  | 
simplifications:  | 
11  | 
\begin{itemize}
 | 
|
12  | 
\item only the most important primitive types  | 
|
13  | 
\item the null type is regarded as reference type  | 
|
14  | 
\end{itemize}
 | 
|
| 62042 | 15  | 
\<close>  | 
| 12854 | 16  | 
|
| 62042 | 17  | 
datatype prim_ty \<comment>\<open>primitive type, cf. 4.2\<close>  | 
18  | 
= Void \<comment>\<open>result type of void methods\<close>  | 
|
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
16417 
diff
changeset
 | 
19  | 
| Boolean  | 
| 
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
16417 
diff
changeset
 | 
20  | 
| Integer  | 
| 12854 | 21  | 
|
22  | 
||
| 62042 | 23  | 
datatype ref_ty \<comment>\<open>reference type, cf. 4.3\<close>  | 
24  | 
= NullT \<comment>\<open>null type, cf. 4.1\<close>  | 
|
25  | 
| IfaceT qtname \<comment>\<open>interface type\<close>  | 
|
26  | 
| ClassT qtname \<comment>\<open>class type\<close>  | 
|
27  | 
| ArrayT ty \<comment>\<open>array type\<close>  | 
|
| 12854 | 28  | 
|
| 62042 | 29  | 
and ty \<comment>\<open>any type, cf. 4.1\<close>  | 
30  | 
= PrimT prim_ty \<comment>\<open>primitive type\<close>  | 
|
31  | 
| RefT ref_ty \<comment>\<open>reference type\<close>  | 
|
| 12854 | 32  | 
|
| 
35067
 
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
 
wenzelm 
parents: 
32960 
diff
changeset
 | 
33  | 
abbreviation "NT == RefT NullT"  | 
| 
 
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
 
wenzelm 
parents: 
32960 
diff
changeset
 | 
34  | 
abbreviation "Iface I == RefT (IfaceT I)"  | 
| 
 
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
 
wenzelm 
parents: 
32960 
diff
changeset
 | 
35  | 
abbreviation "Class C == RefT (ClassT C)"  | 
| 
 
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
 
wenzelm 
parents: 
32960 
diff
changeset
 | 
36  | 
abbreviation Array :: "ty \<Rightarrow> ty"  ("_.[]" [90] 90)
 | 
| 
 
af4c18c30593
modernized syntax translations, using mostly abbreviation/notation;
 
wenzelm 
parents: 
32960 
diff
changeset
 | 
37  | 
where "T.[] == RefT (ArrayT T)"  | 
| 12854 | 38  | 
|
| 37956 | 39  | 
definition  | 
40  | 
the_Class :: "ty \<Rightarrow> qtname"  | 
|
41  | 
where "the_Class T = (SOME C. T = Class C)" (** primrec not possible here **)  | 
|
| 12854 | 42  | 
|
43  | 
lemma the_Class_eq [simp]: "the_Class (Class C)= C"  | 
|
44  | 
by (auto simp add: the_Class_def)  | 
|
45  | 
||
46  | 
end  |