src/HOL/UNITY/Guar.ML
author paulson
Fri, 12 May 2000 15:05:02 +0200
changeset 8862 78643f8449c6
parent 8251 9be357df93d4
child 9337 58bd51302b21
permissions -rw-r--r--
NatSimprocs is now a theory, not a file
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
7400
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
     1
(*  Title:      HOL/UNITY/Guar.ML
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
     2
    ID:         $Id$
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
     4
    Copyright   1999  University of Cambridge
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
     5
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
     6
Guarantees, etc.
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
     7
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
     8
From Chandy and Sanders, "Reasoning About Program Composition"
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
     9
*)
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    10
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    11
(*** existential properties ***)
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    12
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    13
Goalw [ex_prop_def]
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    14
     "[| ex_prop X; finite GG |] ==> GG Int X ~= {} --> (JN G:GG. G) : X";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    15
by (etac finite_induct 1);
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    16
by (auto_tac (claset(), simpset() addsimps [Int_insert_left]));
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    17
qed_spec_mp "ex1";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    18
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    19
Goalw [ex_prop_def]
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    20
     "ALL GG. finite GG & GG Int X ~= {} --> (JN G:GG. G) : X ==> ex_prop X";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    21
by (Clarify_tac 1);
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    22
by (dres_inst_tac [("x", "{F,G}")] spec 1);
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    23
by Auto_tac;
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    24
qed "ex2";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    25
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    26
(*Chandy & Sanders take this as a definition*)
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    27
Goal "ex_prop X = (ALL GG. finite GG & GG Int X ~= {} --> (JN G:GG. G) : X)";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    28
by (blast_tac (claset() addIs [ex1,ex2]) 1);
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    29
qed "ex_prop_finite";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    30
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    31
(*Their "equivalent definition" given at the end of section 3*)
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    32
Goal "ex_prop X = (ALL G. G:X = (ALL H. G <= H --> H: X))";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    33
by Auto_tac;
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    34
by (rewrite_goals_tac [ex_prop_def, component_def]);
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    35
by (Blast_tac 1);
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    36
by Safe_tac;
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    37
by (stac Join_commute 2);
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    38
by (ALLGOALS Blast_tac);
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    39
qed "ex_prop_equiv";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    40
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    41
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    42
(*** universal properties ***)
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    43
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    44
Goalw [uv_prop_def]
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    45
     "[| uv_prop X; finite GG |] ==> GG <= X --> (JN G:GG. G) : X";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    46
by (etac finite_induct 1);
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    47
by (auto_tac (claset(), simpset() addsimps [Int_insert_left]));
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    48
qed_spec_mp "uv1";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    49
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    50
Goalw [uv_prop_def]
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    51
     "ALL GG. finite GG & GG <= X --> (JN G:GG. G) : X  ==> uv_prop X";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    52
by (rtac conjI 1);
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    53
by (Clarify_tac 2);
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    54
by (dres_inst_tac [("x", "{F,G}")] spec 2);
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    55
by (dres_inst_tac [("x", "{}")] spec 1);
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    56
by Auto_tac;
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    57
qed "uv2";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    58
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    59
(*Chandy & Sanders take this as a definition*)
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    60
Goal "uv_prop X = (ALL GG. finite GG & GG <= X --> (JN G:GG. G) : X)";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    61
by (blast_tac (claset() addIs [uv1,uv2]) 1);
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    62
qed "uv_prop_finite";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    63
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    64
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    65
(*** guarantees ***)
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    66
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    67
val prems = Goal
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
    68
     "(!!G. [| G : preserves v; F Join G : X |] ==> F Join G : Y) \
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
    69
\     ==> F : X guarantees[v] Y";
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
    70
by (simp_tac (simpset() addsimps [guar_def, component_def]) 1);
7400
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    71
by (blast_tac (claset() addIs prems) 1);
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    72
qed "guaranteesI";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    73
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    74
Goalw [guar_def, component_def]
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
    75
     "[| F : X guarantees[v] Y;  G : preserves v;  F Join G : X |] \
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
    76
