src/HOL/ex/Subarray.thy
author haftmann
Thu, 08 Jan 2009 17:10:41 +0100
changeset 29399 ebcd69a00872
parent 27656 src/HOL/Library/Subarray.thy@d4f6e64ee7cc
permissions -rw-r--r--
split of Imperative_HOL theories from HOL-Library
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
27656
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
     1
theory Subarray
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
     2
imports Array Sublist
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
     3
begin
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
     4
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
     5
definition subarray :: "nat \<Rightarrow> nat \<Rightarrow> ('a::heap) array \<Rightarrow> heap \<Rightarrow> 'a list"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
     6
where
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
     7
  "subarray n m a h \<equiv> sublist' n m (get_array a h)"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
     8
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
     9
lemma subarray_upd: "i \<ge> m \<Longrightarrow> subarray n m a (Heap.upd a i v h) = subarray n m a h"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    10
apply (simp add: subarray_def Heap.upd_def)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    11
apply (simp add: sublist'_update1)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    12
done
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    13
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    14
lemma subarray_upd2: " i < n  \<Longrightarrow> subarray n m a (Heap.upd a i v h) = subarray n m a h"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    15
apply (simp add: subarray_def Heap.upd_def)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    16
apply (subst sublist'_update2)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    17
apply fastsimp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    18
apply simp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    19
done
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    20
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    21
lemma subarray_upd3: "\<lbrakk> n \<le> i; i < m\<rbrakk> \<Longrightarrow> subarray n m a (Heap.upd a i v h) = subarray n m a h[i - n := v]"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    22
unfolding subarray_def Heap.upd_def
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    23
by (simp add: sublist'_update3)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    24
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    25
lemma subarray_Nil: "n \<ge> m \<Longrightarrow> subarray n m a h = []"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    26
by (simp add: subarray_def sublist'_Nil')
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    27
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    28
lemma subarray_single: "\<lbrakk> n < Heap.length a h \<rbrakk> \<Longrightarrow> subarray n (Suc n) a h = [get_array a h ! n]" 
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    29
by (simp add: subarray_def Heap.length_def sublist'_single)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    30
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    31
lemma length_subarray: "m \<le> Heap.length a h \<Longrightarrow> List.length (subarray n m a h) = m - n"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    32
by (simp add: subarray_def Heap.length_def length_sublist')
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    33
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    34
lemma length_subarray_0: "m \<le> Heap.length a h \<Longrightarrow> List.length (subarray 0 m a h) = m"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    35
by (simp add: length_subarray)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    36
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    37
lemma subarray_nth_array_Cons: "\<lbrakk> i < Heap.length a h; i < j \<rbrakk> \<Longrightarrow> (get_array a h ! i) # subarray (Suc i) j a h = subarray i j a h"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    38
unfolding Heap.length_def subarray_def
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    39
by (simp add: sublist'_front)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    40
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    41
lemma subarray_nth_array_back: "\<lbrakk> i < j; j \<le> Heap.length a h\<rbrakk> \<Longrightarrow> subarray i j a h = subarray i (j - 1) a h @ [get_array a h ! (j - 1)]"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    42
unfolding Heap.length_def subarray_def
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    43
by (simp add: sublist'_back)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    44
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    45
lemma subarray_append: "\<lbrakk> i < j; j < k \<rbrakk> \<Longrightarrow> subarray i j a h @ subarray j k a h = subarray i k a h"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    46
unfolding subarray_def
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    47
by (simp add: sublist'_append)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    48
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    49
lemma subarray_all: "subarray 0 (Heap.length a h) a h = get_array a h"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    50
unfolding Heap.length_def subarray_def
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    51
by (simp add: sublist'_all)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    52
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    53
lemma nth_subarray: "\<lbrakk> k < j - i; j \<le> Heap.length a h \<rbrakk> \<Longrightarrow> subarray i j a h ! k = get_array a h ! (i + k)"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    54
unfolding Heap.length_def subarray_def
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    55
by (simp add: nth_sublist')
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    56
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    57
lemma subarray_eq_samelength_iff: "Heap.length a h = Heap.length a h' \<Longrightarrow> (subarray i j a h = subarray i j a h') = (\<forall>i'. i \<le> i' \<and> i' < j \<longrightarrow> get_array a h ! i' = get_array a h' ! i')"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    58
unfolding Heap.length_def subarray_def by (rule sublist'_eq_samelength_iff)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    59
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    60
lemma all_in_set_subarray_conv: "(\<forall>j. j \<in> set (subarray l r a h) \<longrightarrow> P j) = (\<forall>k. l \<le> k \<and> k < r \<and> k < Heap.length a h \<longrightarrow> P (get_array a h ! k))"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    61
unfolding subarray_def Heap.length_def by (rule all_in_set_sublist'_conv)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    62
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    63
lemma ball_in_set_subarray_conv: "(\<forall>j \<in> set (subarray l r a h). P j) = (\<forall>k. l \<le> k \<and> k < r \<and> k < Heap.length a h \<longrightarrow> P (get_array a h ! k))"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    64
unfolding subarray_def Heap.length_def by (rule ball_in_set_sublist'_conv)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    65
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    66
end