| author | wenzelm | 
| Thu, 07 Apr 2016 22:09:23 +0200 | |
| changeset 62912 | 745d31e63c21 | 
| parent 61115 | 3a4400985780 | 
| child 63649 | e690d6f2185b | 
| permissions | -rw-r--r-- | 
| 58881 | 1 | section "Immutable Arrays with Code Generation" | 
| 50138 | 2 | |
| 3 | theory IArray | |
| 51094 
84b03c49c223
IArray ignorant of particular representation of nat
 haftmann parents: 
50138diff
changeset | 4 | imports Main | 
| 50138 | 5 | begin | 
| 6 | ||
| 60500 | 7 | text\<open>Immutable arrays are lists wrapped up in an additional constructor. | 
| 50138 | 8 | There are no update operations. Hence code generation can safely implement | 
| 9 | this type by efficient target language arrays. Currently only SML is | |
| 10 | provided. Should be extended to other target languages and more operations. | |
| 11 | ||
| 12 | Note that arrays cannot be printed directly but only by turning them into | |
| 13 | lists first. Arrays could be converted back into lists for printing if they | |
| 60500 | 14 | were wrapped up in an additional constructor.\<close> | 
| 50138 | 15 | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 16 | context | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 17 | begin | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 18 | |
| 58310 | 19 | datatype 'a iarray = IArray "'a list" | 
| 50138 | 20 | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 21 | qualified primrec list_of :: "'a iarray \<Rightarrow> 'a list" where | 
| 50138 | 22 | "list_of (IArray xs) = xs" | 
| 23 | ||
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 24 | qualified definition of_fun :: "(nat \<Rightarrow> 'a) \<Rightarrow> nat \<Rightarrow> 'a iarray" where | 
| 51094 
84b03c49c223
IArray ignorant of particular representation of nat
 haftmann parents: 
50138diff
changeset | 25 | [simp]: "of_fun f n = IArray (map f [0..<n])" | 
| 
84b03c49c223
IArray ignorant of particular representation of nat
 haftmann parents: 
50138diff
changeset | 26 | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 27 | qualified definition sub :: "'a iarray \<Rightarrow> nat \<Rightarrow> 'a" (infixl "!!" 100) where | 
| 51094 
84b03c49c223
IArray ignorant of particular representation of nat
 haftmann parents: 
50138diff
changeset | 28 | [simp]: "as !! n = IArray.list_of as ! n" | 
| 
84b03c49c223
IArray ignorant of particular representation of nat
 haftmann parents: 
50138diff
changeset | 29 | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 30 | qualified definition length :: "'a iarray \<Rightarrow> nat" where | 
| 51094 
84b03c49c223
IArray ignorant of particular representation of nat
 haftmann parents: 
50138diff
changeset | 31 | [simp]: "length as = List.length (IArray.list_of as)" | 
| 
84b03c49c223
IArray ignorant of particular representation of nat
 haftmann parents: 
50138diff
changeset | 32 | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 33 | qualified fun all :: "('a \<Rightarrow> bool) \<Rightarrow> 'a iarray \<Rightarrow> bool" where
 | 
| 54459 | 34 | "all p (IArray as) = (ALL a : set as. p a)" | 
| 35 | ||
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 36 | qualified fun exists :: "('a \<Rightarrow> bool) \<Rightarrow> 'a iarray \<Rightarrow> bool" where
 | 
| 54459 | 37 | "exists p (IArray as) = (EX a : set as. p a)" | 
| 38 | ||
| 51094 
84b03c49c223
IArray ignorant of particular representation of nat
 haftmann parents: 
50138diff
changeset | 39 | lemma list_of_code [code]: | 
| 
84b03c49c223
IArray ignorant of particular representation of nat
 haftmann parents: 
50138diff
changeset | 40 | "IArray.list_of as = map (\<lambda>n. as !! n) [0 ..< IArray.length as]" | 
| 
84b03c49c223
IArray ignorant of particular representation of nat
 haftmann parents: 
