src/HOL/IMP/VC.ML
author nipkow
Mon, 09 Feb 1998 14:40:59 +0100
changeset 4612 26764de50c74
parent 4089 96fba19bcbe2
child 5069 3ea049f7979d
permissions -rw-r--r--
Used THEN_ALL_NEW.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1451
diff changeset
     1
(*  Title:      HOL/IMP/VC.ML
1447
bc2c0acbbf29 Added a verified verification-condition generator.
nipkow
parents:
diff changeset
     2
    ID:         $Id$
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1451
diff changeset
     3
    Author:     Tobias Nipkow
1447
bc2c0acbbf29 Added a verified verification-condition generator.
nipkow
parents:
diff changeset
     4
    Copyright   1996 TUM
bc2c0acbbf29 Added a verified verification-condition generator.
nipkow
parents:
diff changeset
     5
bc2c0acbbf29 Added a verified verification-condition generator.
nipkow
parents:
diff changeset
     6
Soundness and completeness of vc
bc2c0acbbf29 Added a verified verification-condition generator.
nipkow
parents:
diff changeset
     7
*)
bc2c0acbbf29 Added a verified verification-condition generator.
nipkow
parents:
diff changeset
     8
bc2c0acbbf29 Added a verified verification-condition generator.
nipkow
parents:
diff changeset
     9
open VC;
bc2c0acbbf29 Added a verified verification-condition generator.
nipkow
parents:
diff changeset
    10
1973
8c94c9a5be10 Converted proofs to use default clasets.
nipkow
parents: 1940
diff changeset
    11
AddIs hoare.intrs;
8c94c9a5be10 Converted proofs to use default clasets.
nipkow
parents: 1940
diff changeset
    12
3842
b55686a7b22c fixed dots;
wenzelm
parents: 2810
diff changeset
    13
val lemma = prove_goal HOL.thy "!s. P s --> P s" (K[Fast_tac 1]);
1447
bc2c0acbbf29 Added a verified verification-condition generator.
nipkow
parents:
diff changeset
    14
2810
c4e16b36bc57 Added wp_while.
nipkow
parents: 2055
diff changeset
    15
goal VC.thy "!Q. (!s. vc c Q s) --> |- {awp c Q} astrip c {Q}";
2031
03a843f0f447 Ran expandshort
paulson
parents: 1973
diff changeset
    16
by (acom.induct_tac "c" 1);
2055
cc274e47f607 Ran expandshort
paulson
parents: 2031
diff changeset
    17
    by (ALLGOALS Simp_tac);
cc274e47f607 Ran expandshort
paulson
parents: 2031
diff changeset
    18
    by (Fast_tac 1);
cc274e47f607 Ran expandshort
paulson
parents: 2031
diff changeset
    19
   by (Fast_tac 1);
cc274e47f607 Ran expandshort
paulson
parents: 2031
diff changeset
    20
  by (Fast_tac 1);
1447
bc2c0acbbf29 Added a verified verification-condition generator.
nipkow
parents:
diff changeset
    21
 (* if *)
2055
cc274e47f607 Ran expandshort
paulson
parents: 2031
diff changeset
    22
 by (Deepen_tac 4 1);
1447
bc2c0acbbf29 Added a verified verification-condition generator.
nipkow
parents:
diff changeset
    23
(* while *)
2031
03a843f0f447 Ran expandshort
paulson
parents: 1973
diff changeset
    24
by (safe_tac HOL_cs);
1465
5d7a7e439cec expanded tabs
clasohm
parents: 1451
diff changeset
    25
by (resolve_tac hoare.intrs 1);
2055
cc274e47f607 Ran expandshort
paulson
parents: 2031
diff changeset
    26
  by (rtac lemma 1);
cc274e47f607 Ran expandshort
paulson
parents: 2031
diff changeset
    27
 by (resolve_tac hoare.intrs 1);
cc274e47f607 Ran expandshort
paulson
parents: 2031
diff changeset
    28
 by (resolve_tac hoare.intrs 1);
cc274e47f607 Ran expandshort
paulson
parents: 2031
diff changeset
    29
   by (ALLGOALS(fast_tac HOL_cs));
1447
bc2c0acbbf29 Added a verified verification-condition generator.
nipkow
parents:
diff changeset
    30
qed "vc_sound";
bc2c0acbbf29 Added a verified verification-condition generator.
nipkow
parents:
diff changeset
    31
2810
c4e16b36bc57 Added wp_while.
nipkow
parents: 2055
diff changeset
    32
goal VC.thy "!P Q. (!s. P s --> Q s) --> (!s. awp c P s --> awp c Q s)";
2031
03a843f0f447 Ran expandshort
paulson
parents: 1973
diff changeset
    33
by (acom.induct_tac "c" 1);
2055
cc274e47f607 Ran expandshort
paulson
parents: 2031
diff changeset
    34
    by (ALLGOALS Asm_simp_tac);
2031
03a843f0f447 Ran expandshort
paulson
parents: 1973
diff changeset
    35
by (EVERY1[rtac allI, rtac allI, rtac impI]);
03a843f0f447 Ran expandshort
paulson
parents: 1973
diff changeset
    36
by (EVERY1[etac allE, etac allE, etac mp]);
03a843f0f447 Ran expandshort
paulson
parents: 1973
diff changeset
    37
by (EVERY1[etac allE, etac allE, etac mp, atac]);
2810
c4e16b36bc57 Added wp_while.
nipkow
parents: 2055
diff changeset
    38
qed_spec_mp "awp_mono";
1447
bc2c0acbbf29 Added a verified verification-condition generator.
nipkow
parents:
diff changeset
    39
