| author | wenzelm | 
| Sun, 16 Jul 2023 15:53:13 +0200 | |
| changeset 78366 | aa4ea5398ab8 | 
| parent 75539 | 2e8a2dc7a1e6 | 
| child 80914 | d97fdabd9e2b | 
| permissions | -rw-r--r-- | 
| 
72987
 
b1be35908165
clarified modules: avoid multiple uses of the same ML file;
 
wenzelm 
parents:  
diff
changeset
 | 
1  | 
(* Title: HOL/Hoare/Hoare_Syntax.thy  | 
| 
 
b1be35908165
clarified modules: avoid multiple uses of the same ML file;
 
wenzelm 
parents:  
diff
changeset
 | 
2  | 
Author: Leonor Prensa Nieto & Tobias Nipkow  | 
| 
 
b1be35908165
clarified modules: avoid multiple uses of the same ML file;
 
wenzelm 
parents:  
diff
changeset
 | 
3  | 
Author: Walter Guttmann (extension to total-correctness proofs)  | 
| 72990 | 4  | 
*)  | 
| 
72987
 
b1be35908165
clarified modules: avoid multiple uses of the same ML file;
 
wenzelm 
parents:  
diff
changeset
 | 
5  | 
|
| 72990 | 6  | 
section \<open>Concrete syntax for Hoare logic, with translations for variables\<close>  | 
| 
72987
 
b1be35908165
clarified modules: avoid multiple uses of the same ML file;
 
wenzelm 
parents:  
diff
changeset
 | 
7  | 
|
| 
 
b1be35908165
clarified modules: avoid multiple uses of the same ML file;
 
wenzelm 
parents:  
diff
changeset
 | 
8  | 
theory Hoare_Syntax  | 
| 
 
b1be35908165
clarified modules: avoid multiple uses of the same ML file;
 
wenzelm 
parents:  
diff
changeset
 | 
9  | 
imports Main  | 
| 
 
b1be35908165
clarified modules: avoid multiple uses of the same ML file;
 
wenzelm 
parents:  
diff
changeset
 | 
10  | 
begin  | 
| 
 
b1be35908165
clarified modules: avoid multiple uses of the same ML file;
 
wenzelm 
parents:  
diff
changeset
 | 
11  | 
|
| 
 
b1be35908165
clarified modules: avoid multiple uses of the same ML file;
 
wenzelm 
parents:  
diff
changeset
 | 
12  | 
syntax  | 
| 
 
b1be35908165
clarified modules: avoid multiple uses of the same ML file;
 
wenzelm 
parents:  
diff
changeset
 | 
13  | 
  "_assign" :: "idt \<Rightarrow> 'b \<Rightarrow> 'com"  ("(2_ :=/ _)" [70, 65] 61)
 | 
| 
 
b1be35908165
clarified modules: avoid multiple uses of the same ML file;
 
wenzelm 
parents:  
diff
changeset
 | 
14  | 
  "_Seq" :: "'com \<Rightarrow> 'com \<Rightarrow> 'com"  ("(_;/ _)" [61,60] 60)
 | 
| 75539 | 15  | 
  "_Cond" :: "'bexp \<Rightarrow> 'com \<Rightarrow> 'com \<Rightarrow> 'com"  ("(IF _/ THEN _ / ELSE _/ FI)" [0,0,0] 61)
 | 
16  | 
  "_While" :: "'bexp \<Rightarrow> 'assn \<Rightarrow> 'var \<Rightarrow> 'com \<Rightarrow> 'com"  ("(WHILE _/ INV {(_)} / VAR {(_)} //DO _ /OD)" [0,0,0,0] 61)
 | 
|
| 
72987
 
b1be35908165
clarified modules: avoid multiple uses of the same ML file;
 
wenzelm 
parents:  
diff
changeset
 | 
17  | 
|
| 
74503
 
403ce50e6a2a
separated commands from annotations to be able to abstract about the latter only
 
nipkow 
parents: 
72990 
diff
changeset
 | 
18  | 
text \<open>The \<open>VAR {_}\<close> syntax supports two variants:
 | 