\     ==> F Join G : Y";
7400
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    77
by (Blast_tac 1);
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    78
qed "guaranteesD";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    79
8065
658e0d4e4ed9 first working version to Alloc/System_Client_Progress;
paulson
parents: 8055
diff changeset
    80
(*This version of guaranteesD matches more easily in the conclusion
658e0d4e4ed9 first working version to Alloc/System_Client_Progress;
paulson
parents: 8055
diff changeset
    81
  The major premise can no longer be  F<=H since we need to reason about G*)
7400
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    82
Goalw [guar_def]
8122
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8065
diff changeset
    83
     "[| F : X guarantees[v] Y;  F Join G = H;  H : X;  G : preserves v |] \
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
    84
\     ==> H : Y";
7400
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    85
by (Blast_tac 1);
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    86
qed "component_guaranteesD";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    87
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    88
Goalw [guar_def]
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
    89
     "[| F: X guarantees[v] X'; Y <= X; X' <= Y' |] ==> F: Y guarantees[v] Y'";
7400
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    90
by (Blast_tac 1);
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    91
qed "guarantees_weaken";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    92
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    93
Goalw [guar_def]
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
    94
     "[| F: X guarantees[v] Y;  preserves w <= preserves v |] \
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
    95
\     ==> F: X guarantees[w] Y";
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
    96
by (Blast_tac 1);
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
    97
qed "guarantees_weaken_var";
7400
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
    98
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
    99
Goalw [guar_def] "X <= Y ==> X guarantees[v] Y = UNIV";
7400
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   100
by (Blast_tac 1);
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   101
qed "subset_imp_guarantees_UNIV";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   102
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   103
(*Equivalent to subset_imp_guarantees_UNIV but more intuitive*)
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   104
Goalw [guar_def] "X <= Y ==> F : X guarantees[v] Y";
7400
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   105
by (Blast_tac 1);
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   106
qed "subset_imp_guarantees";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   107
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   108
(*Remark at end of section 4.1
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   109
Goalw [guar_def] "ex_prop Y = (Y = UNIV guarantees[v] Y)";
7400
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   110
by (simp_tac (simpset() addsimps [ex_prop_equiv]) 1);
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   111
by (blast_tac (claset() addEs [equalityE]) 1);
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   112
qed "ex_prop_equiv2";
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   113
*)
7400
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   114
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   115
(** Distributive laws.  Re-orient to perform miniscoping **)
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   116
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   117
Goalw [guar_def]
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   118
     "(UN i:I. X i) guarantees[v] Y = (INT i:I. X i guarantees[v] Y)";
7400
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   119
by (Blast_tac 1);
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   120
qed "guarantees_UN_left";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   121
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   122
Goalw [guar_def]
8251
9be357df93d4 New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents: 8122
diff changeset
   123
     "(X Un Y) guarantees[v] Z = (X guarantees[v] Z) Int (Y guarantees[v] Z)";
7400
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   124
by (Blast_tac 1);
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   125
qed "guarantees_Un_left";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   126
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   127
Goalw [guar_def]
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   128
     "X guarantees[v] (INT i:I. Y i) = (INT i:I. X guarantees[v] Y i)";
7400
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   129
by (Blast_tac 1);
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   130
qed "guarantees_INT_right";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   131
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   132
Goalw [guar_def]
8251
9be357df93d4 New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents: 8122
diff changeset
   133
     "Z guarantees[v] (X Int Y) = (Z guarantees[v] X) Int (Z guarantees[v] Y)";
7400
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   134
by (Blast_tac 1);
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   135
qed "guarantees_Int_right";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   136
8251
9be357df93d4 New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents: 8122
diff changeset
   137
Goal "[| F : Z guarantees[v] X;  F : Z guarantees[v] Y |] \
9be357df93d4 New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents: 8122
diff changeset
   138
\    ==> F : Z guarantees[v] (X Int Y)";
9be357df93d4 New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents: 8122
diff changeset
   139
by (asm_simp_tac (simpset() addsimps [guarantees_Int_right]) 1);
9be357df93d4 New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents: 8122
diff changeset
   140
qed "guarantees_Int_right_I";
9be357df93d4 New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents: 8122
diff changeset
   141
