src/HOL/UNITY/Union.ML
author paulson
Fri, 06 Nov 1998 13:20:29 +0100
changeset 5804 8e0a4c4fd67b
parent 5785 e58bb0f57c0c
child 5867 1c4806b4bf43
permissions -rw-r--r--
Revising the Client proof as suggested by Michel Charpentier. New lemmas
about composition (in Union.ML), etc. Also changed "length" to "size"
because it is displayed as "size" in any event.
paulson@5252
     1
(*  Title:      HOL/UNITY/Union.ML
paulson@5252
     2
    ID:         $Id$
paulson@5252
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
paulson@5252
     4
    Copyright   1998  University of Cambridge
paulson@5252
     5
paulson@5252
     6
Unions of programs
paulson@5252
     7
paulson@5252
     8
From Misra's Chapter 5: Asynchronous Compositions of Programs
paulson@5252
     9
*)
paulson@5252
    10
paulson@5252
    11
paulson@5648
    12
Goal "Init SKIP = UNIV";
paulson@5648
    13
by (simp_tac (simpset() addsimps [SKIP_def]) 1);
paulson@5611
    14
qed "Init_SKIP";
paulson@5611
    15
paulson@5648
    16
Goal "Acts SKIP = {Id}";
paulson@5648
    17
by (simp_tac (simpset() addsimps [SKIP_def]) 1);
paulson@5648
    18
qed "Acts_SKIP";
paulson@5648
    19
paulson@5648
    20
Addsimps [Init_SKIP, Acts_SKIP];
paulson@5648
    21
paulson@5611
    22
Goal "Init (F Join G) = Init F Int Init G";
paulson@5313
    23
by (simp_tac (simpset() addsimps [Join_def]) 1);
paulson@5313
    24
qed "Init_Join";
paulson@5252
    25
paulson@5611
    26
Goal "Acts (F Join G) = Acts F Un Acts G";
paulson@5596
    27
by (auto_tac (claset(), simpset() addsimps [Join_def]));
paulson@5313
    28
qed "Acts_Join";
paulson@5252
    29
paulson@5611
    30
Goal "Init (JN i:I. F i) = (INT i:I. Init (F i))";
paulson@5313
    31
by (simp_tac (simpset() addsimps [JOIN_def]) 1);
paulson@5313
    32
qed "Init_JN";
paulson@5252
    33
paulson@5611
    34
Goal "Acts (JN i:I. F i) = insert Id (UN i:I. Acts (F i))";
paulson@5596
    35
by (auto_tac (claset(), simpset() addsimps [JOIN_def]));
paulson@5313
    36
qed "Acts_JN";
paulson@5252
    37
paulson@5313
    38
Addsimps [Init_Join, Init_JN];
paulson@5259
    39
paulson@5611
    40
Goal "(JN x:insert a A. B x) = (B a) Join (JN x:A. B x)";
paulson@5611
    41
by (rtac program_equalityI 1);
paulson@5596
    42
by (ALLGOALS (simp_tac (simpset() addsimps [JOIN_def, Join_def])));
paulson@5584
    43
qed "JN_insert";
paulson@5584
    44
Addsimps[JN_insert];
paulson@5584
    45
paulson@5584
    46
paulson@5611
    47
(** Algebraic laws **)
paulson@5259
    48
paulson@5611
    49
Goal "F Join G = G Join F";
paulson@5313
    50
by (simp_tac (simpset() addsimps [Join_def, Un_commute, Int_commute]) 1);
paulson@5259
    51
qed "Join_commute";
paulson@5259
    52
paulson@5611
    53
Goal "(F Join G) Join H = F Join (G Join H)";
paulson@5596
    54
by (simp_tac (simpset() addsimps Un_ac@[Join_def, Int_assoc]) 1);
paulson@5259
    55
qed "Join_assoc";
paulson@5596
    56
 
paulson@5611
    57
Goalw [Join_def, SKIP_def] "SKIP Join F = F";
paulson@5611
    58
by (rtac program_equalityI 1);
paulson@5596
    59
by (ALLGOALS (simp_tac (simpset() addsimps [insert_absorb])));
paulson@5611
    60
qed "Join_SKIP_left";
paulson@5584
    61
paulson@5611
    62
Goalw [Join_def, SKIP_def] "F Join SKIP = F";
paulson@5611
    63
by (rtac program_equalityI 1);
paulson@5611
    64
by (ALLGOALS (simp_tac (simpset() addsimps [insert_absorb])));
paulson@5611
    65
qed "Join_SKIP_right";
paulson@5611
    66
paulson@5611
    67
Addsimps [Join_SKIP_left, Join_SKIP_right];
paulson@5611
    68
paulson@5611
    69
Goalw [Join_def] "F Join F = F";
paulson@5611
    70
