src/HOL/ex/IArray_Examples.thy
author nipkow
Sat, 16 Nov 2013 18:26:57 +0100
changeset 54459 f43ae1afd08a
parent 51143 0a2371e7ced3
child 66453 cc19f7ca2ed6
permissions -rw-r--r--
added functions all and exists to IArray
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
51143
0a2371e7ced3 two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents: 51094
diff changeset
     2
imports "~~/src/HOL/Library/IArray" "~~/src/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
ca989d793b34 new theory of immutable arrays
nipkow
parents:
diff changeset
    32
end
51094
84b03c49c223 IArray ignorant of particular representation of nat
haftmann
parents: 50138
diff changeset
    33