src/HOL/UNITY/Project.ML
author nipkow
Mon, 13 May 2002 15:27:28 +0200
changeset 13145 59bc43b51aa2
parent 11170 015af2fc7026
permissions -rw-r--r--
*** empty log message ***
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
     1
(*  Title:      HOL/UNITY/Project.ML
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
     2
    ID:         $Id$
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
     4
    Copyright   1999  University of Cambridge
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
     5
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
     6
Projections of state sets (also of actions and programs)
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
     7
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
     8
Inheritance of GUARANTEES properties under extension
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
     9
*)
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
    10
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
    11
Open_locale "Extend";
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
    12
7880
62fb24e28e5e exchanged the first two args of "project" and "drop_prog"
paulson
parents: 7878
diff changeset
    13
Goal "F : A co B ==> project h C (extend h F) : A co B";
7660
7e38237edfcb now with (weak safety) guarantees (weak progress) with Extend
paulson
parents: 7630
diff changeset
    14
by (auto_tac (claset(), 
7e38237edfcb now with (weak safety) guarantees (weak progress) with Extend
paulson
parents: 7630
diff changeset
    15
      simpset() addsimps [extend_act_def, project_act_def, constrains_def]));
7e38237edfcb now with (weak safety) guarantees (weak progress) with Extend
paulson
parents: 7630
diff changeset
    16
qed "project_extend_constrains_I";
7e38237edfcb now with (weak safety) guarantees (weak progress) with Extend
paulson
parents: 7630
diff changeset
    17
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
    18
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
    19
(** Safety **)
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
    20
7660
7e38237edfcb now with (weak safety) guarantees (weak progress) with Extend
paulson
parents: 7630
diff changeset
    21
(*used below to prove Join_project_ensures*)
7880
62fb24e28e5e exchanged the first two args of "project" and "drop_prog"
paulson
parents: 7878
diff changeset
    22
Goal "[| G : stable C;  project h C G : A unless B |] \
7660
7e38237edfcb now with (weak safety) guarantees (weak progress) with Extend
paulson
parents: 7630
diff changeset
    23
\     ==> G : (C Int extend_set h A) unless (extend_set h B)";
7e38237edfcb now with (weak safety) guarantees (weak progress) with Extend
paulson
parents: 7630
diff changeset
    24
by (asm_full_simp_tac
7e38237edfcb now with (weak safety) guarantees (weak progress) with Extend
paulson
parents: 7630
diff changeset
    25
    (simpset() addsimps [unless_def, project_constrains]) 1);
7e38237edfcb now with (weak safety) guarantees (weak progress) with Extend
paulson
parents: 7630
diff changeset
    26
by (blast_tac (claset() addDs [stable_constrains_Int]
7e38237edfcb now with (weak safety) guarantees (weak progress) with Extend
paulson
parents: 7630
diff changeset
    27
			addIs [constrains_weaken]) 1);
7e38237edfcb now with (weak safety) guarantees (weak progress) with Extend
paulson
parents: 7630
diff changeset
    28
qed_spec_mp "project_unless";
7e38237edfcb now with (weak safety) guarantees (weak progress) with Extend
paulson
parents: 7630
diff changeset
    29
9337
58bd51302b21 used bounded quantification in definition of guarantees and other minor
paulson
parents: 9190
diff changeset
    30
(*Generalizes project_constrains to the program F Join project h C G;
58bd51302b21 used bounded quantification in definition of guarantees and other minor
paulson
parents: 9190
diff changeset
    31
  useful with guarantees reasoning*)
7880
62fb24e28e5e exchanged the first two args of "project" and "drop_prog"
paulson
parents: 7878
diff changeset
    32
Goal "(F Join project h C G : A co B)  =  \
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
    33
\       (extend h F Join G : (C Int extend_set h A) co (extend_set h B) &  \
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
    34
\        F : A co B)";
9403
aad13b59b8d9 much tidying in connection with the 2nd UNITY paper
paulson
parents: 9337
diff changeset
    35
by (simp_tac (simpset() addsimps [project_constrains]) 1);
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
    36
by (blast_tac (claset() addIs [extend_constrains RS iffD2 RS constrains_weaken]
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
    37
                        addDs [constrains_imp_subset]) 1);
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
    38
qed "Join_project_constrains";
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
    39
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
    40
(*The condition is required to prove the left-to-right direction;
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
    41
  could weaken it to G : (C Int extend_set h A) co C*)
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
    42
Goalw [stable_def]
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
    43
     "extend h F Join G : stable C \
7880
62fb24e28e5e exchanged the first two args of "project" and "drop_prog"
paulson
parents: 7878
diff changeset
    44
\     ==> (F Join project h C G : stable A)  =  \
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
    45
\         (extend h F Join G : stable (C Int extend_set h A) &  \
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
    46
\          F : stable A)";
9403
aad13b59b8d9 much tidying in connection with the 2nd UNITY paper
paulson
parents: 9337
diff changeset
    47
by (simp_tac (HOL_ss addsimps [Join_project_constrains]) 1);
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
    48
by (blast_tac (claset() addIs [constrains_weaken] addDs [constrains_Int]) 1);
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
    49
qed "Join_project_stable";
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
    50
7689
affe0c2fdfbf working snapshot (even Alloc)
paulson
parents: 7660
diff changeset
    51
(*For using project_guarantees in particular cases*)
affe0c2fdfbf working snapshot (even Alloc)
paulson
parents: 7660
diff changeset
    52
Goal "extend h F Join G : extend_set h A co extend_set h B \
7880
62fb24e28e5e exchanged the first two args of "project" and "drop_prog"
paulson
parents: 7878
diff changeset
    53
\     ==> F Join project h C G : A co B";
7826
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
    54
by (asm_full_simp_tac
9403
aad13b59b8d9 much tidying in connection with the 2nd UNITY paper
paulson
parents: 9337
diff changeset
    55
    (simpset() addsimps [project_constrains, extend_constrains]) 1);
7826
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
    56
by (blast_tac (claset() addIs [constrains_weaken]
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
    57
			addDs [constrains_imp_subset]) 1);
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
    58
qed "project_constrains_I";
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
    59
8069
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8065
diff changeset
    60
Goalw [increasing_def, stable_def]
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8065
diff changeset
    61
     "extend h F Join G : increasing (func o f) \
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8065
diff changeset
    62
\     ==> F Join project h C G : increasing func";
9403
aad13b59b8d9 much tidying in connection with the 2nd UNITY paper
paulson
parents: 9337
diff changeset
    63
by (asm_full_simp_tac (simpset_of SubstAx.thy
aad13b59b8d9 much tidying in connection with the 2nd UNITY paper
paulson
parents: 9337
diff changeset
    64
		 addsimps [project_constrains_I, extend_set_eq_Collect]) 1);
8069
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8065
diff changeset
    65
qed "project_increasing_I";
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8065
diff changeset
    66
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8065
diff changeset
    67
Goal "(F Join project h UNIV G : increasing func)  =  \
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8065
diff changeset
    68
\     (extend h F Join G : increasing (func o f))";
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8065
diff changeset
    69
by (rtac iffI 1);
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8065
diff changeset
    70
by (etac project_increasing_I 2);
9403
aad13b59b8d9 much tidying in connection with the 2nd UNITY paper
paulson
parents: 9337
diff changeset
    71
by (asm_full_simp_tac (simpset_of SubstAx.thy
aad13b59b8d9 much tidying in connection with the 2nd UNITY paper
paulson
parents: 9337
diff changeset
    72
		         addsimps [increasing_def, Join_project_stable]) 1);
8069
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8065
diff changeset
    73
by (auto_tac (claset(),
9403
aad13b59b8d9 much tidying in connection with the 2nd UNITY paper
paulson
parents: 9337
diff changeset
    74
	      simpset() addsimps [extend_set_eq_Collect,
8069
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8065
diff changeset
    75
				  extend_stable RS iffD1]));
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8065
diff changeset
    76
qed "Join_project_increasing";
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8065
diff changeset
    77
7826
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
    78
(*The UNIV argument is essential*)
7880
62fb24e28e5e exchanged the first two args of "project" and "drop_prog"
paulson
parents: 7878
diff changeset
    79