by (rtac program_equalityI 1);
paulson@5313
    71
by Auto_tac;
paulson@5313
    72
qed "Join_absorb";
paulson@5313
    73
paulson@5611
    74
Addsimps [Join_absorb];
paulson@5313
    75
paulson@5584
    76
Goal "(JN G:{}. G) = SKIP";
paulson@5584
    77
by (simp_tac (simpset() addsimps [JOIN_def, SKIP_def]) 1);
paulson@5584
    78
qed "JN_empty";
paulson@5584
    79
Addsimps [JN_empty];
paulson@5584
    80
paulson@5584
    81
paulson@5584
    82
paulson@5584
    83
paulson@5313
    84
(*** Safety: constrains, stable, FP ***)
paulson@5313
    85
paulson@5313
    86
Goalw [constrains_def, JOIN_def]
paulson@5584
    87
    "I ~= {} ==> \
paulson@5648
    88
\    (JN i:I. F i) : constrains A B = (ALL i:I. F i : constrains A B)";
paulson@5596
    89
by (Simp_tac 1);
paulson@5596
    90
by (Blast_tac 1);
paulson@5313
    91
qed "constrains_JN";
paulson@5313
    92
paulson@5648
    93
Goal "F Join G : constrains A B =  \
paulson@5648
    94
\     (F : constrains A B & G : constrains A B)";
paulson@5620
    95
by (auto_tac
paulson@5620
    96
    (claset(),
paulson@5620
    97
     simpset() addsimps [constrains_def, Join_def]));
paulson@5620
    98
qed "constrains_Join";
paulson@5620
    99
paulson@5313
   100
(**FAILS, I think; see 5.4.1, Substitution Axiom.
paulson@5620
   101
   The problem is to relate reachable (F Join G) with 
paulson@5620
   102
   reachable F and reachable G
paulson@5620
   103
paulson@5313
   104
Goalw [Constrains_def]
paulson@5648
   105
    "(JN i:I. F i) : Constrains A B = (ALL i:I. F i : Constrains A B)";
paulson@5313
   106
by (simp_tac (simpset() addsimps [constrains_JN]) 1);
paulson@5313
   107
by (Blast_tac 1);
paulson@5313
   108
qed "Constrains_JN";
paulson@5313
   109
paulson@5648
   110
Goal "F Join G : Constrains A B = (Constrains F A B & Constrains G A B)";
paulson@5584
   111
by (auto_tac
paulson@5584
   112
    (claset(),
paulson@5620
   113
     simpset() addsimps [Constrains_def, constrains_Join]));
paulson@5620
   114
qed "Constrains_Join";
paulson@5620
   115
**)
paulson@5313
   116
paulson@5648
   117
Goal "[| F : constrains A A';  G : constrains B B' |] \
paulson@5648
   118
\     ==> F Join G : constrains (A Int B) (A' Un B')";
paulson@5313
   119
by (simp_tac (simpset() addsimps [constrains_Join]) 1);
paulson@5313
   120
by (blast_tac (claset() addIs [constrains_weaken]) 1);
paulson@5313
   121
qed "constrains_Join_weaken";
paulson@5313
   122
paulson@5584
   123
Goal "I ~= {} ==> \
paulson@5648
   124
\     (JN i:I. F i) : stable A = (ALL i:I. F i : stable A)";
paulson@5584
   125
by (asm_simp_tac (simpset() addsimps [stable_def, constrains_JN]) 1);
paulson@5313
   126
qed "stable_JN";
paulson@5313
   127
paulson@5648
   128
Goal "F Join G : stable A = \
paulson@5648
   129
\     (F : stable A & G : stable A)";
paulson@5313
   130
by (simp_tac (simpset() addsimps [stable_def, constrains_Join]) 1);
paulson@5313
   131
qed "stable_Join";
paulson@5313
   132
paulson@5648
   133
Goal "I ~= {} ==> FP (JN i:I. F i) = (INT i:I. FP (F i))";
paulson@5584
   134
by (asm_simp_tac (simpset() addsimps [FP_def, stable_JN, INTER_def]) 1);
paulson@5313
   135
qed "FP_JN";
paulson@5313
   136
paulson@5313
   137
paulson@5313
   138
(*** Progress: transient, ensures ***)
paulson@5313
   139
paulson@5584
   140
Goal "I ~= {} ==> \
paulson@5648
   141
\   (JN i:I. F i) : transient A = (EX i:I. F i : transient A)";
paulson@5584
   142
by (auto_tac (claset(),
paulson@5584
   143
	      simpset() addsimps [transient_def, JOIN_def]));
paulson@5313
   144
qed "transient_JN";
paulson@5313
   145
paulson@5648
   146
