src/HOL/NanoJava/AxSem.thy
author blanchet
Mon, 23 Aug 2010 12:13:58 +0200
changeset 38649 14c207135eff
parent 35431 8758fe1fc9f8
child 41589 bbd861837ebc
permissions -rw-r--r--
"no_atp" fact that leads to unsound proofs
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
     1
(*  Title:      HOL/NanoJava/AxSem.thy
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
     2
    ID:         $Id$
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
     3
    Author:     David von Oheimb
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
     4
    Copyright   2001 Technische Universitaet Muenchen
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
     5
*)
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
     6
11560
46d0bde121ab marginally improved comments
oheimb
parents: 11559
diff changeset
     7
header "Axiomatic Semantics"
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
     8
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 12934
diff changeset
     9
theory AxSem imports State begin
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    10
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    11
types assn   = "state => bool"
11486
8f32860eac3a layout, subscripts
oheimb
parents: 11476
diff changeset
    12
     vassn   = "val => assn"
11476
06c1998340a8 changed to full expressions with side effects
oheimb
parents: 11449
diff changeset
    13
      triple = "assn \<times> stmt \<times>  assn"
06c1998340a8 changed to full expressions with side effects
oheimb
parents: 11449
diff changeset
    14
     etriple = "assn \<times> expr \<times> vassn"
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    15
translations
35431
8758fe1fc9f8 cleanup type translations;
wenzelm
parents: 23894
diff changeset
    16
  (type) "assn" \<leftharpoondown> (type) "state => bool"
8758fe1fc9f8 cleanup type translations;
wenzelm
parents: 23894
diff changeset
    17
  (type) "vassn" \<leftharpoondown> (type) "val => assn"
8758fe1fc9f8 cleanup type translations;
wenzelm
parents: 23894
diff changeset
    18
  (type) "triple" \<leftharpoondown> (type) "assn \<times> stmt \<times> assn"
8758fe1fc9f8 cleanup type translations;
wenzelm
parents: 23894
diff changeset
    19
  (type) "etriple" \<leftharpoondown> (type) "assn \<times> expr \<times> vassn"
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    20
11476
06c1998340a8 changed to full expressions with side effects
oheimb
parents: 11449
diff changeset
    21
11560
46d0bde121ab marginally improved comments
oheimb
parents: 11559
diff changeset
    22