Goal "F Join project h UNIV G : A co B \
7826
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
    80
\     ==> extend h F Join G : extend_set h A co extend_set h B";
7689
affe0c2fdfbf working snapshot (even Alloc)
paulson
parents: 7660
diff changeset
    81
by (asm_full_simp_tac
9403
aad13b59b8d9 much tidying in connection with the 2nd UNITY paper
paulson
parents: 9337
diff changeset
    82
    (simpset() addsimps [project_constrains, extend_constrains]) 1);
7826
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
    83
qed "project_constrains_D";
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
    84
8122
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
    85
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
    86
(*** "projecting" and union/intersection (no converses) ***)
7841
65d91d13fc13 working snapshot; more steps in Alloc
paulson
parents: 7826
diff changeset
    87
7826
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
    88
Goalw [projecting_def]
7841
65d91d13fc13 working snapshot; more steps in Alloc
paulson
parents: 7826
diff changeset
    89
     "[| projecting C h F XA' XA;  projecting C h F XB' XB |] \
65d91d13fc13 working snapshot; more steps in Alloc
paulson
parents: 7826
diff changeset
    90
\     ==> projecting C h F (XA' Int XB') (XA Int XB)";
65d91d13fc13 working snapshot; more steps in Alloc
paulson
parents: 7826
diff changeset
    91
by (Blast_tac 1);
65d91d13fc13 working snapshot; more steps in Alloc
paulson
parents: 7826
diff changeset
    92
qed "projecting_Int";
65d91d13fc13 working snapshot; more steps in Alloc
paulson
parents: 7826
diff changeset
    93
65d91d13fc13 working snapshot; more steps in Alloc
paulson
parents: 7826
diff changeset
    94
Goalw [projecting_def]
65d91d13fc13 working snapshot; more steps in Alloc
paulson
parents: 7826
diff changeset
    95
     "[| projecting C h F XA' XA;  projecting C h F XB' XB |] \
65d91d13fc13 working snapshot; more steps in Alloc
paulson
parents: 7826
diff changeset
    96
\     ==> projecting C h F (XA' Un XB') (XA Un XB)";
65d91d13fc13 working snapshot; more steps in Alloc
paulson
parents: 7826
diff changeset
    97
by (Blast_tac 1);
65d91d13fc13 working snapshot; more steps in Alloc
paulson
parents: 7826
diff changeset
    98
qed "projecting_Un";
65d91d13fc13 working snapshot; more steps in Alloc
paulson
parents: 7826
diff changeset
    99
65d91d13fc13 working snapshot; more steps in Alloc
paulson
parents: 7826
diff changeset
   100
