| author | desharna | 
| Mon, 19 Dec 2022 08:01:31 +0100 | |
| changeset 76684 | 3eda063a20a4 | 
| parent 75936 | d2e6a1342c90 | 
| child 80914 | d97fdabd9e2b | 
| permissions | -rw-r--r-- | 
| 68657 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 1 | (* Title: HOL/Library/IArray.thy | 
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 2 | Author: Tobias Nipkow, TU Muenchen | 
| 68659 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 3 | Author: Jose Divasón <jose.divasonm at unirioja.es> | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 4 | Author: Jesús Aransay <jesus-maria.aransay at unirioja.es> | 
| 68657 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 5 | *) | 
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 6 | |
| 68654 | 7 | section \<open>Immutable Arrays with Code Generation\<close> | 
| 50138 | 8 | |
| 9 | theory IArray | |
| 51094 
84b03c49c223
IArray ignorant of particular representation of nat
 haftmann parents: 
50138diff
changeset | 10 | imports Main | 
| 50138 | 11 | begin | 
| 12 | ||
| 68657 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 13 | subsection \<open>Fundamental operations\<close> | 
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 14 | |
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 15 | text \<open>Immutable arrays are lists wrapped up in an additional constructor. | 
| 50138 | 16 | There are no update operations. Hence code generation can safely implement | 
| 17 | this type by efficient target language arrays. Currently only SML is | |
| 68657 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 18 | provided. Could be extended to other target languages and more operations.\<close> | 
| 50138 | 19 | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 20 | context | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 21 | begin | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 22 | |
| 58310 | 23 | datatype 'a iarray = IArray "'a list" | 
| 50138 | 24 | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 25 | qualified primrec list_of :: "'a iarray \<Rightarrow> 'a list" where | 
| 50138 | 26 | "list_of (IArray xs) = xs" | 
| 27 | ||
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 28 | 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 | 29 | [simp]: "of_fun f n = IArray (map f [0..<n])" | 
| 
84b03c49c223
IArray ignorant of particular representation of nat
 haftmann parents: 
50138diff
changeset | 30 | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 31 | 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 | 32 | [simp]: "as !! n = IArray.list_of as ! n" | 
| 
84b03c49c223
IArray ignorant of particular representation of nat
 haftmann parents: 
50138diff
changeset | 33 | |
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 34 | qualified definition length :: "'a iarray \<Rightarrow> nat" where | 
| 51094 
84b03c49c223
IArray ignorant of particular representation of nat
 haftmann parents: 
50138diff
changeset | 35 | [simp]: "length as = List.length (IArray.list_of as)" | 
| 
84b03c49c223
IArray ignorant of particular representation of nat
 haftmann parents: 
50138diff
changeset | 36 | |
| 68659 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 37 | qualified definition all :: "('a \<Rightarrow> bool) \<Rightarrow> 'a iarray \<Rightarrow> bool" where
 | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 38 | [simp]: "all p as \<longleftrightarrow> (\<forall>a \<in> set (list_of as). p a)" | 
| 54459 | 39 | |
| 68659 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 40 | qualified definition exists :: "('a \<Rightarrow> bool) \<Rightarrow> 'a iarray \<Rightarrow> bool" where
 | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 41 | [simp]: "exists p as \<longleftrightarrow> (\<exists>a \<in> set (list_of as). p a)" | 
| 51094 
84b03c49c223
IArray ignorant of particular representation of nat
 haftmann parents: 
50138diff
changeset | 42 | |
| 68655 | 43 | lemma of_fun_nth: | 
| 44 | "IArray.of_fun f n !! i = f i" if "i < n" | |
| 45 | using that by (simp add: map_nth) | |
| 46 | ||
| 61115 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 47 | end | 
| 
3a4400985780
modernized name space management -- more uniform qualification;
 wenzelm parents: 
60500diff
changeset | 48 | |
| 50138 | 49 | |
| 68657 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 50 | subsection \<open>Generic code equations\<close> | 
| 51094 
84b03c49c223
IArray ignorant of particular representation of nat
 haftmann parents: 
50138diff
changeset | 51 | |
| 51161 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
51143diff
changeset | 52 | lemma [code]: | 
| 68658 | 53 | "size (as :: 'a iarray) = Suc (IArray.length as)" | 
| 58269 | 54 | 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 | 55 | |
| 
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_iarray f as = Suc (size_list f (IArray.list_of as))" | 
| 58 | by (cases as) simp | |
| 59 | ||
| 60 | lemma [code]: | |
| 61 | "rec_iarray f as = f (IArray.list_of as)" | |
| 62 | by (cases as) simp | |
| 63 | ||
| 64 | lemma [code]: | |
| 65 | "case_iarray f as = f (IArray.list_of as)" | |
| 66 | 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 | 67 | |
| 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
51143diff
changeset | 68 | lemma [code]: | 
| 58269 | 69 | "set_iarray as = set (IArray.list_of as)" | 
| 63649 | 70 | by (cases as) auto | 
| 58269 | 71 | |
| 72 | lemma [code]: | |
| 73 | "map_iarray f as = IArray (map f (IArray.list_of as))" | |
| 63649 | 74 | by (cases as) auto | 
| 51161 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
51143diff
changeset | 75 | |
| 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
51143diff
changeset | 76 | lemma [code]: | 
| 58269 | 77 | "rel_iarray r as bs = list_all2 r (IArray.list_of as) (IArray.list_of bs)" | 
| 63649 | 78 | by (cases as, cases bs) auto | 
| 51161 
6ed12ae3b3e1
attempt to re-establish conventions which theories are loaded into the grand unified library theory;
 haftmann parents: 