subsection "Hoare Logic Rules"
46d0bde121ab marginally improved comments
oheimb
parents: 11559
diff changeset
    23
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    24
inductive
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    25
 hoare :: "[triple set, triple set] => bool"  ("_ |\<turnstile>/ _" [61, 61] 60)
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    26
 and ehoare :: "[triple set, etriple] => bool"  ("_ |\<turnstile>\<^sub>e/ _" [61, 61] 60)
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    27
 and hoare1 :: "[triple set, assn,stmt,assn] => bool" 
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    28
   ("_ \<turnstile>/ ({(1_)}/ (_)/ {(1_)})" [61, 3, 90, 3] 60)
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    29
 and ehoare1 :: "[triple set, assn,expr,vassn]=> bool"
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    30
   ("_ \<turnstile>\<^sub>e/ ({(1_)}/ (_)/ {(1_)})" [61, 3, 90, 3] 60)
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    31
where
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    32
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    33
  "A  \<turnstile> {P}c{Q} \<equiv> A |\<turnstile> {(P,c,Q)}"
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    34
| "A  \<turnstile>\<^sub>e {P}e{Q} \<equiv> A |\<turnstile>\<^sub>e (P,e,Q)"
11476
06c1998340a8 changed to full expressions with side effects
oheimb
parents: 11449
diff changeset
    35
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    36
| Skip:  "A \<turnstile> {P} Skip {P}"
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    37
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    38
| Comp: "[| A \<turnstile> {P} c1 {Q}; A \<turnstile> {Q} c2 {R} |] ==> A \<turnstile> {P} c1;;c2 {R}"
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    39
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    40
| Cond: "[| A \<turnstile>\<^sub>e {P} e {Q}; 
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    41
            \<forall>v. A \<turnstile> {Q v} (if v \<noteq> Null then c1 else c2) {R} |] ==>
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    42
            A \<turnstile> {P} If(e) c1 Else c2 {R}"
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    43
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    44
| Loop: "A \<turnstile> {\<lambda>s. P s \<and> s<x> \<noteq> Null} c {P} ==>
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    45
         A \<turnstile> {P} While(x) c {\<lambda>s. P s \<and> s<x> = Null}"
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    46
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    47
| LAcc: "A \<turnstile>\<^sub>e {\<lambda>s. P (s<x>) s} LAcc x {P}"
11476
06c1998340a8 changed to full expressions with side effects
oheimb
parents: 11449
diff changeset
    48
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    49
| LAss: "A \<turnstile>\<^sub>e {P} e {\<lambda>v s.  Q (lupd(x\<mapsto>v) s)} ==>
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    50
         A \<turnstile>  {P} x:==e {Q}"
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    51
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    52
| FAcc: "A \<turnstile>\<^sub>e {P} e {\<lambda>v s. \<forall>a. v=Addr a --> Q (get_field s a f) s} ==>
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    53
         A \<turnstile>\<^sub>e {P} e..f {Q}"
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    54
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    55
| FAss: "[| A \<turnstile>\<^sub>e {P} e1 {\<lambda>v s. \<forall>a. v=Addr a --> Q a s};
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    56
        \<forall>a. A \<turnstile>\<^sub>e {Q a} e2 {\<lambda>v s. R (upd_obj a f v s)} |] ==>
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    57
            A \<turnstile>  {P} e1..f:==e2 {R}"
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    58
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    59
| NewC: "A \<turnstile>\<^sub>e {\<lambda>s. \<forall>a. new_Addr s = Addr a --> P (Addr a) (new_obj a C s)}
11476
06c1998340a8 changed to full expressions with side effects
oheimb
parents: 11449
diff changeset
    60
                new C {P}"
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    61
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    62
| Cast: "A \<turnstile>\<^sub>e {P} e {\<lambda>v s. (case v of Null => True 
11476
06c1998340a8 changed to full expressions with side effects
oheimb
parents: 11449
diff changeset
    63
                                 | Addr a => obj_class s a <=C C) --> Q v s} ==>
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    64
         A \<turnstile>\<^sub>e {P} Cast C e {Q}"
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    65
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    66
| Call: "[| A \<turnstile>\<^sub>e {P} e1 {Q}; \<forall>a. A \<turnstile>\<^sub>e {Q a} e2 {R a};
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    67
    \<forall>a p ls. A \<turnstile> {\<lambda>s'. \<exists>s. R a p s \<and> ls = s \<and> 
11772
cf618fe8facd renamed reset_locs to del_locs
oheimb
parents: 11565
diff changeset
    68
                    s' = lupd(This\<mapsto>a)(lupd(Par\<mapsto>p)(del_locs s))}
11565
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
    69
                  Meth (C,m) {\<lambda>s. S (s<Res>) (set_locs ls s)} |] ==>
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    70
             A \<turnstile>\<^sub>e {P} {C}e1..m(e2) {S}"
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    71
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    72
| Meth: "\<forall>D. A \<turnstile> {\<lambda>s'. \<exists>s a. s<This> = Addr a \<and> D = obj_class s a \<and> D <=C C \<and> 
11497
0e66e0114d9a corrected initialization of locals, streamlined Impl
oheimb
parents: 11486
diff changeset
    73
                        P s \<and> s' = init_locs D m s}
0e66e0114d9a corrected initialization of locals, streamlined Impl
oheimb
parents: 11486
diff changeset
    74
                  Impl (D,m) {Q} ==>
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    75
             A \<turnstile> {P} Meth (C,m) {Q}"
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    76
11565
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
    77
  --{* @{text "\<Union>Z"} instead of @{text "\<forall>Z"} in the conclusion and\\
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
    78
       Z restricted to type state due to limitations of the inductive package *}
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    79
| Impl: "\<forall>Z::state. A\<union> (\<Union>Z. (\<lambda>Cm. (P Z Cm, Impl Cm, Q Z Cm))`Ms) |\<turnstile> 
11565
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
    80
                            (\<lambda>Cm. (P Z Cm, body Cm, Q Z Cm))`Ms ==>
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    81
                      A |\<turnstile> (\<lambda>Cm. (P Z Cm, Impl Cm, Q Z Cm))`Ms"
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    82
11558
6539627881e8 simplified vnam/vname, introduced fname, improved comments
oheimb
parents: 11508
diff changeset
    83