val [prem] = Goalw [projecting_def]
65d91d13fc13 working snapshot; more steps in Alloc
paulson
parents: 7826
diff changeset
   101
     "[| !!i. i:I ==> projecting C h F (X' i) (X i) |] \
7826
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   102
\     ==> projecting C h F (INT i:I. X' i) (INT i:I. X i)";
7841
65d91d13fc13 working snapshot; more steps in Alloc
paulson
parents: 7826
diff changeset
   103
by (blast_tac (claset() addDs [prem RS spec RS mp]) 1);
7826
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   104
qed "projecting_INT";
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   105
7841
65d91d13fc13 working snapshot; more steps in Alloc
paulson
parents: 7826
diff changeset
   106
val [prem] = Goalw [projecting_def]
65d91d13fc13 working snapshot; more steps in Alloc
paulson
parents: 7826
diff changeset
   107
     "[| !!i. i:I ==> projecting C h F (X' i) (X i) |] \
7826
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   108
\     ==> projecting C h F (UN i:I. X' i) (UN i:I. X i)";
7841
65d91d13fc13 working snapshot; more steps in Alloc
paulson
parents: 7826
diff changeset
   109
by (blast_tac (claset() addDs [prem RS spec RS mp]) 1);
7826
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   110
qed "projecting_UN";
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   111
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   112
Goalw [projecting_def]
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   113
     "[| projecting C h F X' X;  U'<=X';  X<=U |] ==> projecting C h F U' U";
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   114
by Auto_tac;
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   115
qed "projecting_weaken";
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   116
8073
6c99b44b333e new weakening laws
paulson
parents: 8069
diff changeset
   117
Goalw [projecting_def]
6c99b44b333e new weakening laws
paulson
parents: 8069
diff changeset
   118
     "[| projecting C h F X' X;  U'<=X' |] ==> projecting C h F U' X";
6c99b44b333e new weakening laws
paulson
parents: 8069
diff changeset
   119
by Auto_tac;
6c99b44b333e new weakening laws
paulson
parents: 8069
diff changeset
   120
qed "projecting_weaken_L";
6c99b44b333e new weakening laws
paulson
parents: 8069
diff changeset
   121
7841
65d91d13fc13 working snapshot; more steps in Alloc
paulson
parents: 7826
diff changeset
   122
Goalw [extending_def]
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   123
     "[| extending C h F YA' YA;  extending C h F YB' YB |] \
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   124
\     ==> extending C h F (YA' Int YB') (YA Int YB)";
7841
65d91d13fc13 working snapshot; more steps in Alloc
paulson
parents: 7826
diff changeset
   125
by (Blast_tac 1);
65d91d13fc13 working snapshot; more steps in Alloc
paulson
parents: 7826
diff changeset
   126
qed "extending_Int";
65d91d13fc13 working snapshot; more steps in Alloc
paulson
parents: 7826
diff changeset
   127
7826
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   128
Goalw [extending_def]
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   129
     "[| extending C h F YA' YA;  extending C h F YB' YB |] \
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   130
\     ==> extending C h F (YA' Un YB') (YA Un YB)";
7841
65d91d13fc13 working snapshot; more steps in Alloc
paulson
parents: 7826
diff changeset
   131
by (Blast_tac 1);
65d91d13fc13 working snapshot; more steps in Alloc
paulson
parents: 7826
diff changeset
   132
qed "extending_Un";
65d91d13fc13 working snapshot; more steps in Alloc
paulson
parents: 7826
diff changeset
   133
65d91d13fc13 working snapshot; more steps in Alloc
paulson
parents: 7826
diff changeset
   134
val [prem] = Goalw [extending_def]
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   135
     "[| !!i. i:I ==> extending C h F (Y' i) (Y i) |] \
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   136
\     ==> extending C h F (INT i:I. Y' i) (INT i:I. Y i)";
7841
65d91d13fc13 working snapshot; more steps in Alloc
paulson
parents: 7826
diff changeset
   137
by (blast_tac (claset() addDs [prem RS spec RS mp]) 1);
7826
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   138
qed "extending_INT";
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   139
7841
65d91d13fc13 working snapshot; more steps in Alloc
paulson
parents: 7826
diff changeset
   140
val [prem] = Goalw [extending_def]
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   141
     "[| !!i. i:I ==> extending C h F (Y' i) (Y i) |] \
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   142
\     ==> extending C h F (UN i:I. Y' i) (UN i:I. Y i)";
7841
65d91d13fc13 working snapshot; more steps in Alloc
paulson
parents: 7826
diff changeset
   143
by (blast_tac (claset() addDs [prem RS spec RS mp]) 1);
7826
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   144
qed "extending_UN";
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   145
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   146
Goalw [extending_def]
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   147
     "[| extending C h F Y' Y;  Y'<=V';  V<=Y |] ==> extending C h F V' V";
7826
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   148
by Auto_tac;
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   149
qed "extending_weaken";
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   150
8073
6c99b44b333e new weakening laws
paulson
parents: 8069
diff changeset
   151
Goalw [extending_def]
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   152
     "[| extending C h F Y' Y;  Y'<=V' |] ==> extending C h F V' Y";
8073
6c99b44b333e new weakening laws
paulson
parents: 8069
diff changeset
   153
by Auto_tac;
6c99b44b333e new weakening laws
paulson
parents: 8069
diff changeset
   154
qed "extending_weaken_L";
6c99b44b333e new weakening laws
paulson
parents: 8069
diff changeset
   155
7826
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   156
Goal "projecting C h F X' UNIV";
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   157
by (simp_tac (simpset() addsimps [projecting_def]) 1);
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   158
qed "projecting_UNIV";
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   159
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   160
Goalw [projecting_def]
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   161
     "projecting C h F (extend_set h A co extend_set h B) (A co B)";
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   162
by (blast_tac (claset() addIs [project_constrains_I]) 1);
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   163
qed "projecting_constrains";
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   164
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   165
Goalw [stable_def]
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   166
     "projecting C h F (stable (extend_set h A)) (stable A)";
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   167
by (rtac projecting_constrains 1);
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   168
qed "projecting_stable";
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   169
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   170
Goalw [projecting_def]
8069
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8065
diff changeset
   171
     "projecting C h F (increasing (func o f)) (increasing func)";
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8065
diff changeset
   172
by (blast_tac (claset() addIs [project_increasing_I]) 1);
7826
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   173
qed "projecting_increasing";
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   174
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   175
Goal "extending C h F UNIV Y";
7826
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   176
by (simp_tac (simpset() addsimps [extending_def]) 1);
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   177
qed "extending_UNIV";
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   178
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   179
Goalw [extending_def]
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   180
     "extending (%G. UNIV) h F (extend_set h A co extend_set h B) (A co B)";
7826
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   181
by (blast_tac (claset() addIs [project_constrains_D]) 1);
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   182
qed "extending_constrains";
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   183
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   184
Goalw [stable_def]
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   185
     "extending (%G. UNIV) h F (stable (extend_set h A)) (stable A)";
7826
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   186
by (rtac extending_constrains 1);
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   187
qed "extending_stable";
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   188
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   189
Goalw [extending_def]
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   190
     "extending (%G. UNIV) h F (increasing (func o f)) (increasing func)";
9403
aad13b59b8d9 much tidying in connection with the 2nd UNITY paper
paulson
parents: 9337
diff changeset
   191
by (simp_tac (HOL_ss addsimps [Join_project_increasing]) 1);
7826
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   192
qed "extending_increasing";
7689
affe0c2fdfbf working snapshot (even Alloc)
paulson
parents: 7660
diff changeset
   193
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   194
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   195
(** Reachability and project **)
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   196
9337
58bd51302b21 used bounded quantification in definition of guarantees and other minor
paulson
parents: 9190
diff changeset
   197
(*In practice, C = reachable(...): the inclusion is equality*)
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   198
Goal "[| reachable (extend h F Join G) <= C;  \
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   199
\        z : reachable (extend h F Join G) |] \
7880
62fb24e28e5e exchanged the first two args of "project" and "drop_prog"
paulson
parents: 7878
diff changeset
   200
\     ==> f z : reachable (F Join project h C G)";
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   201
by (etac reachable.induct 1);
8069
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8065
diff changeset
   202
by (force_tac (claset() addSIs [reachable.Init],
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8065
diff changeset
   203
	       simpset() addsimps [split_extended_all]) 1);
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   204
by Auto_tac;
11170
015af2fc7026 simplified proofs for splitI and splitD, added splitD'
oheimb
parents: 10834
diff changeset
   205
by (force_tac (claset() delSWrapper "split_all_tac" addSbefore 
015af2fc7026 simplified proofs for splitI and splitD, added splitD'
oheimb
parents: 10834
diff changeset
   206
   ("unsafe_split_all_tac", unsafe_split_all_tac) 
015af2fc7026 simplified proofs for splitI and splitD, added splitD'
oheimb
parents: 10834
diff changeset
   207
   addIs [project_act_I RSN (3,reachable.Acts)], simpset()) 2);
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   208
by (res_inst_tac [("act","x")] reachable.Acts 1);
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   209
by Auto_tac;
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   210
by (etac extend_act_D 1);
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   211
qed "reachable_imp_reachable_project";
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   212
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   213
Goalw [Constrains_def]
8128
3a5864b465e2 still working; a bit of polishing
paulson
parents: 8122
diff changeset
   214
     "F Join project h (reachable (extend h F Join G)) G : A Co B  \
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   215
\     ==> extend h F Join G : (extend_set h A) Co (extend_set h B)";
9403
aad13b59b8d9 much tidying in connection with the 2nd UNITY paper
paulson
parents: 9337
diff changeset
   216
by (full_simp_tac (simpset_of SubstAx.thy addsimps [Join_project_constrains]) 1);
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   217
by (Clarify_tac 1);
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   218
by (etac constrains_weaken 1);
8128
3a5864b465e2 still working; a bit of polishing
paulson
parents: 8122
diff changeset
   219
by (auto_tac (claset() addIs [reachable_imp_reachable_project], simpset()));
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   220
qed "project_Constrains_D";
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   221
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   222
Goalw [Stable_def]
8128
3a5864b465e2 still working; a bit of polishing
paulson
parents: 8122
diff changeset
   223
     "F Join project h (reachable (extend h F Join G)) G : Stable A  \
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   224
\     ==> extend h F Join G : Stable (extend_set h A)";
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   225
by (asm_simp_tac (simpset() addsimps [project_Constrains_D]) 1);
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   226
qed "project_Stable_D";
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   227
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   228
Goalw [Always_def]
8128
3a5864b465e2 still working; a bit of polishing
paulson
parents: 8122
diff changeset
   229
     "F Join project h (reachable (extend h F Join G)) G : Always A  \
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   230
\     ==> extend h F Join G : Always (extend_set h A)";
8069
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8065
diff changeset
   231
by (force_tac (claset() addIs [reachable.Init],
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8065
diff changeset
   232
               simpset() addsimps [project_Stable_D, split_extended_all]) 1);
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   233
qed "project_Always_D";
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   234
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   235
Goalw [Increasing_def]
8128
3a5864b465e2 still working; a bit of polishing
paulson
parents: 8122
diff changeset
   236
     "F Join project h (reachable (extend h F Join G)) G : Increasing func  \
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   237
\     ==> extend h F Join G : Increasing (func o f)";
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   238
by Auto_tac;
8251
9be357df93d4 New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents: 8128
diff changeset
   239
by (stac (extend_set_eq_Collect RS sym) 1);
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   240
by (asm_simp_tac (simpset() addsimps [project_Stable_D]) 1); 
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   241
qed "project_Increasing_D";
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   242
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   243
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   244
(** Converse results for weak safety: benefits of the argument C *)
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   245
9337
58bd51302b21 used bounded quantification in definition of guarantees and other minor
paulson
parents: 9190
diff changeset
   246
(*In practice, C = reachable(...): the inclusion is equality*)
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   247
Goal "[| C <= reachable(extend h F Join G);   \
7880
62fb24e28e5e exchanged the first two args of "project" and "drop_prog"
paulson
parents: 7878
diff changeset
   248
\        x : reachable (F Join project h C G) |] \
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   249
\     ==> EX y. h(x,y) : reachable (extend h F Join G)";
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   250
by (etac reachable.induct 1);
11170
015af2fc7026 simplified proofs for splitI and splitD, added splitD'
oheimb
parents: 10834
diff changeset
   251
by  (force_tac (claset() addIs [reachable.Init], simpset()) 1);
015af2fc7026 simplified proofs for splitI and splitD, added splitD'
oheimb
parents: 10834
diff changeset
   252
by (auto_tac (claset(), simpset()addsimps [project_act_def]));
015af2fc7026 simplified proofs for splitI and splitD, added splitD'
oheimb
parents: 10834
diff changeset
   253
by (ALLGOALS (force_tac (claset() delrules [Id_in_Acts]
015af2fc7026 simplified proofs for splitI and splitD, added splitD'
oheimb
parents: 10834
diff changeset
   254
		        addIs [reachable.Acts, extend_act_D], simpset())));
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   255
qed "reachable_project_imp_reachable";
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   256
7660
7e38237edfcb now with (weak safety) guarantees (weak progress) with Extend
paulson
parents: 7630
diff changeset
   257
Goal "project_set h (reachable (extend h F Join G)) = \
7880
62fb24e28e5e exchanged the first two args of "project" and "drop_prog"
paulson
parents: 7878
diff changeset
   258
\     reachable (F Join project h (reachable (extend h F Join G)) G)";
7660
7e38237edfcb now with (weak safety) guarantees (weak progress) with Extend
paulson
parents: 7630
diff changeset
   259
by (auto_tac (claset() addDs [subset_refl RS reachable_imp_reachable_project,
7e38237edfcb now with (weak safety) guarantees (weak progress) with Extend
paulson
parents: 7630
diff changeset
   260
			      subset_refl RS reachable_project_imp_reachable], 
8069
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8065
diff changeset
   261
	      simpset()));
7660
7e38237edfcb now with (weak safety) guarantees (weak progress) with Extend
paulson
parents: 7630
diff changeset
   262
qed "project_set_reachable_extend_eq";
7e38237edfcb now with (weak safety) guarantees (weak progress) with Extend
paulson
parents: 7630
diff changeset
   263
9337
58bd51302b21 used bounded quantification in definition of guarantees and other minor
paulson
parents: 9190
diff changeset
   264
(*UNUSED*)
7947
b999c1ab9327 working again; new treatment of LocalTo
paulson
parents: 7915
diff changeset
   265
Goal "reachable (extend h F Join G) <= C  \
b999c1ab9327 working again; new treatment of LocalTo
paulson
parents: 7915
diff changeset
   266
\     ==> reachable (extend h F Join G) <= \
b999c1ab9327 working again; new treatment of LocalTo
paulson
parents: 7915
diff changeset
   267
\         extend_set h (reachable (F Join project h C G))";
b999c1ab9327 working again; new treatment of LocalTo
paulson
parents: 7915
diff changeset
   268
by (auto_tac (claset() addDs [reachable_imp_reachable_project], 
b999c1ab9327 working again; new treatment of LocalTo
paulson
parents: 7915
diff changeset
   269
	      simpset()));
b999c1ab9327 working again; new treatment of LocalTo
paulson
parents: 7915
diff changeset
   270
qed "reachable_extend_Join_subset";
b999c1ab9327 working again; new treatment of LocalTo
paulson
parents: 7915
diff changeset
   271
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   272
Goalw [Constrains_def]
8128
3a5864b465e2 still working; a bit of polishing
paulson
parents: 8122
diff changeset
   273
     "extend h F Join G : (extend_set h A) Co (extend_set h B)  \
3a5864b465e2 still working; a bit of polishing
paulson
parents: 8122
diff changeset
   274
\     ==> F Join project h (reachable (extend h F Join G)) G : A Co B";
9403
aad13b59b8d9 much tidying in connection with the 2nd UNITY paper
paulson
parents: 9337
diff changeset
   275
by (full_simp_tac (simpset_of SubstAx.thy addsimps [Join_project_constrains, 
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   276
				       extend_set_Int_distrib]) 1);
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   277
by (rtac conjI 1);
8128
3a5864b465e2 still working; a bit of polishing
paulson
parents: 8122
diff changeset
   278
by (force_tac
3a5864b465e2 still working; a bit of polishing
paulson
parents: 8122
diff changeset
   279
    (claset() addEs [constrains_weaken_L]
3a5864b465e2 still working; a bit of polishing
paulson
parents: 8122
diff changeset
   280
              addSDs [extend_constrains_project_set,
3a5864b465e2 still working; a bit of polishing
paulson
parents: 8122
diff changeset
   281
		      subset_refl RS reachable_project_imp_reachable], 
9403
aad13b59b8d9 much tidying in connection with the 2nd UNITY paper
paulson
parents: 9337
diff changeset
   282
     simpset()) 2);
8128
3a5864b465e2 still working; a bit of polishing
paulson
parents: 8122
diff changeset
   283
by (blast_tac (claset() addIs [constrains_weaken_L]) 1);
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   284
qed "project_Constrains_I";
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   285
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   286
Goalw [Stable_def]
8128
3a5864b465e2 still working; a bit of polishing
paulson
parents: 8122
diff changeset
   287
     "extend h F Join G : Stable (extend_set h A)  \
3a5864b465e2 still working; a bit of polishing
paulson
parents: 8122
diff changeset
   288
\     ==> F Join project h (reachable (extend h F Join G)) G : Stable A";
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   289
by (asm_simp_tac (simpset() addsimps [project_Constrains_I]) 1);
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   290
qed "project_Stable_I";
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   291
7689
affe0c2fdfbf working snapshot (even Alloc)
paulson
parents: 7660
diff changeset
   292
Goalw [Always_def]
8128
3a5864b465e2 still working; a bit of polishing
paulson
parents: 8122
diff changeset
   293
     "extend h F Join G : Always (extend_set h A)  \
3a5864b465e2 still working; a bit of polishing
paulson
parents: 8122
diff changeset
   294
\     ==> F Join project h (reachable (extend h F Join G)) G : Always A";
7689
affe0c2fdfbf working snapshot (even Alloc)
paulson
parents: 7660
diff changeset
   295
by (auto_tac (claset(), simpset() addsimps [project_Stable_I]));
8069
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8065
diff changeset
   296
by (rewtac extend_set_def);
7689
affe0c2fdfbf working snapshot (even Alloc)
paulson
parents: 7660
diff changeset
   297
by (Blast_tac 1);
affe0c2fdfbf working snapshot (even Alloc)
paulson
parents: 7660
diff changeset
   298
qed "project_Always_I";
affe0c2fdfbf working snapshot (even Alloc)
paulson
parents: 7660
diff changeset
   299
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   300
Goalw [Increasing_def]
8128
3a5864b465e2 still working; a bit of polishing
paulson
parents: 8122
diff changeset
   301
    "extend h F Join G : Increasing (func o f)  \
3a5864b465e2 still working; a bit of polishing
paulson
parents: 8122
diff changeset
   302
\    ==> F Join project h (reachable (extend h F Join G)) G : Increasing func";
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   303
by Auto_tac;
8251
9be357df93d4 New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents: 8128
diff changeset
   304
by (asm_simp_tac (simpset() addsimps [extend_set_eq_Collect,
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   305
				      project_Stable_I]) 1); 
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   306
qed "project_Increasing_I";
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   307
7880
62fb24e28e5e exchanged the first two args of "project" and "drop_prog"
paulson
parents: 7878
diff changeset
   308
Goal "(F Join project h (reachable (extend h F Join G)) G : A Co B)  =  \
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   309
\     (extend h F Join G : (extend_set h A) Co (extend_set h B))";
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   310
by (blast_tac (claset() addIs [project_Constrains_I, project_Constrains_D]) 1);
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   311
qed "project_Constrains";
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   312
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   313
Goalw [Stable_def]
7880
62fb24e28e5e exchanged the first two args of "project" and "drop_prog"
paulson
parents: 7878
diff changeset
   314
     "(F Join project h (reachable (extend h F Join G)) G : Stable A)  =  \
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   315
\     (extend h F Join G : Stable (extend_set h A))";
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   316
by (rtac project_Constrains 1);
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   317
qed "project_Stable";
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   318
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   319
Goal
7880
62fb24e28e5e exchanged the first two args of "project" and "drop_prog"
paulson
parents: 7878
diff changeset
   320
   "(F Join project h (reachable (extend h F Join G)) G : Increasing func)  = \
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   321
\   (extend h F Join G : Increasing (func o f))";
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   322
by (asm_simp_tac (simpset() addsimps [Increasing_def, project_Stable,
8251
9be357df93d4 New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents: 8128
diff changeset
   323
				      extend_set_eq_Collect]) 1);
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   324
qed "project_Increasing";
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   325
7826
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   326
(** A lot of redundant theorems: all are proved to facilitate reasoning
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   327
    about guarantees. **)
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   328
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   329
Goalw [projecting_def]
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   330
     "projecting (%G. reachable (extend h F Join G)) h F \
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   331
\                (extend_set h A Co extend_set h B) (A Co B)";
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   332
by (blast_tac (claset() addIs [project_Constrains_I]) 1);
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   333
qed "projecting_Constrains";
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   334
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   335
Goalw [Stable_def]
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   336
     "projecting (%G. reachable (extend h F Join G)) h F \
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   337
\                (Stable (extend_set h A)) (Stable A)";
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   338
by (rtac projecting_Constrains 1);
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   339
qed "projecting_Stable";
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   340
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   341
Goalw [projecting_def]
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   342
     "projecting (%G. reachable (extend h F Join G)) h F \
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   343
\                (Always (extend_set h A)) (Always A)";
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   344
by (blast_tac (claset() addIs [project_Always_I]) 1);
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   345
qed "projecting_Always";
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   346
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   347
Goalw [projecting_def]
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   348
     "projecting (%G. reachable (extend h F Join G)) h F \
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   349
\                (Increasing (func o f)) (Increasing func)";
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   350
by (blast_tac (claset() addIs [project_Increasing_I]) 1);
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   351
qed "projecting_Increasing";
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   352
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   353
Goalw [extending_def]
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   354
     "extending (%G. reachable (extend h F Join G)) h F \
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8041
diff changeset
   355
\                 (extend_set h A Co extend_set h B) (A Co B)";
7826
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   356
by (blast_tac (claset() addIs [project_Constrains_D]) 1);
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   357
qed "extending_Constrains";
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   358
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   359
Goalw [extending_def]
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   360
     "extending (%G. reachable (extend h F Join G)) h F \
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8041
diff changeset
   361
\                 (Stable (extend_set h A)) (Stable A)";
7826
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   362
by (blast_tac (claset() addIs [project_Stable_D]) 1);
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   363
qed "extending_Stable";
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   364
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   365
Goalw [extending_def]
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   366
     "extending (%G. reachable (extend h F Join G)) h F \
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8041
diff changeset
   367
\                 (Always (extend_set h A)) (Always A)";
7826
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   368
by (blast_tac (claset() addIs [project_Always_D]) 1);
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   369
qed "extending_Always";
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   370
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   371
Goalw [extending_def]
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   372
     "extending (%G. reachable (extend h F Join G)) h F \
8128
3a5864b465e2 still working; a bit of polishing
paulson
parents: 8122
diff changeset
   373
\                 (Increasing (func o f)) (Increasing func)";
3a5864b465e2 still working; a bit of polishing
paulson
parents: 8122
diff changeset
   374
by (blast_tac (claset() addIs [project_Increasing_D]) 1);
7826
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   375
qed "extending_Increasing";
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   376
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   377
8122
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   378
(*** leadsETo in the precondition (??) ***)
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   379
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   380
(** transient **)
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   381
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   382
Goalw [transient_def]
8110
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   383
     "[| G : transient (C Int extend_set h A);  G : stable C |]  \
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   384
\     ==> project h C G : transient (project_set h C Int A)";
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   385
by (auto_tac (claset(), simpset() addsimps [Domain_project_act]));
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10797
diff changeset
   386
by (subgoal_tac "act `` (C Int extend_set h A) <= - extend_set h A" 1);
8110
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   387
by (asm_full_simp_tac 
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   388
    (simpset() addsimps [stable_def, constrains_def]) 2);
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   389
by (Blast_tac 2);
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   390
(*back to main goal*)
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   391
by (thin_tac "?AA <= -C Un ?BB" 1);
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   392
by (ball_tac 1);
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   393
by (asm_full_simp_tac 
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   394
    (simpset() addsimps [extend_set_def, project_act_def]) 1);
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   395
by (Blast_tac 1);
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   396
qed "transient_extend_set_imp_project_transient";
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   397
8110
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   398
(*converse might hold too?*)
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   399
Goalw [transient_def]
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   400
     "project h C (extend h F) : transient (project_set h C Int D) \
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   401
\     ==> F : transient (project_set h C Int D)";
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   402
by (auto_tac (claset(), simpset() addsimps [Domain_project_act]));
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   403
by (rtac bexI 1);
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   404
by (assume_tac 2);
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   405
by Auto_tac;
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   406
by (rewtac extend_act_def);
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   407
by (Blast_tac 1);
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   408
qed "project_extend_transient_D";
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   409
8110
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   410
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   411
(** ensures -- a primitive combining progress with safety **)
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   412
8110
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   413
(*Used to prove project_leadsETo_I*)
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   414
Goal "[| extend h F : stable C;  G : stable C;  \
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   415
\        extend h F Join G : A ensures B;  A-B = C Int extend_set h D |]  \
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   416
\     ==> F Join project h C G  \
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   417
\           : (project_set h C Int project_set h A) ensures (project_set h B)";
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   418
by (asm_full_simp_tac
9403
aad13b59b8d9 much tidying in connection with the 2nd UNITY paper
paulson
parents: 9337
diff changeset
   419
    (simpset() addsimps [ensures_def, project_constrains, 
8110
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   420
			 Join_transient, extend_transient]) 1);
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   421
by (Clarify_tac 1);
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   422
by (REPEAT_FIRST (rtac conjI));
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   423
(*first subgoal*)
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   424
by (blast_tac (claset() addIs [extend_stable_project_set RS stableD RS 
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   425
			       constrains_Int RS constrains_weaken]
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   426
	                addSDs [extend_constrains_project_set]
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   427
			addSDs [equalityD1]) 1);
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   428
(*2nd subgoal*)
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   429
by (etac (stableD RS constrains_Int RS constrains_weaken) 1);
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   430
by (assume_tac 1);
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   431
by (Blast_tac 3);
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   432
by (full_simp_tac (simpset() addsimps [extend_set_Int_distrib,
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   433
				       extend_set_Un_distrib]) 2);
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   434
by (blast_tac (claset() addSIs [impOfSubs extend_set_project_set]) 2);
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   435
by (full_simp_tac (simpset() addsimps [extend_set_def]) 1);
9190
b86ff604729f tidied proofs using default rule equalityCE
paulson
parents: 9083
diff changeset
   436
by (Blast_tac 1);
8110
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   437
(*The transient part*)
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   438
by Auto_tac;
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   439
by (force_tac (claset() addSDs [equalityD1]
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   440
	                addIs [transient_extend_set_imp_project_transient RS
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   441
			       transient_strengthen], 
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   442
              simpset()) 2);
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   443
by (full_simp_tac (simpset() addsimps [Int_Diff]) 1);
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   444
by (force_tac (claset() addSDs [equalityD1]
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   445
	                addIs [transient_extend_set_imp_project_transient RS 
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   446
			 project_extend_transient_D RS transient_strengthen], 
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   447
              simpset()) 1);
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   448
qed "ensures_extend_set_imp_project_ensures";
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   449
8110
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   450
(*Used to prove project_leadsETo_D*)
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8041
diff changeset
   451
Goal "[| project h C G ~: transient (A-B) | A<=B;  \
7660
7e38237edfcb now with (weak safety) guarantees (weak progress) with Extend
paulson
parents: 7630
diff changeset
   452
\        extend h F Join G : stable C;  \
7880
62fb24e28e5e exchanged the first two args of "project" and "drop_prog"
paulson
parents: 7878
diff changeset
   453
\        F Join project h C G : A ensures B |] \
7660
7e38237edfcb now with (weak safety) guarantees (weak progress) with Extend
paulson
parents: 7630
diff changeset
   454
