author | urbanc |
Tue, 16 May 2006 14:11:39 +0200 | |
changeset 19651 | 247ca17caddd |
parent 17477 | ceb42ea2f223 |
child 19803 | aa2581752afb |
permissions | -rw-r--r-- |
8177 | 1 |
(* Title: HOL/IMPP/EvenOdd.thy |
2 |
ID: $Id$ |
|
3 |
Author: David von Oheimb |
|
4 |
Copyright 1999 TUM |
|
5 |
*) |
|
6 |
||
17477 | 7 |
header {* Example of mutually recursive procedures verified with Hoare logic *} |
8177 | 8 |
|
17477 | 9 |
theory EvenOdd |
10 |
imports Misc |
|
11 |
begin |
|
12 |
||
13 |
constdefs |
|
14 |
even :: "nat => bool" |
|
11704
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
wenzelm
parents:
11701
diff
changeset
|
15 |
"even n == 2 dvd n" |
8177 | 16 |
|
17 |
consts |
|
17477 | 18 |
Even :: pname |
19 |
Odd :: pname |
|
20 |
axioms |
|
21 |
Even_neq_Odd: "Even ~= Odd" |
|
22 |
Arg_neq_Res: "Arg ~= Res" |
|
8177 | 23 |
|
24 |
constdefs |
|
25 |
evn :: com |
|
15354 | 26 |
"evn == IF (%s. s<Arg> = 0) |
8177 | 27 |
THEN Loc Res:==(%s. 0) |
28 |
ELSE(Loc Res:=CALL Odd(%s. s<Arg> - 1);; |
|
29 |
Loc Arg:=CALL Odd(%s. s<Arg> - 1);; |
|
30 |
Loc Res:==(%s. s<Res> * s<Arg>))" |
|
31 |
odd :: com |
|
15354 | 32 |
"odd == IF (%s. s<Arg> = 0) |
8177 | 33 |
THEN Loc Res:==(%s. 1) |
11701
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
wenzelm
parents:
8791
diff
changeset
|
34 |
ELSE(Loc Res:=CALL Even (%s. s<Arg> - 1))" |
8177 | 35 |
|
36 |
defs |
|
17477 | 37 |
bodies_def: "bodies == [(Even,evn),(Odd,odd)]" |
38 |
||
8177 | 39 |
consts |
17477 | 40 |
Z_eq_Arg_plus :: "nat => nat assn" ("Z=Arg+_" [50]50) |
41 |
"even_Z=(Res=0)" :: "nat assn" ("Res'_ok") |
|
8177 | 42 |
defs |
17477 | 43 |
Z_eq_Arg_plus_def: "Z=Arg+n == %Z s. Z = s<Arg>+n" |
44 |
Res_ok_def: "Res_ok == %Z s. even Z = (s<Res> = 0)" |
|
45 |
||
46 |
ML {* use_legacy_bindings (the_context ()) *} |
|
8177 | 47 |
|
48 |
end |