bc2c0acbbf29 Added a verified verification-condition generator.
nipkow
parents:
diff changeset
    40
goal VC.thy "!P Q. (!s. P s --> Q s) --> (!s. vc c P s --> vc c Q s)";
2031
03a843f0f447 Ran expandshort
paulson
parents: 1973
diff changeset
    41
by (acom.induct_tac "c" 1);
2055
cc274e47f607 Ran expandshort
paulson
parents: 2031
diff changeset
    42
    by (ALLGOALS Asm_simp_tac);
2031
03a843f0f447 Ran expandshort
paulson
parents: 1973
diff changeset
    43
by (safe_tac HOL_cs);
03a843f0f447 Ran expandshort
paulson
parents: 1973
diff changeset
    44
by (EVERY[etac allE 1,etac allE 1,etac impE 1,etac allE 2,etac mp 2,atac 2]);
2810
c4e16b36bc57 Added wp_while.
nipkow
parents: 2055
diff changeset
    45
by (fast_tac (HOL_cs addEs [awp_mono]) 1);
1486
7b95d7b49f7a Introduced qed_spec_mp.
nipkow
parents: 1465
diff changeset
    46
qed_spec_mp "vc_mono";
1447
bc2c0acbbf29 Added a verified verification-condition generator.
nipkow
parents:
diff changeset
    47
bc2c0acbbf29 Added a verified verification-condition generator.
nipkow
parents:
diff changeset
    48
goal VC.thy
1743
f7feaacd33d3 Updated for new form of induction rules
paulson
parents: 1486
diff changeset
    49
  "!!P c Q. |- {P}c{Q} ==> \
2810
c4e16b36bc57 Added wp_while.
nipkow
parents: 2055
diff changeset
    50
\          (? ac. astrip ac = c & (!s. vc ac Q s) & (!s. P s --> awp ac Q s))";
1743
f7feaacd33d3 Updated for new form of induction rules
paulson
parents: 1486
diff changeset
    51
by (etac hoare.induct 1);
2055
cc274e47f607 Ran expandshort
paulson
parents: 2031
diff changeset
    52
     by (res_inst_tac [("x","Askip")] exI 1);
cc274e47f607 Ran expandshort
paulson
parents: 2031
diff changeset
    53
     by (Simp_tac 1);
cc274e47f607 Ran expandshort
paulson
parents: 2031
diff changeset
    54
    by (res_inst_tac [("x","Aass x a")] exI 1);
cc274e47f607 Ran expandshort
paulson
parents: 2031
diff changeset
    55
    by (Simp_tac 1);
cc274e47f607 Ran expandshort
paulson
parents: 2031
diff changeset
    56
   by (SELECT_GOAL(safe_tac HOL_cs)1);
cc274e47f607 Ran expandshort
paulson
parents: 2031
diff changeset
    57
   by (res_inst_tac [("x","Asemi ac aca")] exI 1);
cc274e47f607 Ran expandshort
paulson
parents: 2031
diff changeset
    58
   by (Asm_simp_tac 1);
2810
c4e16b36bc57 Added wp_while.
nipkow
parents: 2055
diff changeset
    59
   by (fast_tac (HOL_cs addSEs [awp_mono,vc_mono]) 1);
2055
cc274e47f607 Ran expandshort
paulson
parents: 2031
diff changeset
    60
  by (SELECT_GOAL(safe_tac HOL_cs)1);
cc274e47f607 Ran expandshort
paulson
parents: 2031
diff changeset
    61
  by (res_inst_tac [("x","Aif b ac aca")] exI 1);
cc274e47f607 Ran expandshort
paulson
parents: 2031
diff changeset
    62
  by (Asm_simp_tac 1);
cc274e47f607 Ran expandshort
paulson
parents: 2031
diff changeset
    63
 by (SELECT_GOAL(safe_tac HOL_cs)1);
cc274e47f607 Ran expandshort
paulson
parents: 2031
diff changeset
    64
 by (res_inst_tac [("x","Awhile b Pa ac")] exI 1);
cc274e47f607 Ran expandshort
paulson
parents: 2031
diff changeset
    65
 by (Asm_simp_tac 1);
1743
f7feaacd33d3 Updated for new form of induction rules
paulson
parents: 1486
diff changeset
    66
by (safe_tac HOL_cs);
2031
03a843f0f447 Ran expandshort
paulson
parents: 1973
diff changeset
    67
by (res_inst_tac [("x","ac")] exI 1);
03a843f0f447 Ran expandshort
paulson
parents: 1973
diff changeset
    68
by (Asm_simp_tac 1);
2810
c4e16b36bc57 Added wp_while.
nipkow
parents: 2055
diff changeset
    69
by (fast_tac (HOL_cs addSEs [awp_mono,vc_mono]) 1);
1743
f7feaacd33d3 Updated for new form of induction rules
paulson
parents: 1486
diff changeset
    70
qed "vc_complete";
1451
6803ecb9b198 Added vcwp
nipkow
parents: 1447
diff changeset
    71
2810
c4e16b36bc57 Added wp_while.
nipkow
parents: 2055
diff changeset
    72
goal VC.thy "!Q. vcawp c Q = (vc c Q, awp c Q)";
2031
03a843f0f447 Ran expandshort
paulson
parents: 1973
diff changeset
    73
by (acom.induct_tac "c" 1);
4089
96fba19bcbe2 isatool fixclasimp;
wenzelm
parents: 3842
diff changeset
    74
by (ALLGOALS (asm_simp_tac (simpset() addsimps [Let_def])));
2810
c4e16b36bc57 Added wp_while.
nipkow
parents: 2055
diff changeset
    75
qed "vcawp_vc_awp";