src/HOL/ex/IArray_Examples.thy
author wenzelm
Fri, 16 Sep 2022 22:33:14 +0200
changeset 76178 1f95e9424341
parent 69908 1bd74a0944b3
permissions -rw-r--r--
more robust: snap version of docker cannot access /tmp;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
50138
ca989d793b34 new theory of immutable arrays
nipkow
parents:
diff changeset
     1
theory IArray_Examples
66453
cc19f7ca2ed6 session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents: 54459
diff changeset
     2
imports "HOL-Library.IArray" "HOL-Library.Code_Target_Numeral"
50138
ca989d793b34 new theory of immutable arrays
nipkow
parents:
diff changeset
     3
begin
ca989d793b34 new theory of immutable arrays
nipkow
parents:
diff changeset
     4
ca989d793b34 new theory of immutable arrays
nipkow
parents:
diff changeset
     5
lemma "IArray [True,False] !! 1 = False"
ca989d793b34 new theory of immutable arrays
nipkow
parents:
diff changeset
     6
by eval
ca989d793b34 new theory of immutable arrays
nipkow
parents:
diff changeset
     7
ca989d793b34 new theory of immutable arrays
nipkow
parents:
diff changeset
     8
lemma "IArray.length (IArray [[],[]]) = 2"
ca989d793b34 new theory of immutable arrays
nipkow
parents:
diff changeset
     9
by eval
ca989d793b34 new theory of immutable arrays
nipkow
parents:
diff changeset
    10
ca989d793b34 new theory of immutable arrays
nipkow
parents:
diff changeset
    11
lemma "IArray.list_of (IArray [1,3::int]) = [1,3]"
ca989d793b34 new theory of immutable arrays
nipkow
parents:
diff changeset
    12
by eval
ca989d793b34 new theory of immutable arrays
nipkow
parents:
diff changeset
    13
ca989d793b34 new theory of immutable arrays
nipkow
parents:
diff changeset
    14
lemma "IArray.list_of (IArray.of_fun (%n. n*n) 5) = [0,1,4,9,16]"
ca989d793b34 new theory of immutable arrays
nipkow
parents:
diff changeset
    15
by eval
ca989d793b34 new theory of immutable arrays
nipkow
parents:
diff changeset
    16
54459
f43ae1afd08a added functions all and exists to IArray
nipkow
parents: 51143
diff changeset
    17
lemma "\<not> IArray.all (\<lambda>x. x > 2) (IArray [1,3::int])"
f43ae1afd08a added functions all and exists to IArray
nipkow
parents: 51143
diff changeset
    18
by eval
f43ae1afd08a added functions all and exists to IArray
nipkow
parents: 51143
diff changeset
    19
f43ae1afd08a added functions all and exists to IArray
nipkow
parents: 51143
diff changeset
    20
lemma "IArray.exists (\<lambda>x. x > 2) (IArray [1,3::int])"
f43ae1afd08a added functions all and exists to IArray
nipkow
parents: 51143
diff changeset
    21
by eval
f43ae1afd08a added functions all and exists to IArray
nipkow
parents: 51143
diff changeset
    22
50138
ca989d793b34 new theory of immutable arrays
nipkow
parents:
diff changeset
    23
fun sum2 :: "'a::monoid_add iarray \<Rightarrow> nat \<Rightarrow> 'a \<Rightarrow> 'a" where
ca989d793b34 new theory of immutable arrays
nipkow
parents:
diff changeset
    24
"sum2 A n s = (if n=0 then s else sum2 A (n - 1) (s + A!!(n - 1)))"
ca989d793b34 new theory of immutable arrays
nipkow
parents:
diff changeset
    25
ca989d793b34 new theory of immutable arrays
nipkow
parents:
diff changeset
    26
definition sum :: "'a::monoid_add iarray \<Rightarrow> 'a" where
ca989d793b34 new theory of immutable arrays
nipkow
parents:
diff changeset
    27
"sum A = sum2 A (IArray.length A) 0"
ca989d793b34 new theory of immutable arrays
nipkow
parents:
diff changeset
    28
ca989d793b34 new theory of immutable arrays
nipkow
parents:
diff changeset
    29
lemma "sum (IArray [1,2,3,4,5,6,7,8,9::int]) = 45"
ca989d793b34 new theory of immutable arrays
nipkow
parents:
diff changeset
    30
by eval
ca989d793b34 new theory of immutable arrays
nipkow
parents:
diff changeset
    31
69908
1bd74a0944b3 slightly more complete check of code generation for immutable arrays
haftmann
parents: 66453
diff changeset
    32
definition partial_geometric_sum :: "'a::comm_ring_1 list \<Rightarrow> 'a"
1bd74a0944b3 slightly more complete check of code generation for immutable arrays
haftmann
parents: 66453
diff changeset
    33
  where "partial_geometric_sum xs = (let
1bd74a0944b3 slightly more complete check of code generation for immutable arrays
haftmann
parents: 66453
diff changeset
    34
    values = IArray xs;
1bd74a0944b3 slightly more complete check of code generation for immutable arrays
haftmann
parents: 66453
diff changeset
    35
    coefficients = IArray.of_fun of_nat (length xs)
1bd74a0944b3 slightly more complete check of code generation for immutable arrays
haftmann
parents: 66453
diff changeset
    36
  in sum_list (map (\<lambda>i. values !! i * coefficients !! i) [0..<IArray.length values]))"
1bd74a0944b3 slightly more complete check of code generation for immutable arrays
haftmann
parents: 66453
diff changeset
    37
1bd74a0944b3 slightly more complete check of code generation for immutable arrays
haftmann
parents: 66453
diff changeset
    38
export_code partial_geometric_sum checking SML OCaml? Haskell? Scala
1bd74a0944b3 slightly more complete check of code generation for immutable arrays
haftmann
parents: 66453
diff changeset
    39
50138
ca989d793b34 new theory of immutable arrays
nipkow
parents:
diff changeset
    40
end
51094
84b03c49c223 IArray ignorant of particular representation of nat
haftmann
parents: 50138
diff changeset
    41