Goal "F Join G : transient A = \
paulson@5648
   147
\     (F : transient A | G : transient A)";
paulson@5584
   148
by (auto_tac (claset(),
paulson@5596
   149
	      simpset() addsimps [bex_Un, transient_def,
paulson@5584
   150
				  Join_def]));
paulson@5313
   151
qed "transient_Join";
paulson@5313
   152
paulson@5584
   153
Goal "I ~= {} ==> \
paulson@5648
   154
\     (JN i:I. F i) : ensures A B = \
paulson@5648
   155
\     ((ALL i:I. F i : constrains (A-B) (A Un B)) & \
paulson@5648
   156
\      (EX i:I. F i : ensures A B))";
paulson@5313
   157
by (auto_tac (claset(),
paulson@5584
   158
	      simpset() addsimps [ensures_def, constrains_JN, transient_JN]));
paulson@5313
   159
qed "ensures_JN";
paulson@5313
   160
paulson@5313
   161
Goalw [ensures_def]
paulson@5648
   162
     "F Join G : ensures A B =     \
paulson@5648
   163
\     (F : constrains (A-B) (A Un B) & \
paulson@5648
   164
\      G : constrains (A-B) (A Un B) & \
paulson@5648
   165
\      (F : ensures A B | G : ensures A B))";
paulson@5313
   166
by (auto_tac (claset(),
paulson@5313
   167
	      simpset() addsimps [constrains_Join, transient_Join]));
paulson@5313
   168
qed "ensures_Join";
paulson@5313
   169
paulson@5313
   170
Goalw [stable_def, constrains_def, Join_def]
paulson@5648
   171
    "[| F : stable A;  G : constrains A A' |] \
paulson@5648
   172
\    ==> F Join G : constrains A A'";
paulson@5648
   173
by (asm_full_simp_tac (simpset() addsimps [ball_Un]) 1);
paulson@5313
   174
by (Blast_tac 1);
paulson@5313
   175
qed "stable_constrains_Join";
paulson@5313
   176
paulson@5648
   177
(*Premise for G cannot use Invariant because  Stable F A  is weaker than
paulson@5648
   178
  G : stable A *)
paulson@5648
   179
Goal "[| F : stable A;  G : invariant A |] ==> F Join G : Invariant A";
paulson@5648
   180
by (full_simp_tac (simpset() addsimps [Invariant_def, invariant_def, 
paulson@5648
   181
				       Stable_eq_stable, stable_Join]) 1);
paulson@5313
   182
by (force_tac(claset() addIs [stable_reachable, stable_Int],
paulson@5313
   183
	      simpset() addsimps [Acts_Join]) 1);
paulson@5313
   184
qed "stable_Join_Invariant";
paulson@5313
   185
paulson@5648
   186
Goal "[| F : stable A;  G : ensures A B |] ==> F Join G : ensures A B";
paulson@5313
   187
by (asm_simp_tac (simpset() addsimps [ensures_Join]) 1);
paulson@5313
   188
by (asm_full_simp_tac (simpset() addsimps [stable_def, ensures_def]) 1);
paulson@5313
   189
by (etac constrains_weaken 1);
paulson@5313
   190
by Auto_tac;
paulson@5313
   191
qed "ensures_stable_Join1";
paulson@5313
   192
paulson@5313
   193
(*As above, but exchanging the roles of F and G*)
paulson@5648
   194
Goal "[| F : ensures A B;  G : stable A |] ==> F Join G : ensures A B";
paulson@5313
   195
by (stac Join_commute 1);
paulson@5313
   196
by (blast_tac (claset() addIs [ensures_stable_Join1]) 1);
paulson@5313
   197
qed "ensures_stable_Join2";
paulson@5313
   198
paulson@5648
   199
paulson@5804
   200
(** Diff and localTo **)
paulson@5648
   201
paulson@5804
   202
Goalw [Join_def, Diff_def] "F Join (Diff G (Acts F)) = F Join G";
paulson@5804
   203
by (rtac program_equalityI 1);
paulson@5804
   204
by Auto_tac;
paulson@5804
   205
qed "Join_Diff2";
paulson@5804
   206
paulson@5804
   207
Goalw [Diff_def, Disjoint_def] "Disjoint F (Diff G (Acts F))";
paulson@5804
   208
by Auto_tac;
paulson@5804
   209
qed "Diff_Disjoint";
paulson@5804
   210
paulson@5804
   211
Goal "[| F Join G : v localTo F;  Disjoint F G |] \
paulson@5804
   212
\     ==> G : (INT z. stable {s. v s = z})";
paulson@5648
   213
by (asm_full_simp_tac 
paulson@5804
   214
    (simpset() addsimps [localTo_def, Diff_def, Disjoint_def,
paulson@5804
   215
			 Acts_Join, stable_def, constrains_def]) 1);
