| author | wenzelm | 
| Mon, 09 Mar 2020 14:30:09 +0100 | |
| changeset 71529 | dd56597e026b | 
| parent 67443 | 3abf6a722518 | 
| permissions | -rw-r--r-- | 
| 11376 | 1 | (* Title: HOL/NanoJava/Decl.thy | 
| 2 | Author: David von Oheimb | |
| 3 | Copyright 2001 Technische Universitaet Muenchen | |
| 4 | *) | |
| 5 | ||
| 58889 | 6 | section "Types, class Declarations, and whole programs" | 
| 11376 | 7 | |
| 16417 | 8 | theory Decl imports Term begin | 
| 11376 | 9 | |
| 58310 | 10 | datatype ty | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
63167diff
changeset | 11 | = NT \<comment> \<open>null type\<close> | 
| 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
63167diff
changeset | 12 | | Class cname \<comment> \<open>class type\<close> | 
| 11376 | 13 | |
| 63167 | 14 | text\<open>Field declaration\<close> | 
| 42463 | 15 | type_synonym fdecl | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
18551diff
changeset | 16 | = "fname \<times> ty" | 
| 11376 | 17 | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
18551diff
changeset | 18 | record methd | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
18551diff
changeset | 19 | = par :: ty | 
| 11376 | 20 | res :: ty | 
| 21 | lcl ::"(vname \<times> ty) list" | |
| 22 | bdy :: stmt | |
| 23 | ||
| 63167 | 24 | text\<open>Method declaration\<close> | 
| 42463 | 25 | type_synonym mdecl | 
| 11376 | 26 | = "mname \<times> methd" | 
| 27 | ||
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
18551diff
changeset | 28 | record "class" | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
18551diff
changeset | 29 | = super :: cname | 
| 12264 
9c356e2da72f
renamed "fields" to "flds" (avoid clash with new "fields" operation);
 wenzelm parents: 
11565diff
changeset | 30 | flds ::"fdecl list" | 
| 11376 | 31 | methods ::"mdecl list" | 
| 32 | ||
| 63167 | 33 | text\<open>Class declaration\<close> | 
| 42463 | 34 | type_synonym cdecl | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
18551diff
changeset | 35 | = "cname \<times> class" | 
| 11376 | 36 | |
| 42463 | 37 | type_synonym prog | 
| 11376 | 38 | = "cdecl list" | 
| 39 | ||
| 40 | translations | |
| 35431 | 41 | (type) "fdecl" \<leftharpoondown> (type) "fname \<times> ty" | 
| 42 | (type) "mdecl" \<leftharpoondown> (type) "mname \<times> ty \<times> ty \<times> stmt" | |
| 43 | (type) "class" \<leftharpoondown> (type) "cname \<times> fdecl list \<times> mdecl list" | |
| 44 | (type) "cdecl" \<leftharpoondown> (type) "cname \<times> class" | |
| 45 | (type) "prog " \<leftharpoondown> (type) "cdecl list" | |
| 11376 | 46 | |
| 44375 | 47 | axiomatization | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
63167diff
changeset | 48 | Prog :: prog \<comment> \<open>program as a global value\<close> | 
| 44375 | 49 | and | 
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
63167diff
changeset | 50 | Object :: cname \<comment> \<open>name of root class\<close> | 
| 11376 | 51 | |
| 52 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 53 | definition "class" :: "cname \<rightharpoonup> class" where | 
| 11376 | 54 | "class \<equiv> map_of Prog" | 
| 55 | ||
| 35416 
d8d7d1b785af
replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
 haftmann parents: 
32960diff
changeset | 56 | definition is_class :: "cname => bool" where | 
| 11376 | 57 | "is_class C \<equiv> class C \<noteq> None" | 
| 58 | ||
| 59 | lemma finite_is_class: "finite {C. is_class C}"
 | |
| 60 | apply (unfold is_class_def class_def) | |
| 61 | apply (fold dom_def) | |
| 62 | apply (rule finite_dom_map_of) | |
| 63 | done | |
| 64 | ||
| 65 | end |