\     ==> extend h F Join G : (C Int extend_set h A) ensures (extend_set h B)";
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8041
diff changeset
   455
by (etac disjE 1);
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8041
diff changeset
   456
by (blast_tac (claset() addIs [subset_imp_ensures]) 2);
8002
paulson
parents: 7947
diff changeset
   457
by (auto_tac (claset() addDs [extend_transient RS iffD2] 
8041
e3237d8c18d6 working version with new theory ELT
paulson
parents: 8002
diff changeset
   458
                       addIs [transient_strengthen, project_set_I,
8002
paulson
parents: 7947
diff changeset
   459
			      project_unless RS unlessD, unlessI, 
paulson
parents: 7947
diff changeset
   460
			      project_extend_constrains_I], 
9403
aad13b59b8d9 much tidying in connection with the 2nd UNITY paper
paulson
parents: 9337
diff changeset
   461
	      simpset() addsimps [ensures_def, Join_transient]));
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   462
qed_spec_mp "Join_project_ensures";
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   463
8122
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   464
(** Lemma useful for both STRONG and WEAK progress, but the transient
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   465
    condition's very strong **)
7660
7e38237edfcb now with (weak safety) guarantees (weak progress) with Extend
paulson
parents: 7630
diff changeset
   466
7947
b999c1ab9327 working again; new treatment of LocalTo
paulson
parents: 7915
diff changeset
   467
(*The strange induction formula allows induction over the leadsTo
b999c1ab9327 working again; new treatment of LocalTo
paulson
parents: 7915
diff changeset
   468
  assumption's non-atomic precondition*)
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8041
diff changeset
   469
Goal "[| ALL D. project h C G : transient D --> D={};  \
7660
7e38237edfcb now with (weak safety) guarantees (weak progress) with Extend
paulson
parents: 7630
diff changeset
   470
\        extend h F Join G : stable C;  \
7880
62fb24e28e5e exchanged the first two args of "project" and "drop_prog"
paulson
parents: 7878
diff changeset
   471
\        F Join project h C G : (project_set h C Int A) leadsTo B |] \
7947
b999c1ab9327 working again; new treatment of LocalTo
paulson
parents: 7915
diff changeset
   472
\     ==> extend h F Join G : \
b999c1ab9327 working again; new treatment of LocalTo
paulson
parents: 7915
diff changeset
   473
\         C Int extend_set h (project_set h C Int A) leadsTo (extend_set h B)";
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   474
by (etac leadsTo_induct 1);
7660
7e38237edfcb now with (weak safety) guarantees (weak progress) with Extend
paulson
parents: 7630
diff changeset
   475
by (asm_simp_tac (simpset() delsimps UN_simps
7e38237edfcb now with (weak safety) guarantees (weak progress) with Extend
paulson
parents: 7630
diff changeset
   476
		  addsimps [Int_UN_distrib, leadsTo_UN, extend_set_Union]) 3);
8041
e3237d8c18d6 working version with new theory ELT
paulson
parents: 8002
diff changeset
   477
by (blast_tac (claset() addIs [psp_stable2 RS leadsTo_weaken_L, 
7660
7e38237edfcb now with (weak safety) guarantees (weak progress) with Extend
paulson
parents: 7630
diff changeset
   478
			       leadsTo_Trans]) 2);
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   479
by (blast_tac (claset() addIs [leadsTo_Basis, Join_project_ensures]) 1);
7660
7e38237edfcb now with (weak safety) guarantees (weak progress) with Extend
paulson
parents: 7630
diff changeset
   480
val lemma = result();
7e38237edfcb now with (weak safety) guarantees (weak progress) with Extend
paulson
parents: 7630
diff changeset
   481
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8041
diff changeset
   482
Goal "[| ALL D. project h C G : transient D --> D={};  \
7660
7e38237edfcb now with (weak safety) guarantees (weak progress) with Extend
paulson
parents: 7630
diff changeset
   483
\        extend h F Join G : stable C;  \
7880
62fb24e28e5e exchanged the first two args of "project" and "drop_prog"
paulson
parents: 7878
diff changeset
   484
\        F Join project h C G : (project_set h C Int A) leadsTo B |] \
7660
7e38237edfcb now with (weak safety) guarantees (weak progress) with Extend
paulson
parents: 7630
diff changeset
   485