--{* structural rules *}
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    84
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    85
| Asm:  "   a \<in> A ==> A |\<turnstile> {a}"
11476
06c1998340a8 changed to full expressions with side effects
oheimb
parents: 11449
diff changeset
    86
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    87
| ConjI: " \<forall>c \<in> C. A |\<turnstile> {c} ==> A |\<turnstile> C"
11476
06c1998340a8 changed to full expressions with side effects
oheimb
parents: 11449
diff changeset
    88
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    89
| ConjE: "[|A |\<turnstile> C; c \<in> C |] ==> A |\<turnstile> {c}"
11476
06c1998340a8 changed to full expressions with side effects
oheimb
parents: 11449
diff changeset
    90
11565
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
    91
  --{* Z restricted to type state due to limitations of the inductive package *}
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    92
| Conseq:"[| \<forall>Z::state. A \<turnstile> {P' Z} c {Q' Z};
11565
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
    93
             \<forall>s t. (\<forall>Z. P' Z s --> Q' Z t) --> (P s --> Q t) |] ==>
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    94
            A \<turnstile> {P} c {Q }"
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
    95
11565
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
    96
  --{* Z restricted to type state due to limitations of the inductive package *}
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    97
| eConseq:"[| \<forall>Z::state. A \<turnstile>\<^sub>e {P' Z} e {Q' Z};
11565
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
    98
             \<forall>s v t. (\<forall>Z. P' Z s --> Q' Z v t) --> (P s --> Q v t) |] ==>
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
    99
            A \<turnstile>\<^sub>e {P} e {Q }"
11565
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   100
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   101
12934
6003b4f916c0 Clarification wrt. use of polymorphic variants of Hoare logic rules
oheimb
parents: 11772
diff changeset
   102
subsection "Fully polymorphic variants, required for Example only"
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
   103
11565
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   104
axioms
12934
6003b4f916c0 Clarification wrt. use of polymorphic variants of Hoare logic rules
oheimb
parents: 11772
diff changeset
   105
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
   106
  Conseq:"[| \<forall>Z. A \<turnstile> {P' Z} c {Q' Z};
11565
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   107
             \<forall>s t. (\<forall>Z. P' Z s --> Q' Z t) --> (P s --> Q t) |] ==>
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
   108
                 A \<turnstile> {P} c {Q }"
11565
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   109
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
   110
 eConseq:"[| \<forall>Z. A \<turnstile>\<^sub>e {P' Z} e {Q' Z};
11565
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   111
             \<forall>s v t. (\<forall>Z. P' Z s --> Q' Z v t) --> (P s --> Q v t) |] ==>
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
   112
                 A \<turnstile>\<^sub>e {P} e {Q }"
11565
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   113
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
   114
 Impl: "\<forall>Z. A\<union> (\<Union>Z. (\<lambda>Cm. (P Z Cm, Impl Cm, Q Z Cm))`Ms) |\<turnstile> 
11565
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   115
                          (\<lambda>Cm. (P Z Cm, body Cm, Q Z Cm))`Ms ==>
23755
1c4672d130b1 Adapted to new inductive definition package.
berghofe
parents: 16417
diff changeset
   116
                    A |\<turnstile> (\<lambda>Cm. (P Z Cm, Impl Cm, Q Z Cm))`Ms"
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
   117
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
   118
subsection "Derived Rules"
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
   119
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
   120
lemma Conseq1: "\<lbrakk>A \<turnstile> {P'} c {Q}; \<forall>s. P s \<longrightarrow> P' s\<rbrakk> \<Longrightarrow> A \<turnstile> {P} c {Q}"
11476
06c1998340a8 changed to full expressions with side effects
oheimb
parents: 11449
diff changeset
   121
apply (rule hoare_ehoare.Conseq)
06c1998340a8 changed to full expressions with side effects
oheimb
parents: 11449
diff changeset
   122