| 
 
403ce50e6a2a
separated commands from annotations to be able to abstract about the latter only
 
nipkow 
parents: 
72990 
diff
changeset
 | 
19  | 
\<^item> \<open>VAR {x = t}\<close> where \<open>t::nat\<close> is the decreasing expression,
 | 
| 
 
403ce50e6a2a
separated commands from annotations to be able to abstract about the latter only
 
nipkow 
parents: 
72990 
diff
changeset
 | 
20  | 
the variant, and \<open>x\<close> a variable that can be referred to from inner annotations.  | 
| 
 
403ce50e6a2a
separated commands from annotations to be able to abstract about the latter only
 
nipkow 
parents: 
72990 
diff
changeset
 | 
21  | 
The \<open>x\<close> can be necessary for nested loops, e.g. to prove that the inner loops do not mess with \<open>t\<close>.  | 
| 
 
403ce50e6a2a
separated commands from annotations to be able to abstract about the latter only
 
nipkow 
parents: 
72990 
diff
changeset
 | 
22  | 
\<^item> \<open>VAR {t}\<close> where the variable is omitted because it is not needed.
 | 
| 
 
403ce50e6a2a
separated commands from annotations to be able to abstract about the latter only
 
nipkow 
parents: 
72990 
diff
changeset
 | 
23  | 
\<close>  | 
| 
 
403ce50e6a2a
separated commands from annotations to be able to abstract about the latter only
 
nipkow 
parents: 
72990 
diff
changeset
 | 
24  | 
|
| 
72987
 
b1be35908165
clarified modules: avoid multiple uses of the same ML file;
 
wenzelm 
parents:  
diff
changeset
 | 
25  | 
syntax  | 
| 
 
b1be35908165
clarified modules: avoid multiple uses of the same ML file;
 
wenzelm 
parents:  
diff
changeset
 | 
26  | 
  "_While0" :: "'bexp \<Rightarrow> 'assn \<Rightarrow> 'com \<Rightarrow> 'com"  ("(1WHILE _/ INV {_} //DO _ /OD)" [0,0,0] 61)
 | 
| 
74503
 
403ce50e6a2a
separated commands from annotations to be able to abstract about the latter only
 
nipkow 
parents: 
72990 
diff
changeset
 | 
27  | 
|
| 
 
403ce50e6a2a
separated commands from annotations to be able to abstract about the latter only
 
nipkow 
parents: 
72990 
diff
changeset
 | 
28  | 
text \<open>The \<open>_While0\<close> syntax is translated into the \<open>_While\<close> syntax with the trivial variant 0.  | 
| 
 
403ce50e6a2a
separated commands from annotations to be able to abstract about the latter only
 
nipkow 
parents: 
72990 
diff
changeset
 | 
29  | 
This is ok because partial correctness proofs do not make use of the variant.\<close>  | 
| 
72987
 
b1be35908165
clarified modules: avoid multiple uses of the same ML file;
 
wenzelm 
parents:  
diff
changeset
 | 
30  | 
|
| 
 
b1be35908165
clarified modules: avoid multiple uses of the same ML file;
 
wenzelm 
parents:  
diff
changeset
 | 
31  | 
syntax  | 
| 
 
b1be35908165
clarified modules: avoid multiple uses of the same ML file;
 
wenzelm 
parents:  
diff
changeset
 | 
32  | 
 "_hoare_vars" :: "[idts, 'assn,'com, 'assn] \<Rightarrow> bool"  ("VARS _// {_} // _ // {_}" [0,0,55,0] 50)
 | 
| 
 
b1be35908165
clarified modules: avoid multiple uses of the same ML file;
 
wenzelm 
parents:  
diff
changeset
 | 
33  | 
 "_hoare_vars_tc" :: "[idts, 'assn, 'com, 'assn] \<Rightarrow> bool"  ("VARS _// [_] // _ // [_]" [0,0,55,0] 50)
 | 
| 
 
b1be35908165
clarified modules: avoid multiple uses of the same ML file;
 
wenzelm 
parents:  
diff
changeset
 | 