\     ==> extend h F Join G : (C Int extend_set h A) leadsTo (extend_set h B)";
7689
affe0c2fdfbf working snapshot (even Alloc)
paulson
parents: 7660
diff changeset
   486
by (rtac (lemma RS leadsTo_weaken) 1);
8069
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8065
diff changeset
   487
by (auto_tac (claset(), simpset() addsimps [split_extended_all]));
8110
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   488
qed "project_leadsTo_D_lemma";
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   489
7660
7e38237edfcb now with (weak safety) guarantees (weak progress) with Extend
paulson
parents: 7630
diff changeset
   490
Goal "[| C = (reachable (extend h F Join G)); \
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8041
diff changeset
   491
\        ALL D. project h C G : transient D --> D={};  \
7880
62fb24e28e5e exchanged the first two args of "project" and "drop_prog"
paulson
parents: 7878
diff changeset
   492
\        F Join project h C G : A LeadsTo B |] \
7660
7e38237edfcb now with (weak safety) guarantees (weak progress) with Extend
paulson
parents: 7630
diff changeset
   493
\     ==> extend h F Join G : (extend_set h A) LeadsTo (extend_set h B)";
7e38237edfcb now with (weak safety) guarantees (weak progress) with Extend
paulson
parents: 7630
diff changeset
   494