apply  (rule allI, assumption)
06c1998340a8 changed to full expressions with side effects
oheimb
parents: 11449
diff changeset
   123
apply fast
06c1998340a8 changed to full expressions with side effects
oheimb
parents: 11449
diff changeset
   124
done
06c1998340a8 changed to full expressions with side effects
oheimb
parents: 11449
diff changeset
   125
06c1998340a8 changed to full expressions with side effects
oheimb
parents: 11449
diff changeset
   126
lemma Conseq2: "\<lbrakk>A \<turnstile> {P} c {Q'}; \<forall>t. Q' t \<longrightarrow> Q t\<rbrakk> \<Longrightarrow> A \<turnstile> {P} c {Q}"
06c1998340a8 changed to full expressions with side effects
oheimb
parents: 11449
diff changeset
   127
apply (rule hoare_ehoare.Conseq)
06c1998340a8 changed to full expressions with side effects
oheimb
parents: 11449
diff changeset
   128
apply  (rule allI, assumption)
06c1998340a8 changed to full expressions with side effects
oheimb
parents: 11449
diff changeset
   129
apply fast
06c1998340a8 changed to full expressions with side effects
oheimb
parents: 11449
diff changeset
   130
done
06c1998340a8 changed to full expressions with side effects
oheimb
parents: 11449
diff changeset
   131
11486
8f32860eac3a layout, subscripts
oheimb
parents: 11476
diff changeset
   132
lemma eConseq1: "\<lbrakk>A \<turnstile>\<^sub>e {P'} e {Q}; \<forall>s. P s \<longrightarrow> P' s\<rbrakk> \<Longrightarrow> A \<turnstile>\<^sub>e {P} e {Q}"
11476
06c1998340a8 changed to full expressions with side effects
oheimb
parents: 11449
diff changeset
   133
apply (rule hoare_ehoare.eConseq)
06c1998340a8 changed to full expressions with side effects
oheimb
parents: 11449
diff changeset
   134
apply  (rule allI, assumption)
06c1998340a8 changed to full expressions with side effects
oheimb
parents: 11449
diff changeset
   135
apply fast
06c1998340a8 changed to full expressions with side effects
oheimb
parents: 11449
diff changeset
   136
done
06c1998340a8 changed to full expressions with side effects
oheimb
parents: 11449
diff changeset
   137
11486
8f32860eac3a layout, subscripts
oheimb
parents: 11476
diff changeset
   138
lemma eConseq2: "\<lbrakk>A \<turnstile>\<^sub>e {P} e {Q'}; \<forall>v t. Q' v t \<longrightarrow> Q v t\<rbrakk> \<Longrightarrow> A \<turnstile>\<^sub>e {P} e {Q}"
11476
06c1998340a8 changed to full expressions with side effects
oheimb
parents: 11449
diff changeset
   139
apply (rule hoare_ehoare.eConseq)
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
   140
apply  (rule allI, assumption)
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
   141
apply fast
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
   142
done
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
   143
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
   144
lemma Weaken: "\<lbrakk>A |\<turnstile> C'; C \<subseteq> C'\<rbrakk> \<Longrightarrow> A |\<turnstile> C"
11476
06c1998340a8 changed to full expressions with side effects
oheimb
parents: 11449
diff changeset
   145
apply (rule hoare_ehoare.ConjI)
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
   146
apply clarify
11476
06c1998340a8 changed to full expressions with side effects
oheimb
parents: 11449
diff changeset
   147
apply (drule hoare_ehoare.ConjE)
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
   148
apply  fast
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
   149
apply assumption
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
   150
done
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
   151
11565
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   152
lemma Thin_lemma: 
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   153
  "(A' |\<turnstile>  C         \<longrightarrow> (\<forall>A. A' \<subseteq> A \<longrightarrow> A |\<turnstile>  C       )) \<and> 
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   154
   (A'  \<turnstile>\<^sub>e {P} e {Q} \<longrightarrow> (\<forall>A. A' \<subseteq> A \<longrightarrow> A  \<turnstile>\<^sub>e {P} e {Q}))"
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   155
apply (rule hoare_ehoare.induct)
23894
1a4167d761ac tactics: avoid dynamic reference to accidental theory context (via ML_Context.the_context etc.);
wenzelm
parents: 23755
diff changeset
   156
