src/HOL/IMP/Hoare_Total_EX.thy
author nipkow
Mon, 11 Oct 2021 21:55:21 +0200
changeset 74500 40f03761f77f
parent 69505 cc2d676d5395
permissions -rw-r--r--
more examples
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
63070
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
     1
(* Author: Tobias Nipkow *)
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
     2
68776
403dd13cf6e9 avoid session qualification because no tex is generated when used;
nipkow
parents: 67406
diff changeset
     3
subsubsection "\<open>nat\<close>-Indexed Invariant"
403dd13cf6e9 avoid session qualification because no tex is generated when used;
nipkow
parents: 67406
diff changeset
     4
63538
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents: 63070
diff changeset
     5
theory Hoare_Total_EX
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents: 63070
diff changeset
     6
imports Hoare
d7b5e2a222c2 added new vcg based on existentially quantified while-rule
nipkow
parents: 63070
diff changeset
     7
begin
63070
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
     8
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67019
diff changeset
     9
text\<open>This is the standard set of rules that you find in many publications.
63070
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    10
The While-rule is different from the one in Concrete Semantics in that the
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    11
invariant is indexed by natural numbers and goes down by 1 with
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    12
every iteration. The completeness proof is easier but the rule is harder
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67019
diff changeset
    13
to apply in program proofs.\<close>
63070
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    14
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    15
definition hoare_tvalid :: "assn \<Rightarrow> com \<Rightarrow> assn \<Rightarrow> bool"
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    16
  ("\<Turnstile>\<^sub>t {(1_)}/ (_)/ {(1_)}" 50) where
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    17
"\<Turnstile>\<^sub>t {P}c{Q}  \<longleftrightarrow>  (\<forall>s. P s \<longrightarrow> (\<exists>t. (c,s) \<Rightarrow> t \<and> Q t))"
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    18
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    19
inductive
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    20
  hoaret :: "assn \<Rightarrow> com \<Rightarrow> assn \<Rightarrow> bool" ("\<turnstile>\<^sub>t ({(1_)}/ (_)/ {(1_)})" 50)
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    21
where
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    22
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    23
Skip:  "\<turnstile>\<^sub>t {P} SKIP {P}"  |
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    24
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    25
Assign:  "\<turnstile>\<^sub>t {\<lambda>s. P(s[a/x])} x::=a {P}"  |
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    26
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    27
Seq: "\<lbrakk> \<turnstile>\<^sub>t {P\<^sub>1} c\<^sub>1 {P\<^sub>2}; \<turnstile>\<^sub>t {P\<^sub>2} c\<^sub>2 {P\<^sub>3} \<rbrakk> \<Longrightarrow> \<turnstile>\<^sub>t {P\<^sub>1} c\<^sub>1;;c\<^sub>2 {P\<^sub>3}"  |
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    28
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    29
If: "\<lbrakk> \<turnstile>\<^sub>t {\<lambda>s. P s \<and> bval b s} c\<^sub>1 {Q}; \<turnstile>\<^sub>t {\<lambda>s. P s \<and> \<not> bval b s} c\<^sub>2 {Q} \<rbrakk>
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    30
  \<Longrightarrow> \<turnstile>\<^sub>t {P} IF b THEN c\<^sub>1 ELSE c\<^sub>2 {Q}"  |
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    31
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    32
While:
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    33
  "\<lbrakk> \<And>n::nat. \<turnstile>\<^sub>t {P (Suc n)} c {P n};
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    34
     \<forall>n s. P (Suc n) s \<longrightarrow> bval b s;  \<forall>s. P 0 s \<longrightarrow> \<not> bval b s \<rbrakk>
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    35
   \<Longrightarrow> \<turnstile>\<^sub>t {\<lambda>s. \<exists>n. P n s} WHILE b DO c {P 0}"  |
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    36
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    37
conseq: "\<lbrakk> \<forall>s. P' s \<longrightarrow> P s; \<turnstile>\<^sub>t {P}c{Q}; \<forall>s. Q s \<longrightarrow> Q' s  \<rbrakk> \<Longrightarrow>
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    38
           \<turnstile>\<^sub>t {P'}c{Q'}"
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    39
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67019
diff changeset
    40