by (asm_full_simp_tac 
9403
aad13b59b8d9 much tidying in connection with the 2nd UNITY paper
paulson
parents: 9337
diff changeset
   495
    (simpset_of SubstAx.thy addsimps [LeadsTo_def, project_leadsTo_D_lemma, 
7660
7e38237edfcb now with (weak safety) guarantees (weak progress) with Extend
paulson
parents: 7630
diff changeset
   496
			 project_set_reachable_extend_eq]) 1);
7e38237edfcb now with (weak safety) guarantees (weak progress) with Extend
paulson
parents: 7630
diff changeset
   497
qed "Join_project_LeadsTo";
7e38237edfcb now with (weak safety) guarantees (weak progress) with Extend
paulson
parents: 7630
diff changeset
   498
7e38237edfcb now with (weak safety) guarantees (weak progress) with Extend
paulson
parents: 7630
diff changeset
   499
8122
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   500
(*** Towards project_Ensures_D ***)
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   501
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   502
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   503
Goalw [project_set_def, extend_set_def, project_act_def]
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10797
diff changeset
   504
     "act `` (C Int extend_set h A) <= B \
a7897aebbffc *** empty log message ***
nipkow
parents: 10797
diff changeset
   505
\     ==> project_act h (Restrict C act) `` (project_set h C Int A) \
8122
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   506
\         <= project_set h B";
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   507
by (Blast_tac 1);
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   508
qed "act_subset_imp_project_act_subset";
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   509
9610
7dd6a1661bc9 major sharpening of stable_project_transient
paulson
parents: 9403
diff changeset
   510
(*This trivial proof is the complementation part of transferring a transient
7dd6a1661bc9 major sharpening of stable_project_transient
paulson
parents: 9403
diff changeset
   511
  property upwards.  The hard part would be to 
7dd6a1661bc9 major sharpening of stable_project_transient
paulson
parents: 9403
diff changeset
   512
  show that G's action has a big enough domain.*)
7dd6a1661bc9 major sharpening of stable_project_transient
paulson
parents: 9403
diff changeset
   513
Goal "[| act: Acts G;       \
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10797
diff changeset
   514
\        (project_act h (Restrict C act))`` \
9610
7dd6a1661bc9 major sharpening of stable_project_transient
paulson
parents: 9403
diff changeset
   515
\             (project_set h C Int A - B) <= -(project_set h C Int A - B) |] \
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10797
diff changeset
   516
\     ==> act``(C Int extend_set h A - extend_set h B) \
9610
7dd6a1661bc9 major sharpening of stable_project_transient
paulson
parents: 9403
diff changeset
   517
\           <= -(C Int extend_set h A - extend_set h B)"; 
7dd6a1661bc9 major sharpening of stable_project_transient
paulson
parents: 9403
diff changeset
   518
by (auto_tac (claset(), 
7dd6a1661bc9 major sharpening of stable_project_transient
paulson
parents: 9403
diff changeset
   519
     simpset() addsimps [project_set_def, extend_set_def, project_act_def]));  
7dd6a1661bc9 major sharpening of stable_project_transient
paulson
parents: 9403
diff changeset
   520
result();
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8041
diff changeset
   521
8122
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   522
Goal "[| G : stable ((C Int extend_set h A) - (extend_set h B));  \
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   523
\        project h C G : transient (project_set h C Int A - B) |]  \
9610
7dd6a1661bc9 major sharpening of stable_project_transient
paulson
parents: 9403
diff changeset
   524
\     ==> (C Int extend_set h A) - extend_set h B = {}";
8122
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   525
by (auto_tac (claset(), 
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   526
	      simpset() addsimps [transient_def, subset_Compl_self_eq,
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   527
				  Domain_project_act, split_extended_all]));
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   528
by (Blast_tac 1);
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   529
by (auto_tac (claset(), 
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   530
	      simpset() addsimps [stable_def, constrains_def]));
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   531
by (ball_tac 1);
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   532
by (auto_tac (claset(), 
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   533
	      simpset() addsimps [Int_Diff,
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   534
				  extend_set_Diff_distrib RS sym]));
