| 
5069
 | 
     1  | 
Goal
  | 
| 
1363
 | 
     2  | 
"{x=X} \
 | 
| 
 | 
     3  | 
\ y := []; \
  | 
| 
 | 
     4  | 
\ WHILE x ~= [] \
  | 
| 
 | 
     5  | 
\ DO { rev(x)@y = rev(X)} \
 | 
| 
 | 
     6  | 
\    y := hd x # y; x := tl x \
  | 
| 
 | 
     7  | 
\ END \
  | 
| 
 | 
     8  | 
\{y=rev(X)}";
 | 
| 
2031
 | 
     9  | 
by (hoare_tac 1);
  | 
| 
4089
 | 
    10  | 
by (asm_full_simp_tac (simpset() addsimps [neq_Nil_conv]) 1);
  | 
| 
4153
 | 
    11  | 
by Safe_tac;
  | 
| 
2031
 | 
    12  | 
by (Asm_full_simp_tac 1);
  | 
| 
1363
 | 
    13  | 
qed "imperative_reverse";
  | 
| 
 | 
    14  | 
  | 
| 
5069
 | 
    15  | 
Goal
  | 
| 
1363
 | 
    16  | 
"{x=X & y = Y} \
 | 
| 
 | 
    17  | 
\ x := rev(x); \
  | 
| 
 | 
    18  | 
\ WHILE x ~= [] \
  | 
| 
 | 
    19  | 
\ DO { rev(x)@y = X@Y} \
 | 
| 
 | 
    20  | 
\    y := hd x # y; x := tl x \
  | 
| 
 | 
    21  | 
\ END \
  | 
| 
 | 
    22  | 
\{y = X@Y}";
 | 
| 
2031
 | 
    23  | 
by (hoare_tac 1);
  | 
| 
4089
 | 
    24  | 
by (asm_full_simp_tac (simpset() addsimps [neq_Nil_conv]) 1);
  | 
| 
4153
 | 
    25  | 
by Safe_tac;
  | 
| 
2031
 | 
    26  | 
by (Asm_full_simp_tac 1);
  | 
| 
1363
 | 
    27  | 
qed "imperative_append";
  |