8065
658e0d4e4ed9 first working version to Alloc/System_Client_Progress;
paulson
parents: 8055
diff changeset
   142
Goal "(F : X guarantees[v] (INTER I Y)) = \
658e0d4e4ed9 first working version to Alloc/System_Client_Progress;
paulson
parents: 8055
diff changeset
   143
\     (ALL i:I. F : X guarantees[v] (Y i))";
658e0d4e4ed9 first working version to Alloc/System_Client_Progress;
paulson
parents: 8055
diff changeset
   144
by (simp_tac (simpset() addsimps [guarantees_INT_right]) 1);
658e0d4e4ed9 first working version to Alloc/System_Client_Progress;
paulson
parents: 8055
diff changeset
   145
qed "guarantees_INT_right_iff";
658e0d4e4ed9 first working version to Alloc/System_Client_Progress;
paulson
parents: 8055
diff changeset
   146
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   147
Goalw [guar_def] "(X guarantees[v] Y) = (UNIV guarantees[v] (-X Un Y))";
7400
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   148
by (Blast_tac 1);
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   149
qed "shunting";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   150
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   151
Goalw [guar_def] "(X guarantees[v] Y) = -Y guarantees[v] -X";
7400
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   152
by (Blast_tac 1);
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   153
qed "contrapositive";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   154
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   155
(** The following two can be expressed using intersection and subset, which
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   156
    is more faithful to the text but looks cryptic.
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   157
**)
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   158
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   159
Goalw [guar_def]
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   160
    "[| F : V guarantees[v] X;  F : (X Int Y) guarantees[v] Z |]\
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   161
\    ==> F : (V Int Y) guarantees[v] Z";
7400
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   162
by (Blast_tac 1);
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   163
qed "combining1";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   164
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   165
Goalw [guar_def]
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   166
    "[| F : V guarantees[v] (X Un Y);  F : Y guarantees[v] Z |]\
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   167
\    ==> F : V guarantees[v] (X Un Z)";
7400
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   168
by (Blast_tac 1);
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   169
qed "combining2";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   170
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   171
(** The following two follow Chandy-Sanders, but the use of object-quantifiers
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   172
    does not suit Isabelle... **)
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   173
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   174
(*Premise should be (!!i. i: I ==> F: X guarantees[v] Y i) *)
7400
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   175
Goalw [guar_def]
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   176
     "ALL i:I. F : X guarantees[v] (Y i) ==> F : X guarantees[v] (INT i:I. Y i)";
7400
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   177
by (Blast_tac 1);
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   178
qed "all_guarantees";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   179
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   180
(*Premises should be [| F: X guarantees[v] Y i; i: I |] *)
7400
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   181
Goalw [guar_def]
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   182
     "EX i:I. F : X guarantees[v] (Y i) ==> F : X guarantees[v] (UN i:I. Y i)";
7400
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   183
by (Blast_tac 1);
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   184
qed "ex_guarantees";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   185
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   186
(*** Additional guarantees laws, by lcp ***)
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   187
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   188
Goalw [guar_def]
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   189
    "[| F: U guarantees[v] V;  G: X guarantees[v] Y; \
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   190
\       F : preserves v;  G : preserves v |] \
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   191
\    ==> F Join G: (U Int X) guarantees[v] (V Int Y)";
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   192
by (Simp_tac 1);
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   193
by Safe_tac;
8065
658e0d4e4ed9 first working version to Alloc/System_Client_Progress;
paulson
parents: 8055
diff changeset
   194
by (asm_full_simp_tac (simpset() addsimps [Join_assoc]) 1);
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   195
by (subgoal_tac "F Join G Join Ga = G Join (F Join Ga)" 1);
8065
658e0d4e4ed9 first working version to Alloc/System_Client_Progress;
paulson
parents: 8055
diff changeset
   196