8128
3a5864b465e2 still working; a bit of polishing
paulson
parents: 8122
diff changeset
   535
by (dtac act_subset_imp_project_act_subset 1);
8122
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   536
by (subgoal_tac
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10797
diff changeset
   537
    "project_act h (Restrict C act) `` (project_set h C Int (A - B)) = {}" 1);
a7897aebbffc *** empty log message ***
nipkow
parents: 10797
diff changeset
   538
by (REPEAT (thin_tac "?r``?A <= ?B" 1));
8986
paulson
parents: 8251
diff changeset
   539
by (rewrite_goals_tac [project_set_def, extend_set_def, project_act_def]);
8122
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   540
by (Blast_tac 2);
9083
b36787a56a1f this proof needs more detail
paulson
parents: 8986
diff changeset
   541
by (rtac ccontr 1);
b36787a56a1f this proof needs more detail
paulson
parents: 8986
diff changeset
   542
by (dtac subsetD 1);
9610
7dd6a1661bc9 major sharpening of stable_project_transient
paulson
parents: 9403
diff changeset
   543
by (Blast_tac 1);
9083
b36787a56a1f this proof needs more detail
paulson
parents: 8986
diff changeset
   544
by (force_tac (claset(), simpset() addsimps [split_extended_all]) 1);
8122
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   545
qed "stable_project_transient";
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   546
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   547
Goal "[| G : stable C;  project h C G : (project_set h C Int A) unless B |] \
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   548
\     ==> G : (C Int extend_set h A) unless (extend_set h B)";
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   549
by (auto_tac
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   550
    (claset() addDs [stable_constrains_Int]
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   551
              addIs [constrains_weaken],
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   552
     simpset() addsimps [unless_def, project_constrains, Diff_eq, 
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   553
			 Int_assoc, Int_extend_set_lemma]));
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   554
qed_spec_mp "project_unless2";
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8041
diff changeset
   555
8122
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   556
Goal "[| G : stable ((C Int extend_set h A) - (extend_set h B));  \
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   557
\        F Join project h C G : (project_set h C Int A) ensures B;  \
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   558
\        extend h F Join G : stable C |] \
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   559
\     ==> extend h F Join G : (C Int extend_set h A) ensures (extend_set h B)";
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   560
(*unless*)
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   561
by (auto_tac (claset() addSIs [rewrite_rule [unless_def] project_unless2] 
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   562
                       addIs [project_extend_constrains_I], 
9403
aad13b59b8d9 much tidying in connection with the 2nd UNITY paper
paulson
parents: 9337
diff changeset
   563
	      simpset() addsimps [ensures_def]));
8122
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   564
(*transient*)
9610
7dd6a1661bc9 major sharpening of stable_project_transient
paulson
parents: 9403
diff changeset
   565
(*A G-action cannot occur*)
7dd6a1661bc9 major sharpening of stable_project_transient
paulson
parents: 9403
diff changeset
   566
by (force_tac (claset() addDs [stable_project_transient], 
7dd6a1661bc9 major sharpening of stable_project_transient
paulson
parents: 9403
diff changeset
   567
                        simpset() delsimps [Diff_eq_empty_iff]
7dd6a1661bc9 major sharpening of stable_project_transient
paulson
parents: 9403
diff changeset
   568
				 addsimps [Diff_eq_empty_iff RS sym]) 2); 
7dd6a1661bc9 major sharpening of stable_project_transient
paulson
parents: 9403
diff changeset
   569
(*An F-action*)
8122
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   570
by (force_tac (claset() addSEs [extend_transient RS iffD2 RS 
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   571
				transient_strengthen], 
9610
7dd6a1661bc9 major sharpening of stable_project_transient
paulson
parents: 9403
diff changeset
   572
	       simpset() addsimps [split_extended_all]) 1);
8122
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   573
qed "project_ensures_D_lemma";
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8041
diff changeset
   574
8122
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   575
Goal "[| F Join project h UNIV G : A ensures B;  \
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   576
\        G : stable (extend_set h A - extend_set h B) |] \
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   577
\     ==> extend h F Join G : (extend_set h A) ensures (extend_set h B)";
8128
3a5864b465e2 still working; a bit of polishing
paulson
parents: 8122
diff changeset
   578
by (rtac (project_ensures_D_lemma RS revcut_rl) 1);
8122
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   579
by (stac stable_UNIV 3);
8128
3a5864b465e2 still working; a bit of polishing
paulson
parents: 8122
diff changeset
   580
by Auto_tac;
8122
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   581
qed "project_ensures_D";
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   582
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   583
Goalw [Ensures_def]
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   584
     "[| F Join project h (reachable (extend h F Join G)) G : A Ensures B;  \
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   585
\        G : stable (reachable (extend h F Join G) Int extend_set h A - \
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   586
\                    extend_set h B) |] \
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   587
\     ==> extend h F Join G : (extend_set h A) Ensures (extend_set h B)";
8128
3a5864b465e2 still working; a bit of polishing
paulson
parents: 8122
diff changeset
   588
by (rtac (project_ensures_D_lemma RS revcut_rl) 1);
8122
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   589
by (auto_tac (claset(), 
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   590
	      simpset() addsimps [project_set_reachable_extend_eq RS sym]));
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   591
qed "project_Ensures_D";
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   592
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   593
8251
9be357df93d4 New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents: 8128
diff changeset
   594
(*** Guarantees ***)
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   595
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   596
Goal "project_act h (Restrict C act) <= project_act h act";
11170
015af2fc7026 simplified proofs for splitI and splitD, added splitD'
oheimb
parents: 10834
diff changeset
   597
by (auto_tac (claset(), simpset() addsimps [project_act_def]));
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   598
qed "project_act_Restrict_subset_project_act";
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   599
					   
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   600
							   
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   601
Goal "[| extend h F ok G; subset_closed (AllowedActs F) |] \
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   602
\     ==> F ok project h C G";
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   603
by (auto_tac (claset(), simpset() addsimps [ok_def]));
11170
015af2fc7026 simplified proofs for splitI and splitD, added splitD'
oheimb
parents: 10834
diff changeset
   604
by (dtac subsetD 1);
015af2fc7026 simplified proofs for splitI and splitD, added splitD'
oheimb
parents: 10834
diff changeset
   605
by (Blast_tac 1);
015af2fc7026 simplified proofs for splitI and splitD, added splitD'
oheimb
parents: 10834
diff changeset
   606
by (force_tac (claset() delSWrapper "split_all_tac" addSbefore 
015af2fc7026 simplified proofs for splitI and splitD, added splitD'
oheimb
parents: 10834
diff changeset
   607
                    ("unsafe_split_all_tac", unsafe_split_all_tac) 
015af2fc7026 simplified proofs for splitI and splitD, added splitD'
oheimb
parents: 10834
diff changeset
   608
                     addSIs [rev_image_eqI], simpset()) 1);
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   609
by (cut_facts_tac [project_act_Restrict_subset_project_act] 1);
11170
015af2fc7026 simplified proofs for splitI and splitD, added splitD'
oheimb
parents: 10834
diff changeset
   610
by (auto_tac (claset(), simpset() addsimps [subset_closed_def]));
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   611
qed "subset_closed_ok_extend_imp_ok_project";
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   612
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   613
8251
9be357df93d4 New treatment of "guarantees" with polymorphic components and bijections.
paulson
parents: 8128
diff changeset
   614
(*Weak precondition and postcondition
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   615
  Not clear that it has a converse [or that we want one!]*)
7841
65d91d13fc13 working snapshot; more steps in Alloc
paulson
parents: 7826
diff changeset
   616
9337
58bd51302b21 used bounded quantification in definition of guarantees and other minor
paulson
parents: 9190
diff changeset
   617