apply (tactic "ALLGOALS(EVERY'[clarify_tac @{claset}, REPEAT o smp_tac 1])")
11565
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   157
apply (blast intro: hoare_ehoare.Skip)
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   158
apply (blast intro: hoare_ehoare.Comp)
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   159
apply (blast intro: hoare_ehoare.Cond)
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   160
apply (blast intro: hoare_ehoare.Loop)
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   161
apply (blast intro: hoare_ehoare.LAcc)
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   162
apply (blast intro: hoare_ehoare.LAss)
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   163
apply (blast intro: hoare_ehoare.FAcc)
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   164
apply (blast intro: hoare_ehoare.FAss)
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   165
apply (blast intro: hoare_ehoare.NewC)
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   166
apply (blast intro: hoare_ehoare.Cast)
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   167
apply (erule hoare_ehoare.Call)
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   168
apply   (rule, drule spec, erule conjE, tactic "smp_tac 1 1", assumption)
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   169
apply  blast
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   170
apply (blast intro!: hoare_ehoare.Meth)
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   171
apply (blast intro!: hoare_ehoare.Impl)
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   172
apply (blast intro!: hoare_ehoare.Asm)
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   173
apply (blast intro: hoare_ehoare.ConjI)
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   174
apply (blast intro: hoare_ehoare.ConjE)
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   175
apply (rule hoare_ehoare.Conseq)
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   176
apply  (rule, drule spec, erule conjE, tactic "smp_tac 1 1", assumption+)
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   177
apply (rule hoare_ehoare.eConseq)
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   178
apply  (rule, drule spec, erule conjE, tactic "smp_tac 1 1", assumption+)
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   179
done
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   180
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   181
lemma cThin: "\<lbrakk>A' |\<turnstile> C; A' \<subseteq> A\<rbrakk> \<Longrightarrow> A |\<turnstile> C"
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   182
by (erule (1) conjunct1 [OF Thin_lemma, rule_format])
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   183
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   184
lemma eThin: "\<lbrakk>A' \<turnstile>\<^sub>e {P} e {Q}; A' \<subseteq> A\<rbrakk> \<Longrightarrow> A \<turnstile>\<^sub>e {P} e {Q}"
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   185
by (erule (1) conjunct2 [OF Thin_lemma, rule_format])
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   186
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   187
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   188
lemma Union: "A |\<turnstile> (\<Union>Z. C Z) = (\<forall>Z. A |\<turnstile> C Z)"
11476
06c1998340a8 changed to full expressions with side effects
oheimb
parents: 11449
diff changeset
   189
by (auto intro: hoare_ehoare.ConjI hoare_ehoare.ConjE)
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
   190
11565
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   191
lemma Impl1': 
12934
6003b4f916c0 Clarification wrt. use of polymorphic variants of Hoare logic rules
oheimb
parents: 11772
diff changeset
   192
   "\<lbrakk>\<forall>Z::state. A\<union> (\<Union>Z. (\<lambda>Cm. (P Z Cm, Impl Cm, Q Z Cm))`Ms) |\<turnstile> 
11565
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   193
                 (\<lambda>Cm. (P Z Cm, body Cm, Q Z Cm))`Ms; 
11508
168dbdaedb71 cosmetics
oheimb
parents: 11507
diff changeset
   194
    Cm \<in> Ms\<rbrakk> \<Longrightarrow> 
11565
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   195
                A   \<turnstile>  {P Z Cm} Impl Cm {Q Z Cm}"
12934
6003b4f916c0 Clarification wrt. use of polymorphic variants of Hoare logic rules
oheimb
parents: 11772
diff changeset
   196
apply (drule AxSem.Impl)
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
   197
apply (erule Weaken)
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
   198
apply (auto del: image_eqI intro: rev_image_eqI)
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
   199
done
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
   200
11565
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   201
lemmas Impl1 = AxSem.Impl [of _ _ _ "{Cm}", simplified, standard]
ab004c0ecc63 Minor improvements, added Example
oheimb
parents: 11560
diff changeset
   202
11376
bf98ad1c22c6 added NanoJava
oheimb
parents:
diff changeset
   203
end