text\<open>Building in the consequence rule:\<close>
63070
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    41
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    42
lemma strengthen_pre:
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    43
  "\<lbrakk> \<forall>s. P' s \<longrightarrow> P s;  \<turnstile>\<^sub>t {P} c {Q} \<rbrakk> \<Longrightarrow> \<turnstile>\<^sub>t {P'} c {Q}"
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    44
by (metis conseq)
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    45
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    46
lemma weaken_post:
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    47
  "\<lbrakk> \<turnstile>\<^sub>t {P} c {Q};  \<forall>s. Q s \<longrightarrow> Q' s \<rbrakk> \<Longrightarrow>  \<turnstile>\<^sub>t {P} c {Q'}"
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    48
by (metis conseq)
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    49
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    50
lemma Assign': "\<forall>s. P s \<longrightarrow> Q(s[a/x]) \<Longrightarrow> \<turnstile>\<^sub>t {P} x ::= a {Q}"
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    51
by (simp add: strengthen_pre[OF _ Assign])
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    52
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67019
diff changeset
    53
text\<open>The soundness theorem:\<close>
63070
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    54
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    55
theorem hoaret_sound: "\<turnstile>\<^sub>t {P}c{Q}  \<Longrightarrow>  \<Turnstile>\<^sub>t {P}c{Q}"
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    56
proof(unfold hoare_tvalid_def, induction rule: hoaret.induct)
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    57
  case (While P c b)
67019
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents: 63538
diff changeset
    58
  have "P n s \<Longrightarrow> \<exists>t. (WHILE b DO c, s) \<Rightarrow> t \<and> P 0 t" for n s
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents: 63538
diff changeset
    59
  proof(induction "n" arbitrary: s)
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents: 63538
diff changeset
    60
    case 0 thus ?case using While.hyps(3) WhileFalse by blast
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents: 63538
diff changeset
    61
  next
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents: 63538
diff changeset
    62
    case Suc
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents: 63538
diff changeset
    63
    thus ?case by (meson While.IH While.hyps(2) WhileTrue)
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents: 63538
diff changeset
    64
  qed
63070
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    65
  thus ?case by auto
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    66
next
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    67
  case If thus ?case by auto blast
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    68
qed fastforce+
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    69
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    70
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    71
definition wpt :: "com \<Rightarrow> assn \<Rightarrow> assn" ("wp\<^sub>t") where
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    72
"wp\<^sub>t c Q  =  (\<lambda>s. \<exists>t. (c,s) \<Rightarrow> t \<and> Q t)"
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    73
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    74
lemma [simp]: "wp\<^sub>t SKIP Q = Q"
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    75
by(auto intro!: ext simp: wpt_def)
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    76
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    77
lemma [simp]: "wp\<^sub>t (x ::= e) Q = (\<lambda>s. Q(s(x := aval e s)))"
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    78
by(auto intro!: ext simp: wpt_def)
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    79
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    80
lemma [simp]: "wp\<^sub>t (c\<^sub>1;;c\<^sub>2) Q = wp\<^sub>t c\<^sub>1 (wp\<^sub>t c\<^sub>2 Q)"
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    81
unfolding wpt_def
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    82
apply(rule ext)
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    83
apply auto
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    84
done
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    85
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    86
lemma [simp]:
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    87
 "wp\<^sub>t (IF b THEN c\<^sub>1 ELSE c\<^sub>2) Q = (\<lambda>s. wp\<^sub>t (if bval b s then c\<^sub>1 else c\<^sub>2) Q s)"
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    88
apply(unfold wpt_def)
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    89
apply(rule ext)
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    90
apply auto
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    91
done
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    92
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    93
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 68776
diff changeset
    94
text\<open>Function \<open>wpw\<close> computes the weakest precondition of a While-loop
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 67019
diff changeset
    95
