| author | wenzelm | 
| Fri, 14 Mar 2025 23:03:58 +0100 | |
| changeset 82276 | d22e9c5b5dc6 | 
| parent 77106 | 5ef443fa4a5d | 
| permissions | -rw-r--r-- | 
| 37714 | 1 | (* Title: HOL/Imperative_HOL/Heap.thy | 
| 26170 | 2 | Author: John Matthews, Galois Connections; Alexander Krauss, TU Muenchen | 
| 3 | *) | |
| 4 | ||
| 63167 | 5 | section \<open>A polymorphic heap based on cantor encodings\<close> | 
| 26170 | 6 | |
| 7 | theory Heap | |
| 66453 
cc19f7ca2ed6
session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
 wenzelm parents: 
66004diff
changeset | 8 | imports Main "HOL-Library.Countable" | 
| 26170 | 9 | begin | 
| 10 | ||
| 63167 | 11 | subsection \<open>Representable types\<close> | 
| 26170 | 12 | |
| 63167 | 13 | text \<open>The type class of representable types\<close> | 
| 26170 | 14 | |
| 28335 | 15 | class heap = typerep + countable | 
| 26170 | 16 | |
| 38408 | 17 | instance unit :: heap .. | 
| 18 | ||
| 19 | instance bool :: heap .. | |
| 20 | ||
| 26170 | 21 | instance nat :: heap .. | 
| 22 | ||
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
36640diff
changeset | 23 | instance prod :: (heap, heap) heap .. | 
| 26170 | 24 | |
| 37678 
0040bafffdef
"prod" and "sum" replace "*" and "+" respectively
 haftmann parents: 
36640diff
changeset | 25 | instance sum :: (heap, heap) heap .. | 
| 26170 | 26 | |
| 27 | instance list :: (heap) heap .. | |
| 28 | ||
| 29 | instance option :: (heap) heap .. | |
| 30 | ||
| 31 | instance int :: heap .. | |
| 32 | ||
| 31205 
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
 haftmann parents: 
30738diff
changeset | 33 | instance String.literal :: heap .. | 
| 26170 | 34 | |
| 66004 
797ef4889177
Added char, ref, array to heap-storable types
 lammich <lammich@in.tum.de> parents: 
63167diff
changeset | 35 | instance char :: heap .. | 
| 
797ef4889177
Added char, ref, array to heap-storable types
 lammich <lammich@in.tum.de> parents: 
63167diff
changeset | 36 | |
| 28335 | 37 | instance typerep :: heap .. | 
| 26170 | 38 | |
| 39 | ||
| 63167 | 40 | subsection \<open>A polymorphic heap with dynamic arrays and references\<close> | 
| 26170 | 41 | |
| 63167 | 42 | text \<open> | 
| 37719 
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
 haftmann parents: 
37718diff
changeset | 43 | References and arrays are developed in parallel, | 
| 
271ecd4fb9f9
moved "open" operations from Heap.thy to Array.thy and Ref.thy
 haftmann parents: 
37718diff
changeset | 44 | but keeping them separate makes some later proofs simpler. | 
| 63167 | 45 | \<close> | 
| 37714 | 46 | |
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
66453diff
changeset | 47 | type_synonym addr = nat \<comment> \<open>untyped heap references\<close> | 
| 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
66453diff
changeset | 48 | type_synonym heap_rep = nat \<comment> \<open>representable values\<close> | 
| 37714 | 49 | |
| 50 | record heap = | |
| 51 | arrays :: "typerep \<Rightarrow> addr \<Rightarrow> heap_rep list" | |
| 52 | refs :: "typerep \<Rightarrow> addr \<Rightarrow> heap_rep" | |
| 53 | lim :: addr | |
| 54 | ||
| 55 | definition empty :: heap where | |
| 37723 | 56 | "empty = \<lparr>arrays = (\<lambda>_ _. []), refs = (\<lambda>_ _. 0), lim = 0\<rparr>" | 
| 26170 | 57 | |
| 67443 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
66453diff
changeset | 58 | datatype 'a array = Array addr \<comment> \<open>note the phantom type 'a\<close> | 
| 
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
 wenzelm parents: 
66453diff
changeset | 59 | datatype 'a ref = Ref addr \<comment> \<open>note the phantom type 'a\<close> | 
| 26170 | 60 | |
| 61 | primrec addr_of_array :: "'a array \<Rightarrow> addr" where | |
| 62 | "addr_of_array (Array x) = x" | |
| 63 | ||
| 64 | primrec addr_of_ref :: "'a ref \<Rightarrow> addr" where | |
| 65 | "addr_of_ref (Ref x) = x" | |
| 66 | ||
| 67 | lemma addr_of_array_inj [simp]: | |
| 68 | "addr_of_array a = addr_of_array a' \<longleftrightarrow> a = a'" | |
| 69 | by (cases a, cases a') simp_all | |
| 70 | ||
| 71 | lemma addr_of_ref_inj [simp]: | |
| 72 | "addr_of_ref r = addr_of_ref r' \<longleftrightarrow> r = r'" | |
| 73 | by (cases r, cases r') simp_all | |
| 74 | ||
| 75 | instance array :: (type) countable | |
| 76 | by (rule countable_classI [of addr_of_array]) simp | |
| 77 | ||
| 78 | instance ref :: (type) countable | |
| 79 | by (rule countable_classI [of addr_of_ref]) simp | |
| 80 | ||
| 66004 
797ef4889177
Added char, ref, array to heap-storable types
 lammich <lammich@in.tum.de> parents: 
63167diff
changeset | 81 | instance array :: (type) heap .. | 
| 77106 | 82 | |
| 66004 
797ef4889177
Added char, ref, array to heap-storable types
 lammich <lammich@in.tum.de> parents: 
63167diff
changeset | 83 | instance ref :: (type) heap .. | 
| 77106 | 84 | |
| 85 | ||
| 63167 | 86 | text \<open>Syntactic convenience\<close> | 
| 37714 | 87 | |
| 63167 | 88 | setup \<open> | 
| 69597 | 89 | Sign.add_const_constraint (\<^const_name>\<open>Array\<close>, SOME \<^typ>\<open>nat \<Rightarrow> 'a::heap array\<close>) | 
| 90 | #> Sign.add_const_constraint (\<^const_name>\<open>Ref\<close>, SOME \<^typ>\<open>nat \<Rightarrow> 'a::heap ref\<close>) | |
| 91 | #> Sign.add_const_constraint (\<^const_name>\<open>addr_of_array\<close>, SOME \<^typ>\<open>'a::heap array \<Rightarrow> nat\<close>) | |
| 92 | #> Sign.add_const_constraint (\<^const_name>\<open>addr_of_ref\<close>, SOME \<^typ>\<open>'a::heap ref \<Rightarrow> nat\<close>) | |
| 63167 | 93 | \<close> | 
| 26170 | 94 | |
| 37723 | 95 | hide_const (open) empty | 
| 96 | ||
| 26170 | 97 | end |