src/HOL/Library/Assert.thy
author wenzelm
Fri, 15 Aug 2008 15:50:52 +0200
changeset 27885 76b51cd0a37c
parent 27656 d4f6e64ee7cc
permissions -rw-r--r--
renamed T.source_of' to T.source_position_of;
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 Assert
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
     2
imports Heap_Monad
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
section {* The Assert command *}
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
     6
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
     7
text {* We define a command Assert a property P.
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
     8
 This property does not consider any statement about the heap but only about functional values in the program. *}
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
     9
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    10
definition
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    11
  assert :: "('a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> 'a Heap"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    12
where
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    13
  "assert P x = (if P x then return x else raise (''assert''))"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    14
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    15
lemma assert_cong[fundef_cong]:
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    16
assumes "P = P'"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    17
assumes "\<And>x. P' x \<Longrightarrow> f x = f' x"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    18
shows "(assert P x >>= f) = (assert P' x >>= f')"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    19
using assms
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    20
by (auto simp add: assert_def return_bind raise_bind)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    21
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    22
section {* Example : Using Assert for showing termination of functions *}
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    23
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    24
function until_zero :: "int \<Rightarrow> int Heap"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    25
where
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    26
  "until_zero a = (if a \<le> 0 then return a else (do x \<leftarrow> return (a - 1); until_zero x done))" 
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    27
by (pat_completeness, auto)
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    28
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    29
termination
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    30
apply (relation "measure (\<lambda>x. nat x)")
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    31
apply simp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    32
apply simp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    33
oops
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    34
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    35
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    36
function until_zero' :: "int \<Rightarrow> int Heap"
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    37
where
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    38
  "until_zero' a = (if a \<le> 0 then return a else (do x \<leftarrow> return (a - 1); y \<leftarrow> assert (\<lambda>x. x < a) x; until_zero' y done))" 
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    39
by (pat_completeness, auto)
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
termination
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    42
apply (relation "measure (\<lambda>x. nat x)")
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    43
apply simp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    44
apply simp
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    45
done
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    46
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    47
hide (open) const until_zero until_zero'
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
text {* Also look at lemmas about assert in Relational theory. *}
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    50
d4f6e64ee7cc added verification framework for the HeapMonad and quicksort as example for this framework
bulwahn
parents:
diff changeset
    51
end