by (Asm_full_simp_tac 1);
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   197
by (asm_simp_tac (simpset() addsimps Join_ac) 1);
7400
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   198
qed "guarantees_Join_Int";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   199
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   200
Goalw [guar_def]
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   201
    "[| F: U guarantees[v] V;  G: X guarantees[v] Y; \
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   202
\       F : preserves v;  G : preserves v |]  \
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   203
\    ==> F Join G: (U Un X) guarantees[v] (V Un Y)";
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   204
by (Simp_tac 1);
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   205
by Safe_tac;
8065
658e0d4e4ed9 first working version to Alloc/System_Client_Progress;
paulson
parents: 8055
diff changeset
   206
by (asm_full_simp_tac (simpset() addsimps [Join_assoc]) 1);
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   207
by (subgoal_tac "F Join G Join Ga = G Join (F Join Ga)" 1);
8065
658e0d4e4ed9 first working version to Alloc/System_Client_Progress;
paulson
parents: 8055
diff changeset
   208
by (Asm_full_simp_tac 1);
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   209
by (asm_simp_tac (simpset() addsimps Join_ac) 1);
7400
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   210
qed "guarantees_Join_Un";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   211
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   212
Goalw [guar_def]
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   213
    "[| ALL i:I. F i : X i guarantees[v] Y i;  \
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   214
\       ALL i:I. F i : preserves v |] \
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   215
\    ==> (JOIN I F) : (INTER I X) guarantees[v] (INTER I Y)";
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   216
by Auto_tac;
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   217
by (ball_tac 1);
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   218
by (dres_inst_tac [("x", "JOIN I F Join G")] spec 1);
8065
658e0d4e4ed9 first working version to Alloc/System_Client_Progress;
paulson
parents: 8055
diff changeset
   219
by (asm_full_simp_tac (simpset() addsimps [Join_assoc RS sym, JN_absorb]) 1);
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   220
qed "guarantees_JN_INT";
7400
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   221
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   222
Goalw [guar_def]
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   223
    "[| ALL i:I. F i : X i guarantees[v] Y i;  \
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   224
\       ALL i:I. F i : preserves v |] \
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   225
\    ==> (JOIN I F) : (UNION I X) guarantees[v] (UNION I Y)";
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   226
by Auto_tac;
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   227
by (ball_tac 1);
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   228
by (dres_inst_tac [("x", "JOIN I F Join G")] spec 1);
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   229
by (auto_tac
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   230
    (claset(),
8065
658e0d4e4ed9 first working version to Alloc/System_Client_Progress;
paulson
parents: 8055
diff changeset
   231
     simpset() addsimps [Join_assoc RS sym, JN_absorb]));
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   232
qed "guarantees_JN_UN";
7400
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   233
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   234
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   235
(*** guarantees[v] laws for breaking down the program, by lcp ***)
7400
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   236
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   237
Goalw [guar_def]
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   238
     "[| F: X guarantees[v] Y;  G: preserves v |] \
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   239
\     ==> F Join G: X guarantees[v] Y";
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   240
by (Simp_tac 1);
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   241
by Safe_tac;
8065
658e0d4e4ed9 first working version to Alloc/System_Client_Progress;
paulson
parents: 8055
diff changeset
   242
by (asm_full_simp_tac (simpset() addsimps [Join_assoc]) 1);
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   243
qed "guarantees_Join_I1";
7400
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   244
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   245
Goal "[| G: X guarantees[v] Y;  F: preserves v |] \
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   246
\    ==> F Join G: X guarantees[v] Y";
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   247
by (stac Join_commute 1);
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   248
by (blast_tac (claset() addIs [guarantees_Join_I1]) 1);
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   249
qed "guarantees_Join_I2";
7689
affe0c2fdfbf working snapshot (even Alloc)
paulson
parents: 7537
diff changeset
   250
7400
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   251
Goalw [guar_def]
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   252
     "[| i : I;  F i: X guarantees[v] Y;  \
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   253
\        ALL j:I. i~=j --> F j : preserves v |] \
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   254
\     ==> (JN i:I. (F i)) : X guarantees[v] Y";
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   255
by (Clarify_tac 1);
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   256
by (dres_inst_tac [("x", "JOIN (I-{i}) F Join G")] spec 1);
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 7947
diff changeset
   257
