# HG changeset patch # User nipkow # Date 1370080086 -7200 # Node ID 780b3870319f93fc1de975e3ab38177fb75b85d9 # Parent c3f837d9253777a223f43369d977c0912070f679 tuned rules diff -r c3f837d92537 -r 780b3870319f src/HOL/IMP/HoareT.thy --- a/src/HOL/IMP/HoareT.thy Fri May 31 14:08:48 2013 +0200 +++ b/src/HOL/IMP/HoareT.thy Sat Jun 01 11:48:06 2013 +0200 @@ -9,7 +9,7 @@ definition hoare_tvalid :: "assn \ com \ assn \ bool" ("\\<^sub>t {(1_)}/ (_)/ {(1_)}" 50) where -"\\<^sub>t {P}c{Q} \ \s. P s \ (\t. (c,s) \ t \ Q t)" +"\\<^sub>t {P}c{Q} \ (\s. P s \ (\t. (c,s) \ t \ Q t))" text{* Provability of Hoare triples in the proof system for total correctness is written @{text"\\<^sub>t {P}c{Q}"} and defined @@ -20,20 +20,35 @@ inductive hoaret :: "assn \ com \ assn \ bool" ("\\<^sub>t ({(1_)}/ (_)/ {(1_)})" 50) where -Skip: "\\<^sub>t {P} SKIP {P}" | -Assign: "\\<^sub>t {\s. P(s[a/x])} x::=a {P}" | -Seq: "\ \\<^sub>t {P\<^isub>1} c\<^isub>1 {P\<^isub>2}; \\<^sub>t {P\<^isub>2} c\<^isub>2 {P\<^isub>3} \ \ \\<^sub>t {P\<^isub>1} c\<^isub>1;;c\<^isub>2 {P\<^isub>3}" | + +Skip: "\\<^sub>t {P} SKIP {P}" | + +Assign: "\\<^sub>t {\s. P(s[a/x])} x::=a {P}" | + +Seq: "\ \\<^sub>t {P\<^isub>1} c\<^isub>1 {P\<^isub>2}; \\<^sub>t {P\<^isub>2} c\<^isub>2 {P\<^isub>3} \ \ \\<^sub>t {P\<^isub>1} c\<^isub>1;;c\<^isub>2 {P\<^isub>3}" | + If: "\ \\<^sub>t {\s. P s \ bval b s} c\<^isub>1 {Q}; \\<^sub>t {\s. P s \ \ bval b s} c\<^isub>2 {Q} \ - \ \\<^sub>t {P} IF b THEN c\<^isub>1 ELSE c\<^isub>2 {Q}" | + \ \\<^sub>t {P} IF b THEN c\<^isub>1 ELSE c\<^isub>2 {Q}" | + While: - "\ \n::nat. \\<^sub>t {\s. P s \ bval b s \ T s n} c {\s. P s \ (\n'. T s n' \ n' < n)} \ - \ \\<^sub>t {\s. P s \ (\n. T s n)} WHILE b DO c {\s. P s \ \bval b s}" | + "(\n::nat. + \\<^sub>t {\s. P s \ bval b s \ T s n} c {\s. P s \ (\n'. T s n' \ n' < n)}) + \ \\<^sub>t {\s. P s \ (\n. T s n)} WHILE b DO c {\s. P s \ \bval b s}" | + conseq: "\ \s. P' s \ P s; \\<^sub>t {P}c{Q}; \s. Q s \ Q' s \ \ \\<^sub>t {P'}c{Q'}" text{* The @{term While}-rule is like the one for partial correctness but it requires additionally that with every execution of the loop body some measure -function @{term[source]"f :: state \ nat"} decreases. *} +relation @{term[source]"T :: state \ nat \ bool"} decreases. +The following functional version is more intuitive: *} + +lemma While_fun: + "\ \n::nat. \\<^sub>t {\s. P s \ bval b s \ n = f s} c {\s. P s \ f s < n}\ + \ \\<^sub>t {P} WHILE b DO c {\s. P s \ \bval b s}" + by (rule While [where T="\s n. n = f s", simplified]) + +text{* Building in the consequence rule: *} lemma strengthen_pre: "\ \s. P' s \ P s; \\<^sub>t {P} c {Q} \ \ \\<^sub>t {P'} c {Q}" @@ -46,24 +61,20 @@ lemma Assign': "\s. P s \ Q(s[a/x]) \ \\<^sub>t {P} x ::= a {Q}" by (simp add: strengthen_pre[OF _ Assign]) -lemma While': -assumes "\n::nat. \\<^sub>t {\s. P s \ bval b s \ T s n} c {\s. P s \ (\n'. T s n' \ n' < n)}" +lemma While_fun': +assumes "\n::nat. \\<^sub>t {\s. P s \ bval b s \ n = f s} c {\s. P s \ f s < n}" and "\s. P s \ \ bval b s \ Q s" -shows "\\<^sub>t {\s. P s \ (\n. T s n)} WHILE b DO c {Q}" -by(blast intro: assms(1) weaken_post[OF While assms(2)]) +shows "\\<^sub>t {P} WHILE b DO c {Q}" +by(blast intro: assms(1) weaken_post[OF While_fun assms(2)]) -lemma While_fun: - "\ \n::nat. \\<^sub>t {\s. P s \ bval b s \ f s = n} c {\s. P s \ f s < n}\ - \ \\<^sub>t {P} WHILE b DO c {\s. P s \ \bval b s}" - by (rule While [where T="\s n. f s = n", simplified]) text{* Our standard example: *} lemma "\\<^sub>t {\s. s ''x'' = i} ''y'' ::= N 0;; wsum {\s. s ''y'' = sum i}" apply(rule Seq) prefer 2 - apply(rule While' [where P = "\s. (s ''y'' = sum i - sum(s ''x''))" - and T = "\s n. n = nat(s ''x'')"]) + apply(rule While_fun' [where P = "\s. (s ''y'' = sum i - sum(s ''x''))" + and f = "\s. nat(s ''x'')"]) apply(rule Seq) prefer 2 apply(rule Assign)