paulson@5804
   216
by (Blast_tac 1);
paulson@5804
   217
qed "localTo_imp_stable";
paulson@5804
   218
paulson@5804
   219
Goal "[| F Join G : v localTo F;  (s,s') : act;  \
paulson@5804
   220
\        act : Acts G;  Disjoint F G |] ==> v s' = v s";
paulson@5804
   221
by (asm_full_simp_tac 
paulson@5804
   222
    (simpset() addsimps [localTo_def, Diff_def, Disjoint_def,
paulson@5804
   223
			 Acts_Join, stable_def, constrains_def]) 1);
paulson@5648
   224
by (Blast_tac 1);
paulson@5648
   225
qed "localTo_imp_equals";
paulson@5804
   226
paulson@5804
   227
Goalw [localTo_def, stable_def, constrains_def]
paulson@5804
   228
     "v localTo F <= (f o v) localTo F";
paulson@5804
   229
by (Clarify_tac 1);
paulson@5804
   230
by (auto_tac (claset() addSEs [allE, ballE], simpset()));
paulson@5804
   231
qed "localTo_imp_o_localTo";
paulson@5804
   232
paulson@5804
   233
paulson@5804
   234
(*** Higher-level rules involving localTo and Join ***)
paulson@5804
   235
paulson@5804
   236
Goal "[| F : constrains {s. P (v s) (w s)} {s. P' (v s) (w s)};   \
paulson@5804
   237
\        F Join G : v localTo F;       \
paulson@5804
   238
\        F Join G : w localTo F;       \
paulson@5804
   239
\        Disjoint F G |]               \
paulson@5804
   240
\     ==> F Join G : constrains {s. P (v s) (w s)} {s. P' (v s) (w s)}";
paulson@5804
   241
by (auto_tac (claset(), simpset() addsimps [constrains_def, Acts_Join]));
paulson@5804
   242
by (REPEAT_FIRST (dtac localTo_imp_equals THEN' REPEAT1 o atac));
paulson@5804
   243
by Auto_tac;
paulson@5804
   244
qed "constrains_localTo_constrains2";
paulson@5804
   245
paulson@5804
   246
Goalw [stable_def]
paulson@5804
   247
     "[| F : stable {s. P (v s) (w s)};   \
paulson@5804
   248
\        F Join G : v localTo F;       \
paulson@5804
   249
\        F Join G : w localTo F;       \
paulson@5804
   250
\        Disjoint F G |]               \
paulson@5804
   251
\     ==> F Join G : stable {s. P (v s) (w s)}";
paulson@5804
   252
by (blast_tac (claset() addIs [constrains_localTo_constrains2]) 1);
paulson@5804
   253
qed "stable_localTo_stable2";
paulson@5804
   254
paulson@5804
   255
paulson@5804
   256
Goal "(UN k. {s. f s = k}) = UNIV";
paulson@5804
   257
by (Blast_tac 1);
paulson@5804
   258
qed "UN_eq_UNIV";
paulson@5804
   259
paulson@5804
   260
Goal "[| F : stable {s. v s <= w s};   \
paulson@5804
   261
\        F Join G : v localTo F;       \
paulson@5804
   262
\        F Join G : Increasing w;      \
paulson@5804
   263
\        Disjoint F G |]               \
paulson@5804
   264
\     ==> F Join G : Stable {s. v s <= w s}";
paulson@5804
   265
by (safe_tac (claset() addSDs [localTo_imp_stable]));
paulson@5804
   266
by (rewrite_goals_tac [stable_def, Stable_def, Increasing_def]);
paulson@5804
   267
by (subgoal_tac "ALL k: UNIV. ?H : Constrains ({s. v s = k} Int ?AA) ?AA" 1);
paulson@5804
   268
by (dtac ball_Constrains_UN 1);
paulson@5804
   269
by (full_simp_tac (simpset() addsimps [UN_eq_UNIV]) 1);
paulson@5804
   270
by (rtac ballI 1);
paulson@5804
   271
by (subgoal_tac "F Join G : constrains ({s. v s = k} Int {s. v s <= w s}) \
paulson@5804
   272
\                                      ({s. v s = k} Un {s. v s <= w s})" 1);
paulson@5804
   273
by (asm_simp_tac (simpset() addsimps [constrains_Join]) 2);
paulson@5804
   274
by (blast_tac (claset() addIs [constrains_weaken]) 2);
paulson@5804
   275
by (dtac (constrains_imp_Constrains RS Constrains_Int) 1 THEN etac INT_D 1);
paulson@5804
   276
by (etac Constrains_weaken 2);
paulson@5804
   277
by Auto_tac;
paulson@5804
   278
qed "Increasing_localTo_Stable";