that is unfolded a fixed number of times.\<close>
63070
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    96
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    97
fun wpw :: "bexp \<Rightarrow> com \<Rightarrow> nat \<Rightarrow> assn \<Rightarrow> assn" where
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    98
"wpw b c 0 Q s = (\<not> bval b s \<and> Q s)" |
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
    99
"wpw b c (Suc n) Q s = (bval b s \<and> (\<exists>s'. (c,s) \<Rightarrow> s' \<and>  wpw b c n Q s'))"
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   100
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   101
lemma WHILE_Its: "(WHILE b DO c,s) \<Rightarrow> t \<Longrightarrow> Q t \<Longrightarrow> \<exists>n. wpw b c n Q s"
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   102
proof(induction "WHILE b DO c" s t rule: big_step_induct)
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   103
  case WhileFalse thus ?case using wpw.simps(1) by blast 
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   104
next
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   105
  case WhileTrue thus ?case using wpw.simps(2) by blast
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   106
qed
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   107
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   108
lemma wpt_is_pre: "\<turnstile>\<^sub>t {wp\<^sub>t c Q} c {Q}"
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   109
proof (induction c arbitrary: Q)
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   110
  case SKIP show ?case by (auto intro:hoaret.Skip)
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   111
next
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   112
  case Assign show ?case by (auto intro:hoaret.Assign)
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   113
next
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   114
  case Seq thus ?case by (auto intro:hoaret.Seq)
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   115
next
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   116
  case If thus ?case by (auto intro:hoaret.If hoaret.conseq)
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   117
next
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   118
  case (While b c)
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   119
  let ?w = "WHILE b DO c"
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   120
  have c1: "\<forall>s. wp\<^sub>t ?w Q s \<longrightarrow> (\<exists>n. wpw b c n Q s)"
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   121
    unfolding wpt_def by (metis WHILE_Its)
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   122
  have c3: "\<forall>s. wpw b c 0 Q s \<longrightarrow> Q s" by simp
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   123
  have w2: "\<forall>n s. wpw b c (Suc n) Q s \<longrightarrow> bval b s" by simp
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   124
  have w3: "\<forall>s. wpw b c 0 Q s \<longrightarrow> \<not> bval b s" by simp
67019
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents: 63538
diff changeset
   125
  have "\<turnstile>\<^sub>t {wpw b c (Suc n) Q} c {wpw b c n Q}" for n
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents: 63538
diff changeset
   126
  proof -
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents: 63538
diff changeset
   127
    have *: "\<forall>s. wpw b c (Suc n) Q s \<longrightarrow> (\<exists>t. (c, s) \<Rightarrow> t \<and> wpw b c n Q t)" by simp
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents: 63538
diff changeset
   128
    show ?thesis by(rule strengthen_pre[OF * While.IH[of "wpw b c n Q", unfolded wpt_def]])
7a3724078363 Replaced { } proofs by local lemmas; added Hoare logic with logical variables.
nipkow
parents: 63538
diff changeset
   129
  qed
63070
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   130
  from conseq[OF c1 hoaret.While[OF this w2 w3] c3]
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   131
  show ?case .
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   132
qed
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   133
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   134
theorem hoaret_complete: "\<Turnstile>\<^sub>t {P}c{Q} \<Longrightarrow> \<turnstile>\<^sub>t {P}c{Q}"
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   135
apply(rule strengthen_pre[OF _ wpt_is_pre])
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   136
apply(auto simp: hoare_tvalid_def wpt_def)
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   137
done
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   138
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   139
corollary hoaret_sound_complete: "\<turnstile>\<^sub>t {P}c{Q} \<longleftrightarrow> \<Turnstile>\<^sub>t {P}c{Q}"
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   140
by (metis hoaret_sound hoaret_complete)
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   141
74500
40f03761f77f more examples
nipkow
parents: 69505
diff changeset
   142
text \<open>Two examples:\<close>
40f03761f77f more examples
nipkow
parents: 69505
diff changeset
   143
40f03761f77f more examples
nipkow
parents: 69505
diff changeset
   144