50138diff
changeset | 41 | by (cases as) (simp add: map_nth) | 
| 
84b03c49c223
IArray ignorant of particular representation of nat
 haftmann parents: 
50138diff
changeset | 42 | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 43 | end | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 44 | |
| 50138 | 45 | |
| 46 | subsection "Code Generation" | |
| 47 | ||
| 51094 
84b03c49c223
IArray ignorant of particular representation of nat
 haftmann parents: 
50138diff
changeset | 48 | code_reserved SML Vector | 
| 
84b03c49c223
IArray ignorant of particular representation of nat
 haftmann parents: 
50138diff
changeset | 49 | |
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51161diff
changeset | 50 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51161diff
changeset | 51 | type_constructor iarray \<rightharpoonup> (SML) "_ Vector.vector" | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51161diff
changeset | 52 | | constant IArray \<rightharpoonup> (SML) "Vector.fromList" | 
| 54459 | 53 | | constant IArray.all \<rightharpoonup> (SML) "Vector.all" | 
| 54 | | constant IArray.exists \<rightharpoonup> (SML) "Vector.exists" | |
| 51094 
84b03c49c223
IArray ignorant of particular representation of nat
 haftmann parents: 
50138diff
changeset | 55 | |
| 51161 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
51143diff
changeset | 56 | lemma [code]: | 
| 58269 | 57 | "size (as :: 'a iarray) = Suc (length (IArray.list_of as))" | 
| 58 | by (cases as) simp | |
| 51161 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
51143diff
changeset | 59 | |
| 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
51143diff
changeset | 60 | lemma [code]: | 
| 58269 | 61 | "size_iarray f as = Suc (size_list f (IArray.list_of as))" | 
| 62 | by (cases as) simp | |
| 63 | ||
| 64 | lemma [code]: | |
| 65 | "rec_iarray f as = f (IArray.list_of as)" | |
| 66 | by (cases as) simp | |
| 67 | ||
| 68 | lemma [code]: | |
| 69 | "case_iarray f as = f (IArray.list_of as)" | |
| 70 | by (cases as) simp | |
| 51161 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
51143diff
changeset | 71 | |
| 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
51143diff
changeset | 72 | lemma [code]: | 
| 58269 | 73 | "set_iarray as = set (IArray.list_of as)" | 
| 74 | by (case_tac as) auto | |
| 75 | ||
| 76 | lemma [code]: | |
| 77 | "map_iarray f as = IArray (map f (IArray.list_of as))" | |
| 78 | by (case_tac as) auto | |
| 51161 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
51143diff
changeset | 79 | |
| 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
51143diff
changeset | 80 | lemma [code]: | 
| 58269 | 81 | "rel_iarray r as bs = list_all2 r (IArray.list_of as) (IArray.list_of bs)" | 
| 82 | by (case_tac as) (case_tac bs, auto) | |
| 51161 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
51143diff
changeset | 83 | |
| 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
51143diff
changeset | 84 | lemma [code]: | 
| 58269 | 85 | "HOL.equal as bs \<longleftrightarrow> HOL.equal (IArray.list_of as) (IArray.list_of bs)" | 
| 86 | by (cases as, cases bs) (simp add: equal) | |
| 51161 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
51143diff
changeset | 87 | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 88 | context | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 89 | begin | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 90 | |
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 91 | qualified primrec tabulate :: "integer \<times> (integer \<Rightarrow> 'a) \<Rightarrow> 'a iarray" where | 
| 58269 | 92 | "tabulate (n, f) = IArray (map (f \<circ> integer_of_nat) [0..<nat_of_integer n])" | 
| 93 | ||
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 94 | end | 
| 51094 
84b03c49c223
IArray ignorant of particular representation of nat
 haftmann parents: 
50138diff
changeset | 95 | |
| 
84b03c49c223
IArray ignorant of particular representation of nat
 haftmann parents: 