34  | 
syntax (output)  | 
| 75539 | 35  | 
 "_hoare" :: "['assn, 'com, 'assn] \<Rightarrow> bool"  ("({_}//_//{_})" [0,55,0] 50)
 | 
36  | 
 "_hoare_tc" :: "['assn, 'com, 'assn] \<Rightarrow> bool"  ("([_]//_//[_])" [0,55,0] 50)
 | 
|
| 
72987
 
b1be35908165
clarified modules: avoid multiple uses of the same ML file;
 
wenzelm 
parents:  
diff
changeset
 | 
37  | 
|
| 
74503
 
403ce50e6a2a
separated commands from annotations to be able to abstract about the latter only
 
nipkow 
parents: 
72990 
diff
changeset
 | 
38  | 
text \<open>Completeness requires(?) the ability to refer to an outer variant in an inner invariant.  | 
| 
 
403ce50e6a2a
separated commands from annotations to be able to abstract about the latter only
 
nipkow 
parents: 
72990 
diff
changeset
 | 
39  | 
Thus we need to abstract over a variable equated with the variant, the \<open>x\<close> in \<open>VAR {x = t}\<close>.
 | 
| 
 
403ce50e6a2a
separated commands from annotations to be able to abstract about the latter only
 
nipkow 
parents: 
72990 
diff
changeset
 | 
40  | 
But the \<open>x\<close> should only occur in invariants. To enforce this, syntax translations in \<open>hoare_syntax.ML\<close>  | 
| 
 
403ce50e6a2a
separated commands from annotations to be able to abstract about the latter only
 
nipkow 
parents: 
72990 
diff
changeset
 | 
41  | 
separate the program from its annotations and only the latter are abstracted over over \<open>x\<close>.  | 
| 
 
403ce50e6a2a
separated commands from annotations to be able to abstract about the latter only
 
nipkow 
parents: 
72990 
diff
changeset
 | 
42  | 
(Thus \<open>x\<close> can also occur in inner variants, but that neither helps nor hurts.)\<close>  | 
| 
 
403ce50e6a2a
separated commands from annotations to be able to abstract about the latter only
 
nipkow 
parents: 
72990 
diff
changeset
 | 
43  | 
|
| 
 
403ce50e6a2a
separated commands from annotations to be able to abstract about the latter only
 
nipkow 
parents: 
72990 
diff
changeset
 | 
44  | 
datatype 'a anno =  | 
| 
 
403ce50e6a2a
separated commands from annotations to be able to abstract about the latter only
 
nipkow 
parents: 
72990 
diff
changeset
 | 
45  | 
Abasic |  | 
| 
 
403ce50e6a2a
separated commands from annotations to be able to abstract about the latter only
 
nipkow 
parents: 
72990 
diff
changeset
 | 
46  | 
Aseq "'a anno" "'a anno" |  | 
| 
 
403ce50e6a2a
separated commands from annotations to be able to abstract about the latter only
 
nipkow 
parents: 
72990 
diff
changeset
 | 
47  | 
Acond "'a anno" "'a anno" |  | 
| 
 
403ce50e6a2a
separated commands from annotations to be able to abstract about the latter only
 
nipkow 
parents: 
72990 
diff
changeset
 | 
48  | 
Awhile "'a set" "'a \<Rightarrow> nat" "nat \<Rightarrow> 'a anno"  | 
| 
 
403ce50e6a2a
separated commands from annotations to be able to abstract about the latter only
 
nipkow 
parents: 
72990 
diff
changeset
 | 
49  | 
|
| 
72987
 
b1be35908165
clarified modules: avoid multiple uses of the same ML file;
 
wenzelm 
parents:  
diff
changeset
 | 
50  | 
ML_file \<open>hoare_syntax.ML\<close>  | 
| 
 
b1be35908165
clarified modules: avoid multiple uses of the same ML file;
 
wenzelm 
parents:  
diff
changeset
 | 
51  | 
|
| 
 
b1be35908165
clarified modules: avoid multiple uses of the same ML file;
 
wenzelm 
parents:  
diff
changeset
 | 
52  | 
end  |