51143diff
changeset | 79 | |
| 68659 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 80 | lemma list_of_code [code]: | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 81 | "IArray.list_of as = map (\<lambda>n. as !! n) [0 ..< IArray.length as]" | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 82 | by (cases as) (simp add: map_nth) | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 83 | |
| 51161 
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 | |
| 68659 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 88 | lemma [code]: | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 89 | "IArray.all p = Not \<circ> IArray.exists (Not \<circ> p)" | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 90 | by (simp add: fun_eq_iff) | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 91 | |
| 72607 | 92 | context | 
| 93 | includes term_syntax | |
| 59457 | 94 | begin | 
| 95 | ||
| 96 | lemma [code]: | |
| 97 | "Code_Evaluation.term_of (as :: 'a::typerep iarray) = | |
| 98 |     Code_Evaluation.Const (STR ''IArray.iarray.IArray'') (TYPEREP('a list \<Rightarrow> 'a iarray)) <\<cdot>> (Code_Evaluation.term_of (IArray.list_of as))"
 | |
| 99 | by (subst term_of_anything) rule | |
| 100 | ||
| 101 | end | |
| 102 | ||
| 68657 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 103 | |
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 104 | subsection \<open>Auxiliary operations for code generation\<close> | 
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 105 | |
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 106 | context | 
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 107 | begin | 
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 108 | |
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 109 | qualified primrec tabulate :: "integer \<times> (integer \<Rightarrow> 'a) \<Rightarrow> 'a iarray" where | 
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 110 | "tabulate (n, f) = IArray (map (f \<circ> integer_of_nat) [0..<nat_of_integer n])" | 
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 111 | |
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 112 | lemma [code]: | 
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 113 | "IArray.of_fun f n = IArray.tabulate (integer_of_nat n, f \<circ> nat_of_integer)" | 
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 114 | by simp | 
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 115 | |
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 116 | qualified primrec sub' :: "'a iarray \<times> integer \<Rightarrow> 'a" where | 
| 68659 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 117 | "sub' (as, n) = as !! nat_of_integer n" | 
| 68657 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 118 | |
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 119 | lemma [code]: | 
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 120 | "IArray.sub' (IArray as, n) = as ! nat_of_integer n" | 
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 121 | by simp | 
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 122 | |
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 123 | lemma [code]: | 
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 124 | "as !! n = IArray.sub' (as, integer_of_nat n)" | 
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 125 | by simp | 
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 126 | |
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 127 | qualified definition length' :: "'a iarray \<Rightarrow> integer" where | 
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 128 | [simp]: "length' as = integer_of_nat (List.length (IArray.list_of as))" | 
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 129 | |
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 130 | lemma [code]: | 
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 131 | "IArray.length' (IArray as) = integer_of_nat (List.length as)" | 
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 132 | by simp | 
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 133 | |
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 134 | lemma [code]: | 
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 135 | "IArray.length as = nat_of_integer (IArray.length' as)" | 
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 136 | by simp | 
| 50138 | 137 | |
| 68659 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 138 | qualified definition exists_upto :: "('a \<Rightarrow> bool) \<Rightarrow> integer \<Rightarrow> 'a iarray \<Rightarrow> bool" where
 | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 139 | [simp]: "exists_upto p k as \<longleftrightarrow> (\<exists>l. 0 \<le> l \<and> l < k \<and> p (sub' (as, l)))" | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 140 | |
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 141 | lemma exists_upto_of_nat: | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 142 | "exists_upto p (of_nat n) as \<longleftrightarrow> (\<exists>m<n. p (as !! m))" | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 143 | including integer.lifting by (simp, transfer) | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 144 | (metis nat_int nat_less_iff of_nat_0_le_iff) | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 145 | |
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 146 | lemma [code]: | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 147 | "exists_upto p k as \<longleftrightarrow> (if k \<le> 0 then False else | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 148 | let l = k - 1 in p (sub' (as, l)) \<or> exists_upto p l as)" | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 149 | proof (cases "k \<ge> 1") | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 150 | case False | 
| 75936 | 151 | then have \<open>k \<le> 0\<close> | 
| 152 | including integer.lifting by transfer simp | |
| 68659 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 153 | then show ?thesis | 
| 75936 | 154 | by simp | 
| 68659 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 155 | next | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 156 | case True | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 157 | then have less: "k \<le> 0 \<longleftrightarrow> False" | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 158 | by simp | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 159 | define n where "n = nat_of_integer (k - 1)" | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 160 | with True have k: "k - 1 = of_nat n" "k = of_nat (Suc n)" | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 161 | by simp_all | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 162 | show ?thesis unfolding less Let_def k(1) unfolding k(2) exists_upto_of_nat | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 163 | using less_Suc_eq by auto | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 164 | qed | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 165 | |
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 166 | lemma [code]: | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 167 | "IArray.exists p as \<longleftrightarrow> exists_upto p (length' as) as" | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 168 | including integer.lifting by (simp, transfer) | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 169 | (auto, metis in_set_conv_nth less_imp_of_nat_less nat_int of_nat_0_le_iff) | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 170 | |
| 50138 | 171 | end | 
| 68657 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 172 | |
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 173 | |
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 174 | subsection \<open>Code Generation for SML\<close> | 
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 175 | |
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 176 | text \<open>Note that arrays cannot be printed directly but only by turning them into | 
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 177 | lists first. Arrays could be converted back into lists for printing if they | 
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 178 | were wrapped up in an additional constructor.\<close> | 
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 179 | |
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 180 | code_reserved SML Vector | 
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 181 | |
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 182 | code_printing | 
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 183 | type_constructor iarray \<rightharpoonup> (SML) "_ Vector.vector" | 
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 184 | | constant IArray \<rightharpoonup> (SML) "Vector.fromList" | 
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 185 | | constant IArray.all \<rightharpoonup> (SML) "Vector.all" | 
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 186 | | constant IArray.exists \<rightharpoonup> (SML) "Vector.exists" | 
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 187 | | constant IArray.tabulate \<rightharpoonup> (SML) "Vector.tabulate" | 
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 188 | | constant IArray.sub' \<rightharpoonup> (SML) "Vector.sub" | 
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 189 | | constant IArray.length' \<rightharpoonup> (SML) "Vector.length" | 
| 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 190 | |
| 68659 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 191 | |
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 192 | subsection \<open>Code Generation for Haskell\<close> | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 193 | |
| 69593 | 194 | text \<open>We map \<^typ>\<open>'a iarray\<close>s in Isabelle/HOL to \<open>Data.Array.IArray.array\<close> | 
| 69272 | 195 | in Haskell. Performance mapping to \<open>Data.Array.Unboxed.Array\<close> and | 
| 196 | \<open>Data.Array.Array\<close> is similar.\<close> | |
| 68659 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 197 | |
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 198 | code_printing | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 199 | code_module "IArray" \<rightharpoonup> (Haskell) \<open> | 
| 69690 | 200 | module IArray(IArray, tabulate, of_list, sub, length) where {
 | 
| 201 | ||
| 68659 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 202 | import Prelude (Bool(True, False), not, Maybe(Nothing, Just), | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 203 | Integer, (+), (-), (<), fromInteger, toInteger, map, seq, (.)); | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 204 | import qualified Prelude; | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 205 | import qualified Data.Array.IArray; | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 206 | import qualified Data.Array.Base; | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 207 | import qualified Data.Ix; | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 208 | |
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 209 | newtype IArray e = IArray (Data.Array.IArray.Array Integer e); | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 210 | |
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 211 | tabulate :: (Integer, (Integer -> e)) -> IArray e; | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 212 | tabulate (k, f) = IArray (Data.Array.IArray.array (0, k - 1) (map (\i -> let fi = f i in fi `seq` (i, fi)) [0..k - 1])); | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 213 | |
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 214 | of_list :: [e] -> IArray e; | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 215 | of_list l = IArray (Data.Array.IArray.listArray (0, (toInteger . Prelude.length) l - 1) l); | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 216 | |
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 217 | sub :: (IArray e, Integer) -> e; | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 218 | sub (IArray v, i) = v `Data.Array.Base.unsafeAt` fromInteger i; | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 219 | |
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 220 | length :: IArray e -> Integer; | 
| 69690 | 221 | length (IArray v) = toInteger (Data.Ix.rangeSize (Data.Array.IArray.bounds v)); | 
| 222 | ||
| 223 | }\<close> for type_constructor iarray constant IArray IArray.tabulate IArray.sub' IArray.length' | |
| 68659 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 224 | |
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 225 | code_reserved Haskell IArray_Impl | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 226 | |
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 227 | code_printing | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 228 | type_constructor iarray \<rightharpoonup> (Haskell) "IArray.IArray _" | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 229 | | constant IArray \<rightharpoonup> (Haskell) "IArray.of'_list" | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 230 | | constant IArray.tabulate \<rightharpoonup> (Haskell) "IArray.tabulate" | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 231 | | constant IArray.sub' \<rightharpoonup> (Haskell) "IArray.sub" | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 232 | | constant IArray.length' \<rightharpoonup> (Haskell) "IArray.length" | 
| 
db0c70767d86
setup for Haskell taken over from AFP / Gauss_Jordan
 haftmann parents: 
68658diff
changeset | 233 | |
| 68657 
65ad2bfc19d2
restructured for future incorporation of Haskell
 haftmann parents: 
68656diff
changeset | 234 | end |