| author | wenzelm | 
| Tue, 26 Jan 2021 23:34:40 +0100 | |
| changeset 73194 | c0d6d57a9a31 | 
| parent 67443 | 3abf6a722518 | 
| child 80914 | d97fdabd9e2b | 
| permissions | -rw-r--r-- | 
| 8011 | 1 | (* Title: HOL/MicroJava/J/Term.thy | 
| 41589 | 2 | Author: David von Oheimb, Technische Universitaet Muenchen | 
| 11070 | 3 | *) | 
| 8011 | 4 | |
| 61361 | 5 | section \<open>Expressions and Statements\<close> | 
| 8011 | 6 | |
| 16417 | 7 | theory Term imports Value begin | 
| 8011 | 8 | |
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
62042diff
changeset | 9 | datatype binop = Eq | Add \<comment> \<open>function codes for binary operation\<close> | 
| 9240 | 10 | |
| 58310 | 11 | datatype expr | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
62042diff
changeset | 12 | = NewC cname \<comment> \<open>class instance creation\<close> | 
| 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
62042diff
changeset | 13 | | Cast cname expr \<comment> \<open>type cast\<close> | 
| 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
62042diff
changeset | 14 | | Lit val \<comment> \<open>literal value, also references\<close> | 
| 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
62042diff
changeset | 15 | | BinOp binop expr expr \<comment> \<open>binary operation\<close> | 
| 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
62042diff
changeset | 16 | | LAcc vname \<comment> \<open>local (incl. parameter) access\<close> | 
| 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
62042diff
changeset | 17 |   | LAss vname expr          ("_::=_" [90,90]90)      \<comment> \<open>local assign\<close>
 | 
| 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
62042diff
changeset | 18 |   | FAcc cname expr vname    ("{_}_.._" [10,90,99]90) \<comment> \<open>field access\<close>
 | 
| 10069 | 19 | | FAss cname expr vname | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
62042diff
changeset | 20 |                     expr     ("{_}_.._:=_" [10,90,99,90]90) \<comment> \<open>field ass.\<close>
 | 
| 10763 | 21 | | Call cname expr mname | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
62042diff
changeset | 22 |     "ty list" "expr list"    ("{_}_.._'( {_}_')" [10,90,99,10,10] 90) \<comment> \<open>method call\<close>
 | 
| 58263 | 23 | |
| 24 | datatype_compat expr | |
| 10061 
fe82134773dc
added HTML syntax; added spaces in normal syntax for better documents
 kleing parents: 
10042diff
changeset | 25 | |
| 58310 | 26 | datatype stmt | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
62042diff
changeset | 27 | = Skip \<comment> \<open>empty statement\<close> | 
| 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
62042diff
changeset | 28 | | Expr expr \<comment> \<open>expression statement\<close> | 
| 10061 
fe82134773dc
added HTML syntax; added spaces in normal syntax for better documents
 kleing parents: 
10042diff
changeset | 29 |   | Comp stmt stmt       ("_;; _"             [61,60]60)
 | 
| 
fe82134773dc
added HTML syntax; added spaces in normal syntax for better documents
 kleing parents: 
10042diff
changeset | 30 |   | Cond expr stmt stmt  ("If '(_') _ Else _" [80,79,79]70)
 | 
| 
fe82134773dc
added HTML syntax; added spaces in normal syntax for better documents
 kleing parents: 
10042diff
changeset | 31 |   | Loop expr stmt       ("While '(_') _"     [80,79]70)
 | 
| 
fe82134773dc
added HTML syntax; added spaces in normal syntax for better documents
 kleing parents: 
10042diff
changeset | 32 | |
| 8011 | 33 | end | 
| 11026 
a50365d21144
converted to Isar, simplifying recursion on class hierarchy
 oheimb parents: 
10763diff
changeset | 34 |