by (auto_tac (claset(),
8065
658e0d4e4ed9 first working version to Alloc/System_Client_Progress;
paulson
parents: 8055
diff changeset
   258
	      simpset() addsimps [JN_Join_diff, Join_assoc RS sym]));
7400
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   259
qed "guarantees_JN_I";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   260
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   261
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   262
(*** well-definedness ***)
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   263
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   264
Goalw [welldef_def] "F Join G: welldef ==> F: welldef";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   265
by Auto_tac;
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   266
qed "Join_welldef_D1";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   267
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   268
Goalw [welldef_def] "F Join G: welldef ==> G: welldef";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   269
by Auto_tac;
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   270
qed "Join_welldef_D2";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   271
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   272
(*** refinement ***)
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   273
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   274
Goalw [refines_def] "F refines F wrt X";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   275
by (Blast_tac 1);
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   276
qed "refines_refl";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   277
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   278
Goalw [refines_def]
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   279
     "[| H refines G wrt X;  G refines F wrt X |] ==> H refines F wrt X";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   280
by Auto_tac;
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   281
qed "refines_trans";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   282
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   283
Goalw [strict_ex_prop_def]
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   284
     "strict_ex_prop X \
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   285
\     ==> (ALL H. F Join H : X --> G Join H : X) = (F:X --> G:X)";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   286
by (Blast_tac 1);
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   287
qed "strict_ex_refine_lemma";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   288
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   289
Goalw [strict_ex_prop_def]
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   290
     "strict_ex_prop X \
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   291
\     ==> (ALL H. F Join H : welldef & F Join H : X --> G Join H : X) = \
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   292
\         (F: welldef Int X --> G:X)";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   293
by Safe_tac;
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   294
by (eres_inst_tac [("x","SKIP"), ("P", "%H. ?PP H --> ?RR H")] allE 1);
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   295
by (auto_tac (claset() addDs [Join_welldef_D1, Join_welldef_D2], simpset()));
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   296
qed "strict_ex_refine_lemma_v";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   297
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   298
Goal "[| strict_ex_prop X;  \
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   299
\        ALL H. F Join H : welldef Int X --> G Join H : welldef |] \
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   300
\     ==> (G refines F wrt X) = (G iso_refines F wrt X)";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   301
by (res_inst_tac [("x","SKIP")] allE 1
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   302
    THEN assume_tac 1);
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   303
by (asm_full_simp_tac
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   304
    (simpset() addsimps [refines_def, iso_refines_def,
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   305
			 strict_ex_refine_lemma_v]) 1);
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   306
qed "ex_refinement_thm";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   307
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   308
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   309
Goalw [strict_uv_prop_def]
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   310
     "strict_uv_prop X \
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   311
\     ==> (ALL H. F Join H : X --> G Join H : X) = (F:X --> G:X)";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   312
by (Blast_tac 1);
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   313
qed "strict_uv_refine_lemma";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   314
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   315
Goalw [strict_uv_prop_def]
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   316
     "strict_uv_prop X \
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   317
\     ==> (ALL H. F Join H : welldef & F Join H : X --> G Join H : X) = \
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   318
\         (F: welldef Int X --> G:X)";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   319
by Safe_tac;
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   320
by (eres_inst_tac [("x","SKIP"), ("P", "%H. ?PP H --> ?RR H")] allE 1);
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   321
by (auto_tac (claset() addDs [Join_welldef_D1, Join_welldef_D2],
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   322
	      simpset()));
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   323
qed "strict_uv_refine_lemma_v";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   324
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   325
Goal "[| strict_uv_prop X;  \
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   326
\        ALL H. F Join H : welldef Int X --> G Join H : welldef |] \
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   327
\     ==> (G refines F wrt X) = (G iso_refines F wrt X)";
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   328
by (res_inst_tac [("x","SKIP")] allE 1
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   329
    THEN assume_tac 1);
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   330
by (asm_full_simp_tac (simpset() addsimps [refines_def, iso_refines_def,
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   331
					   strict_uv_refine_lemma_v]) 1);
fbd5582761e6 new files HOL/UNITY/Guar.{thy,ML}: theory file gets the instance declaration
paulson
parents:
diff changeset
   332
qed "uv_refinement_thm";