50138diff
changeset | 96 | lemma [code]: | 
| 58269 | 97 | "IArray.of_fun f n = IArray.tabulate (integer_of_nat n, f \<circ> nat_of_integer)" | 
| 98 | by simp | |
| 51094 
84b03c49c223
IArray ignorant of particular representation of nat
 haftmann parents: 
50138diff
changeset | 99 | |
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51161diff
changeset | 100 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51161diff
changeset | 101 | constant IArray.tabulate \<rightharpoonup> (SML) "Vector.tabulate" | 
| 51094 
84b03c49c223
IArray ignorant of particular representation of nat
 haftmann parents: 
50138diff
changeset | 102 | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 103 | context | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 104 | begin | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 105 | |
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 106 | qualified primrec sub' :: "'a iarray \<times> integer \<Rightarrow> 'a" where | 
| 58269 | 107 | [code del]: "sub' (as, n) = IArray.list_of as ! nat_of_integer n" | 
| 108 | ||
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 109 | end | 
| 51094 
84b03c49c223
IArray ignorant of particular representation of nat
 haftmann parents: 
50138diff
changeset | 110 | |
| 
84b03c49c223
IArray ignorant of particular representation of nat
 haftmann parents: 
50138diff
changeset | 111 | lemma [code]: | 
| 57117 | 112 | "IArray.sub' (IArray as, n) = as ! nat_of_integer n" | 
| 113 | by simp | |
| 114 | ||
| 115 | lemma [code]: | |
| 58269 | 116 | "as !! n = IArray.sub' (as, integer_of_nat n)" | 
| 117 | by simp | |
| 51094 
84b03c49c223
IArray ignorant of particular representation of nat
 haftmann parents: 
50138diff
changeset | 118 | |
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51161diff
changeset | 119 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51161diff
changeset | 120 | constant IArray.sub' \<rightharpoonup> (SML) "Vector.sub" | 
| 51094 
84b03c49c223
IArray ignorant of particular representation of nat
 haftmann parents: 
50138diff
changeset | 121 | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 122 | context | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 123 | begin | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 124 | |
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 125 | qualified definition length' :: "'a iarray \<Rightarrow> integer" where | 
| 58269 | 126 | [code del, simp]: "length' as = integer_of_nat (List.length (IArray.list_of as))" | 
| 127 | ||
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 128 | end | 
| 51094 
84b03c49c223
IArray ignorant of particular representation of nat
 haftmann parents: 
50138diff
changeset | 129 | |
| 
84b03c49c223
IArray ignorant of particular representation of nat
 haftmann parents: 
50138diff
changeset | 130 | lemma [code]: | 
| 57117 | 131 | "IArray.length' (IArray as) = integer_of_nat (List.length as)" | 
| 132 | by simp | |
| 133 | ||
| 134 | lemma [code]: | |
| 58269 | 135 | "IArray.length as = nat_of_integer (IArray.length' as)" | 
| 136 | by simp | |
| 51094 
84b03c49c223
IArray ignorant of particular representation of nat
 haftmann parents: 
50138diff
changeset | 137 | |
| 59457 | 138 | context term_syntax | 
| 139 | begin | |
| 140 | ||
| 141 | lemma [code]: | |
| 142 | "Code_Evaluation.term_of (as :: 'a::typerep iarray) = | |
| 143 |     Code_Evaluation.Const (STR ''IArray.iarray.IArray'') (TYPEREP('a list \<Rightarrow> 'a iarray)) <\<cdot>> (Code_Evaluation.term_of (IArray.list_of as))"
 | |
| 144 | by (subst term_of_anything) rule | |
| 145 | ||
| 146 | end | |
| 147 | ||
| 52435 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51161diff
changeset | 148 | code_printing | 
| 
6646bb548c6b
migration from code_(const|type|class|instance) to code_printing and from code_module to code_identifier
 haftmann parents: 
51161diff
changeset | 149 | constant IArray.length' \<rightharpoonup> (SML) "Vector.length" | 
| 50138 | 150 | |
| 151 | end |