(*The raw version; 3rd premise could be weakened by adding the
58bd51302b21 used bounded quantification in definition of guarantees and other minor
paulson
parents: 9190
diff changeset
   618
  precondition extend h F Join G : X' *)
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   619
val [xguary,closed,project,extend] =
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   620
Goal "[| F : X guarantees Y;  subset_closed (AllowedActs F);  \
8122
b43ad07660b9 working version, with Alloc now working on the same state space as the whole
paulson
parents: 8110
diff changeset
   621
\        !!G. extend h F Join G : X' ==> F Join project h (C G) G : X;  \
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   622
\        !!G. [| F Join project h (C G) G : Y |] \
7841
65d91d13fc13 working snapshot; more steps in Alloc
paulson
parents: 7826
diff changeset
   623
\             ==> extend h F Join G : Y' |] \
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   624
\     ==> extend h F : X' guarantees Y'";
7841
65d91d13fc13 working snapshot; more steps in Alloc
paulson
parents: 7826
diff changeset
   625
by (rtac (xguary RS guaranteesD RS extend RS guaranteesI) 1);
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   626
by (blast_tac (claset() addIs [closed, 
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   627
          subset_closed_ok_extend_imp_ok_project]) 1); 
7841
65d91d13fc13 working snapshot; more steps in Alloc
paulson
parents: 7826
diff changeset
   628
by (etac project 1);
8041
e3237d8c18d6 working version with new theory ELT
paulson
parents: 8002
diff changeset
   629
qed "project_guarantees_raw";
7841
65d91d13fc13 working snapshot; more steps in Alloc
paulson
parents: 7826
diff changeset
   630
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   631
Goal "[| F : X guarantees Y;  subset_closed (AllowedActs F); \
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   632
\        projecting C h F X' X;  extending C h F Y' Y |] \
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   633
\     ==> extend h F : X' guarantees Y'";
7826
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   634
by (rtac guaranteesI 1);
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   635
by (auto_tac (claset(), 
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   636
        simpset() addsimps [guaranteesD, projecting_def, 
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   637
                    extending_def, subset_closed_ok_extend_imp_ok_project]));
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   638
qed "project_guarantees";
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   639
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8041
diff changeset
   640
7841
65d91d13fc13 working snapshot; more steps in Alloc
paulson
parents: 7826
diff changeset
   641
(*It seems that neither "guarantees" law can be proved from the other.*)
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   642
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   643
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   644
(*** guarantees corollaries ***)
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   645
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8041
diff changeset
   646
(** Some could be deleted: the required versions are easy to prove **)
7826
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   647
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   648
Goal "[| F : UNIV guarantees increasing func;  \
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   649
\        subset_closed (AllowedActs F) |] \
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   650
\     ==> extend h F : X' guarantees increasing (func o f)";
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   651
by (etac project_guarantees 1);
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   652
by (rtac extending_increasing 3);
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   653
by (rtac projecting_UNIV 2);
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8041
diff changeset
   654
by Auto_tac;
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   655
qed "extend_guar_increasing";
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   656
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   657
Goal "[| F : UNIV guarantees Increasing func;  \
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   658
\        subset_closed (AllowedActs F) |] \
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   659
\     ==> extend h F : X' guarantees Increasing (func o f)";
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   660
by (etac project_guarantees 1);
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   661
by (rtac extending_Increasing 3);
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   662
by (rtac projecting_UNIV 2);
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   663
by Auto_tac;
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   664
qed "extend_guar_Increasing";
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   665
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   666
Goal "[| F : Always A guarantees Always B;  \
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   667
\        subset_closed (AllowedActs F) |] \
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   668
\     ==> extend h F                   \
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   669
\           : Always(extend_set h A) guarantees Always(extend_set h B)";
7689
affe0c2fdfbf working snapshot (even Alloc)
paulson
parents: 7660
diff changeset
   670
by (etac project_guarantees 1);
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   671
by (rtac extending_Always 3);
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   672
by (rtac projecting_Always 2);
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8041
diff changeset
   673
by Auto_tac;
7689
affe0c2fdfbf working snapshot (even Alloc)
paulson
parents: 7660
diff changeset
   674
qed "extend_guar_Always";
affe0c2fdfbf working snapshot (even Alloc)
paulson
parents: 7660
diff changeset
   675
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8041
diff changeset
   676
Goal "[| G : preserves f;  project h C G : transient D |] ==> D={}";
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8041
diff changeset
   677
by (rtac stable_transient_empty 1);
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8041
diff changeset
   678
by (assume_tac 2);
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8041
diff changeset
   679
by (blast_tac (claset() addIs [project_preserves_id_I,
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8041
diff changeset
   680
			 impOfSubs preserves_id_subset_stable]) 1);
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8041
diff changeset
   681
qed "preserves_project_transient_empty";
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8041
diff changeset
   682
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   683
8069
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8065
diff changeset
   684
(** Guarantees with a leadsTo postcondition 
19b9f92ca503 working with weak LeadsTo in guarantees precondition\!
paulson
parents: 8065
diff changeset
   685
    THESE ARE ALL TOO WEAK because G can't affect F's variables at all**)
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   686
7880
62fb24e28e5e exchanged the first two args of "project" and "drop_prog"
paulson
parents: 7878
diff changeset
   687
Goal "[| F Join project h UNIV G : A leadsTo B;    \
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8041
diff changeset
   688
\        G : preserves f |]  \
7689
affe0c2fdfbf working snapshot (even Alloc)
paulson
parents: 7660
diff changeset
   689
\     ==> extend h F Join G : (extend_set h A) leadsTo (extend_set h B)";
8110
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   690
by (res_inst_tac [("C1", "UNIV")] 
f7651ede12b7 moved some proofs from UNITY/ELT to UNITY/Project
paulson
parents: 8073
diff changeset
   691
    (project_leadsTo_D_lemma RS leadsTo_weaken) 1);
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8041
diff changeset
   692
by (auto_tac (claset() addDs [preserves_project_transient_empty], 
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8041
diff changeset
   693
	      simpset()));
7689
affe0c2fdfbf working snapshot (even Alloc)
paulson
parents: 7660
diff changeset
   694
qed "project_leadsTo_D";
affe0c2fdfbf working snapshot (even Alloc)
paulson
parents: 7660
diff changeset
   695
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   696
Goal "[| F Join project h (reachable (extend h F Join G)) G : A LeadsTo B;    \
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   697
\        G : preserves f |]  \
7689
affe0c2fdfbf working snapshot (even Alloc)
paulson
parents: 7660
diff changeset
   698
\      ==> extend h F Join G : (extend_set h A) LeadsTo (extend_set h B)";
7826
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   699
by (rtac (refl RS Join_project_LeadsTo) 1);
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8041
diff changeset
   700
by (auto_tac (claset() addDs [preserves_project_transient_empty], 
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8041
diff changeset
   701
	      simpset()));
7689
affe0c2fdfbf working snapshot (even Alloc)
paulson
parents: 7660
diff changeset
   702
qed "project_LeadsTo_D";
affe0c2fdfbf working snapshot (even Alloc)
paulson
parents: 7660
diff changeset
   703
7826
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   704
Goalw [extending_def]
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   705
     "(ALL G. extend h F ok G --> G : preserves f) \
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   706
\     ==> extending (%G. UNIV) h F \
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8041
diff changeset
   707
\                 (extend_set h A leadsTo extend_set h B) (A leadsTo B)";
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8041
diff changeset
   708
by (blast_tac (claset() addIs [project_leadsTo_D]) 1);
7826
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   709
qed "extending_leadsTo";
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   710
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   711
Goalw [extending_def]
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   712
     "(ALL G. extend h F ok G --> G : preserves f) \
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 9610
diff changeset
   713
\     ==> extending (%G. reachable (extend h F Join G)) h F \
8055
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8041
diff changeset
   714
\                 (extend_set h A LeadsTo extend_set h B) (A LeadsTo B)";
bb15396278fb abolition of localTo: instead "guarantees" has local vars as extra argument
paulson
parents: 8041
diff changeset
   715
by (blast_tac (claset() addIs [project_LeadsTo_D]) 1);
7826
c6a8b73b6c2a working shapshot with "projecting" and "extending"
paulson
parents: 7689
diff changeset
   716
qed "extending_LeadsTo";
7689
affe0c2fdfbf working snapshot (even Alloc)
paulson
parents: 7660
diff changeset
   717
7630
d0e4a6f1f05c working snapshot with new theory "Project"
paulson
parents:
diff changeset
   718
Close_locale "Extend";