lemma "\<turnstile>\<^sub>t 
40f03761f77f more examples
nipkow
parents: 69505
diff changeset
   145
{\<lambda>s. \<exists>n. n = nat(s ''x'')}
40f03761f77f more examples
nipkow
parents: 69505
diff changeset
   146
 WHILE Less (N 0) (V ''x'') DO ''x'' ::= Plus (V ''x'') (N (-1))
40f03761f77f more examples
nipkow
parents: 69505
diff changeset
   147
{\<lambda>s. s ''x'' \<le> 0}"
40f03761f77f more examples
nipkow
parents: 69505
diff changeset
   148
apply(rule weaken_post)
40f03761f77f more examples
nipkow
parents: 69505
diff changeset
   149
 apply(rule While)
40f03761f77f more examples
nipkow
parents: 69505
diff changeset
   150
   apply(rule Assign')
40f03761f77f more examples
nipkow
parents: 69505
diff changeset
   151
   apply auto
40f03761f77f more examples
nipkow
parents: 69505
diff changeset
   152
done
40f03761f77f more examples
nipkow
parents: 69505
diff changeset
   153
40f03761f77f more examples
nipkow
parents: 69505
diff changeset
   154
lemma "\<turnstile>\<^sub>t 
40f03761f77f more examples
nipkow
parents: 69505
diff changeset
   155
{\<lambda>s. \<exists>n. n = nat(s ''x'')}
40f03761f77f more examples
nipkow
parents: 69505
diff changeset
   156
 WHILE Less (N 0) (V ''x'')
40f03761f77f more examples
nipkow
parents: 69505
diff changeset
   157
 DO (''x'' ::= Plus (V ''x'') (N (-1));;
40f03761f77f more examples
nipkow
parents: 69505
diff changeset
   158
    (''y'' ::= V ''x'';;
40f03761f77f more examples
nipkow
parents: 69505
diff changeset
   159
     WHILE Less (N 0) (V ''y'') DO ''y'' ::= Plus (V ''y'') (N (-1))))
40f03761f77f more examples
nipkow
parents: 69505
diff changeset
   160
{\<lambda>s. s ''x'' \<le> 0}"
40f03761f77f more examples
nipkow
parents: 69505
diff changeset
   161
apply(rule weaken_post)
40f03761f77f more examples
nipkow
parents: 69505
diff changeset
   162
 apply(rule While)
40f03761f77f more examples
nipkow
parents: 69505
diff changeset
   163
   defer
40f03761f77f more examples
nipkow
parents: 69505
diff changeset
   164
   apply auto[3]
40f03761f77f more examples
nipkow
parents: 69505
diff changeset
   165
apply(rule Seq)
40f03761f77f more examples
nipkow
parents: 69505
diff changeset
   166
 prefer 2
40f03761f77f more examples
nipkow
parents: 69505
diff changeset
   167
 apply(rule Seq)
40f03761f77f more examples
nipkow
parents: 69505
diff changeset
   168
  prefer 2
40f03761f77f more examples
nipkow
parents: 69505
diff changeset
   169
  apply(rule weaken_post)
40f03761f77f more examples
nipkow
parents: 69505
diff changeset
   170
   apply(rule_tac P = "\<lambda>m s. n = nat(s ''x'') \<and> m = nat(s ''y'')" in While)
40f03761f77f more examples
nipkow
parents: 69505
diff changeset
   171
     apply(rule Assign')
40f03761f77f more examples
nipkow
parents: 69505
diff changeset
   172
     apply auto[4]
40f03761f77f more examples
nipkow
parents: 69505
diff changeset
   173
 apply(rule Assign)
40f03761f77f more examples
nipkow
parents: 69505
diff changeset
   174
apply(rule Assign')
40f03761f77f more examples
nipkow
parents: 69505
diff changeset
   175
apply auto
40f03761f77f more examples
nipkow
parents: 69505
diff changeset
   176
done
40f03761f77f more examples
nipkow
parents: 69505
diff changeset
   177
63070
952714a20087 the standard While-rule
nipkow
parents:
diff changeset
   178
end