src/HOL/UNITY/Union.ML
author paulson
Thu, 29 Apr 1999 10:51:58 +0200
changeset 6536 281d44905cab
parent 6309 ca52347e259a
child 6570 a7d7985050a9
permissions -rw-r--r--
made many specification operators infix
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@6012
    11
Addcongs [UN_cong, INT_cong];
paulson@5252
    12
paulson@6295
    13
paulson@5867
    14
(** SKIP **)
paulson@5867
    15
paulson@6295
    16
Goal "Init SKIP = UNIV";
paulson@5648
    17
by (simp_tac (simpset() addsimps [SKIP_def]) 1);
paulson@5611
    18
qed "Init_SKIP";
paulson@5611
    19
paulson@6295
    20
Goal "Acts SKIP = {Id}";
paulson@5648
    21
by (simp_tac (simpset() addsimps [SKIP_def]) 1);
paulson@5648
    22
qed "Acts_SKIP";
paulson@5648
    23
paulson@6295
    24
Addsimps [Init_SKIP, Acts_SKIP];
paulson@5648
    25
paulson@6295
    26
Goal "reachable SKIP = UNIV";
paulson@6012
    27
by (force_tac (claset() addEs [reachable.induct]
paulson@6012
    28
			addIs reachable.intrs, simpset()) 1);
paulson@5867
    29
qed "reachable_SKIP";
paulson@5867
    30
paulson@5867
    31
Addsimps [reachable_SKIP];
paulson@5867
    32
paulson@5867
    33
paulson@6012
    34
(** Join **)
paulson@6012
    35
paulson@6295
    36
Goal "Init (F Join G) = Init F Int Init G";
paulson@6012
    37
by (simp_tac (simpset() addsimps [Join_def]) 1);
paulson@5313
    38
qed "Init_Join";
paulson@5252
    39
paulson@6295
    40
Goal "Acts (F Join G) = Acts F Un Acts G";
paulson@5596
    41
by (auto_tac (claset(), simpset() addsimps [Join_def]));
paulson@5313
    42
qed "Acts_Join";
paulson@5252
    43
paulson@6295
    44
Addsimps [Init_Join];
paulson@6295
    45
paulson@6012
    46
paulson@6012
    47
(** JN **)
paulson@6012
    48
paulson@6295
    49
Goalw [JOIN_def, SKIP_def] "(JN i:{}. F i) = SKIP";
paulson@6012
    50
by Auto_tac;
paulson@6012
    51
qed "JN_empty";
paulson@6012
    52
Addsimps [JN_empty];
paulson@6012
    53
paulson@6295
    54
Goal "(JN x:insert a A. B x) = (B a) Join (JN x:A. B x)";
paulson@6295
    55
by (rtac program_equalityI 1);
paulson@6295
    56
by (ALLGOALS (simp_tac (simpset() addsimps [JOIN_def, Join_def])));
paulson@6295
    57
qed "JN_insert";
paulson@6295
    58
Addsimps[JN_empty, JN_insert];
paulson@6012
    59
paulson@5611
    60
Goal "Init (JN i:I. F i) = (INT i:I. Init (F i))";
paulson@6295
    61
by (simp_tac (simpset() addsimps [JOIN_def]) 1);
paulson@5313
    62
qed "Init_JN";
paulson@5252
    63
paulson@6295
    64
Goal "Acts (JN i:I. F i) = insert Id (UN i:I. Acts (F i))";
paulson@5596
    65
by (auto_tac (claset(), simpset() addsimps [JOIN_def]));
paulson@5313
    66
qed "Acts_JN";
paulson@5252
    67
paulson@6295
    68
Addsimps [Init_JN];
paulson@5259
    69
paulson@6295
    70
val prems = Goalw [JOIN_def]
paulson@6295
    71
    "[| A=B;  !!x. x:B ==> F(x) = G(x) |] ==> \
paulson@6295
    72
\    (JN x:A. F(x)) = (JN x:B. G(x))";
paulson@6295
    73
by (asm_simp_tac (simpset() addsimps prems) 1);
paulson@6295
    74
qed "JN_cong";
paulson@6012
    75
paulson@6295
    76
Addcongs [JN_cong];
paulson@5584
    77
paulson@5584
    78
paulson@5611
    79
(** Algebraic laws **)
paulson@5259
    80
paulson@5611
    81
Goal "F Join G = G Join F";
paulson@5313
    82
by (simp_tac (simpset() addsimps [Join_def, Un_commute, Int_commute]) 1);
paulson@5259
    83
qed "Join_commute";
paulson@5259
    84
paulson@5611
    85
Goal "(F Join G) Join H = F Join (G Join H)";
paulson@6295
    86
by (simp_tac (simpset() addsimps Un_ac@[Join_def, Int_assoc]) 1);
paulson@5259
    87
qed "Join_assoc";
paulson@5596
    88
 
paulson@6295
    89
Goalw [Join_def, SKIP_def] "SKIP Join F = F";
paulson@5611
    90
by (rtac program_equalityI 1);
paulson@6295
    91
by (ALLGOALS (simp_tac (simpset() addsimps [insert_absorb])));
paulson@5611
    92
qed "Join_SKIP_left";
paulson@5584
    93
paulson@6295
    94
Goalw [Join_def, SKIP_def] "F Join SKIP = F";
paulson@6295
    95
by (rtac program_equalityI 1);
paulson@6295
    96
by (ALLGOALS (simp_tac (simpset() addsimps [insert_absorb])));
paulson@5611
    97
qed "Join_SKIP_right";
paulson@5611
    98
paulson@5611
    99
Addsimps [Join_SKIP_left, Join_SKIP_right];
paulson@5611
   100
paulson@5611
   101
Goalw [Join_def] "F Join F = F";
paulson@5611
   102
by (rtac program_equalityI 1);
paulson@6295
   103
by Auto_tac;
paulson@5313
   104
qed "Join_absorb";
paulson@5313
   105
paulson@5611
   106
Addsimps [Join_absorb];
paulson@5313
   107
paulson@6295
   108
paulson@6295
   109
paulson@6295
   110
(*** JN laws ***)
paulson@6295
   111
paulson@6295
   112
paulson@6295
   113
paulson@6295
   114
Goal "i : I ==> (JN i:I. A i Join B) = ((JN i:I. A i) Join B)";
paulson@6295
   115
by (auto_tac (claset() addSIs [program_equalityI],
paulson@6295
   116
	      simpset() addsimps [Acts_JN, Acts_Join]));
paulson@6295
   117
qed "JN_Join_miniscope";
paulson@6295
   118
paulson@6295
   119
Goal "k:I ==> A k Join (JN i:I. A i) = (JN i:I. A i)";
paulson@6295
   120
by (auto_tac (claset() addSIs [program_equalityI],
paulson@6295
   121
	      simpset() addsimps [Acts_JN, Acts_Join]));
paulson@6295
   122
qed "JN_absorb";
paulson@6295
   123
paulson@5970
   124
Goal "(JN i: I Un J. A i) = ((JN i: I. A i) Join (JN i:J. A i))";
paulson@6295
   125
by (auto_tac (claset() addSIs [program_equalityI],
paulson@6295
   126
	      simpset() addsimps [Acts_JN, Acts_Join]));
paulson@5970
   127
qed "JN_Join";
paulson@5970
   128
paulson@5970
   129
Goal "i: I ==> (JN i:I. c) = c";
paulson@5970
   130
by (auto_tac (claset() addSIs [program_equalityI],
paulson@6295
   131
	      simpset() addsimps [Acts_JN]));
paulson@5970
   132
qed "JN_constant";
paulson@5970
   133
paulson@5970
   134
Goal "(JN i:I. A i Join B i) = (JN i:I. A i)  Join  (JN i:I. B i)";
paulson@6295
   135
by (auto_tac (claset() addSIs [program_equalityI],
paulson@6295
   136
	      simpset() addsimps [Acts_JN, Acts_Join]));
paulson@5970
   137
qed "JN_Join_distrib";
paulson@5584
   138
paulson@5584
   139
paulson@6536
   140
(*** Safety: co, stable, FP ***)
paulson@5313
   141
paulson@6295
   142
Goalw [constrains_def, JOIN_def]
paulson@6295
   143
    "i : I ==> \
paulson@6536
   144
\    (JN i:I. F i) : A co B = (ALL i:I. F i : A co B)";
paulson@6295
   145
by (Simp_tac 1);
paulson@6295
   146
by (Blast_tac 1);
paulson@6295
   147
qed "constrains_JN";
paulson@6012
   148
paulson@6536
   149
Goal "F Join G : A co B =  \
paulson@6536
   150
\     (F : A co B & G : A co B)";
paulson@6295
   151
by (auto_tac
paulson@6295
   152
    (claset(),
paulson@6295
   153
     simpset() addsimps [constrains_def, Join_def]));
paulson@6012
   154
qed "constrains_Join";
paulson@6012
   155
paulson@6536
   156
Goal "[| i : I;  ALL i:I. F i : A co B |] \
paulson@6536
   157
\     ==> (JN i:I. F i) : A co B";
paulson@6295
   158
by (full_simp_tac (simpset() addsimps [constrains_def, Acts_JN]) 1);
paulson@6295
   159
by (Blast_tac 1);
paulson@6012
   160
qed "constrains_imp_JN_constrains";
paulson@6012
   161
paulson@6012
   162
    (**FAILS, I think; see 5.4.1, Substitution Axiom.
paulson@6012
   163
       The problem is to relate reachable (F Join G) with 
paulson@6012
   164
       reachable F and reachable G
paulson@5620
   165
paulson@6012
   166
    Goalw [Constrains_def]
paulson@6536
   167
	"(JN i:I. F i) : A Co B = (ALL i:I. F i : A Co B)";
paulson@6012
   168
    by (simp_tac (simpset() addsimps [constrains_JN]) 1);
paulson@6012
   169
    by (Blast_tac 1);
paulson@6012
   170
    qed "Constrains_JN";
paulson@5620
   171
paulson@6536
   172
    Goal "F Join G : A Co B = (F Co A B & G Co A B)";
paulson@6012
   173
    by (auto_tac
paulson@6012
   174
	(claset(),
paulson@6012
   175
	 simpset() addsimps [Constrains_def, constrains_Join]));
paulson@6012
   176
    qed "Constrains_Join";
paulson@6012
   177
    **)
paulson@5313
   178
paulson@6295
   179
paulson@6536
   180
Goal "[| F : A co A';  G : B co B' |] \
paulson@6536
   181
\     ==> F Join G : (A Int B) co (A' Un B')";
paulson@6295
   182
by (simp_tac (simpset() addsimps [constrains_Join]) 1);
paulson@5313
   183
by (blast_tac (claset() addIs [constrains_weaken]) 1);
paulson@5313
   184
qed "constrains_Join_weaken";
paulson@5313
   185
paulson@6295
   186
Goal "i : I ==> \
paulson@6295
   187
\     (JN i:I. F i) : stable A = (ALL i:I. F i : stable A)";
paulson@5584
   188
by (asm_simp_tac (simpset() addsimps [stable_def, constrains_JN]) 1);
paulson@5313
   189
qed "stable_JN";
paulson@5313
   190
paulson@6295
   191
Goal "[| ALL i:I. F i : invariant A;  i : I |]  \
paulson@5970
   192
\      ==> (JN i:I. F i) : invariant A";
paulson@5970
   193
by (asm_full_simp_tac (simpset() addsimps [invariant_def, stable_JN]) 1);
paulson@5970
   194
by (Blast_tac 1);
paulson@5970
   195
bind_thm ("invariant_JN_I", ballI RS result());
paulson@5970
   196
paulson@6295
   197
Goal "F Join G : stable A = \
paulson@6295
   198
\     (F : stable A & G : stable A)";
paulson@6295
   199
by (simp_tac (simpset() addsimps [stable_def, constrains_Join]) 1);
paulson@5313
   200
qed "stable_Join";
paulson@5313
   201
paulson@6295
   202
Goal "[| F : invariant A; G : invariant A |]  \
paulson@5970
   203
\     ==> F Join G : invariant A";
paulson@6295
   204
by (full_simp_tac (simpset() addsimps [invariant_def, stable_Join]) 1);
paulson@5970
   205
by (Blast_tac 1);
paulson@5970
   206
qed "invariant_JoinI";
paulson@5970
   207
paulson@6295
   208
Goal "i : I ==> FP (JN i:I. F i) = (INT i:I. FP (F i))";
paulson@5584
   209
by (asm_simp_tac (simpset() addsimps [FP_def, stable_JN, INTER_def]) 1);
paulson@5313
   210
qed "FP_JN";
paulson@5313
   211
paulson@5313
   212
paulson@5313
   213
(*** Progress: transient, ensures ***)
paulson@5313
   214
paulson@6295
   215
Goal "i : I ==> \
paulson@6295
   216
\   (JN i:I. F i) : transient A = (EX i:I. F i : transient A)";
paulson@6295
   217
by (auto_tac (claset(),
paulson@6295
   218
	      simpset() addsimps [transient_def, JOIN_def]));
paulson@5313
   219
qed "transient_JN";
paulson@5313
   220
paulson@6295
   221
Goal "F Join G : transient A = \
paulson@6295
   222
\     (F : transient A | G : transient A)";
paulson@6295
   223
by (auto_tac (claset(),
paulson@6295
   224
	      simpset() addsimps [bex_Un, transient_def,
paulson@6295
   225
				  Join_def]));
paulson@5313
   226
qed "transient_Join";
paulson@5313
   227
paulson@6295
   228
Goal "i : I ==> \
paulson@6536
   229
\     (JN i:I. F i) : A ensures B = \
paulson@6536
   230
\     ((ALL i:I. F i : (A-B) co (A Un B)) & \
paulson@6536
   231
\      (EX i:I. F i : A ensures B))";
paulson@5313
   232
by (auto_tac (claset(),
paulson@5584
   233
	      simpset() addsimps [ensures_def, constrains_JN, transient_JN]));
paulson@5313
   234
qed "ensures_JN";
paulson@5313
   235
paulson@5313
   236
Goalw [ensures_def]
paulson@6536
   237
     "F Join G : A ensures B =     \
paulson@6536
   238
\     (F : (A-B) co (A Un B) & \
paulson@6536
   239
\      G : (A-B) co (A Un B) & \
paulson@6536
   240
\      (F : A ensures B | G : A ensures B))";
paulson@5313
   241
by (auto_tac (claset(),
paulson@5313
   242
	      simpset() addsimps [constrains_Join, transient_Join]));
paulson@5313
   243
qed "ensures_Join";
paulson@5313
   244
paulson@6295
   245
Goalw [stable_def, constrains_def, Join_def]
paulson@6536
   246
    "[| F : stable A;  G : A co A' |] \
paulson@6536
   247
\    ==> F Join G : A co A'";
paulson@6295
   248
by (asm_full_simp_tac (simpset() addsimps [ball_Un]) 1);
paulson@5313
   249
by (Blast_tac 1);
paulson@5313
   250
qed "stable_constrains_Join";
paulson@5313
   251
paulson@5648
   252
(*Premise for G cannot use Invariant because  Stable F A  is weaker than
paulson@5648
   253
  G : stable A *)
paulson@6295
   254
Goal "[| F : stable A;  G : invariant A |] ==> F Join G : Invariant A";
paulson@6295
   255
by (full_simp_tac (simpset() addsimps [Invariant_def, invariant_def, 
paulson@6295
   256
				       Stable_eq_stable, stable_Join]) 1);
paulson@5313
   257
by (force_tac(claset() addIs [stable_reachable, stable_Int],
paulson@5313
   258
	      simpset() addsimps [Acts_Join]) 1);
paulson@5313
   259
qed "stable_Join_Invariant";
paulson@5313
   260
paulson@6536
   261
Goal "[| F : stable A;  G : A ensures B |] ==> F Join G : A ensures B";
paulson@5313
   262
by (asm_simp_tac (simpset() addsimps [ensures_Join]) 1);
paulson@5313
   263
by (asm_full_simp_tac (simpset() addsimps [stable_def, ensures_def]) 1);
paulson@5313
   264
by (etac constrains_weaken 1);
paulson@5313
   265
by Auto_tac;
paulson@5313
   266
qed "ensures_stable_Join1";
paulson@5313
   267
paulson@5313
   268
(*As above, but exchanging the roles of F and G*)
paulson@6536
   269
Goal "[| F : A ensures B;  G : stable A |] ==> F Join G : A ensures B";
paulson@5313
   270
by (stac Join_commute 1);
paulson@5313
   271
by (blast_tac (claset() addIs [ensures_stable_Join1]) 1);
paulson@5313
   272
qed "ensures_stable_Join2";
paulson@5313
   273
paulson@5648
   274
paulson@5804
   275
(** Diff and localTo **)
paulson@5648
   276
paulson@6295
   277
Goalw [Join_def, Diff_def] "F Join (Diff G (Acts F)) = F Join G";
paulson@5804
   278
by (rtac program_equalityI 1);
paulson@6295
   279
by Auto_tac;
paulson@5804
   280
qed "Join_Diff2";
paulson@5804
   281
paulson@6295
   282
Goalw [Diff_def, Disjoint_def] "Disjoint F (Diff G (Acts F))";
paulson@5804
   283
by Auto_tac;
paulson@5804
   284
qed "Diff_Disjoint";
paulson@5804
   285
paulson@5804
   286
Goal "[| F Join G : v localTo F;  Disjoint F G |] \
paulson@5804
   287
\     ==> G : (INT z. stable {s. v s = z})";
paulson@5648
   288
by (asm_full_simp_tac 
paulson@6295
   289
    (simpset() addsimps [localTo_def, Diff_def, Disjoint_def,
paulson@5804
   290
			 Acts_Join, stable_def, constrains_def]) 1);
paulson@5804
   291
by (Blast_tac 1);
paulson@5804
   292
qed "localTo_imp_stable";
paulson@5804
   293
paulson@5804
   294
Goal "[| F Join G : v localTo F;  (s,s') : act;  \
paulson@6295
   295
\        act : Acts G;  Disjoint F G |] ==> v s' = v s";
paulson@5804
   296
by (asm_full_simp_tac 
paulson@6295
   297
    (simpset() addsimps [localTo_def, Diff_def, Disjoint_def,
paulson@5804
   298
			 Acts_Join, stable_def, constrains_def]) 1);
paulson@5648
   299
by (Blast_tac 1);
paulson@5648
   300
qed "localTo_imp_equals";
paulson@5804
   301
paulson@5804
   302
Goalw [localTo_def, stable_def, constrains_def]
paulson@5804
   303
     "v localTo F <= (f o v) localTo F";
paulson@5804
   304
by (Clarify_tac 1);
paulson@5898
   305
by (force_tac (claset() addSEs [allE, ballE], simpset()) 1);
paulson@5804
   306
qed "localTo_imp_o_localTo";
paulson@5804
   307
paulson@5804
   308
paulson@5804
   309
(*** Higher-level rules involving localTo and Join ***)
paulson@5804
   310
paulson@6536
   311
Goal "[| F : {s. P (v s) (w s)} co {s. P' (v s) (w s)};   \
paulson@5804
   312
\        F Join G : v localTo F;       \
paulson@5804
   313
\        F Join G : w localTo F;       \
paulson@5804
   314
\        Disjoint F G |]               \
paulson@6536
   315
\     ==> F Join G : {s. P (v s) (w s)} co {s. P' (v s) (w s)}";
paulson@6295
   316
by (auto_tac (claset(), simpset() addsimps [constrains_def, Acts_Join]));
paulson@5804
   317
by (REPEAT_FIRST (dtac localTo_imp_equals THEN' REPEAT1 o atac));
paulson@6295
   318
by Auto_tac;
paulson@5804
   319
qed "constrains_localTo_constrains2";
paulson@5804
   320
paulson@5804
   321
Goalw [stable_def]
paulson@5804
   322
     "[| F : stable {s. P (v s) (w s)};   \
paulson@5804
   323
\        F Join G : v localTo F;       \
paulson@5804
   324
\        F Join G : w localTo F;       \
paulson@5804
   325
\        Disjoint F G |]               \
paulson@5804
   326
\     ==> F Join G : stable {s. P (v s) (w s)}";
paulson@5804
   327
by (blast_tac (claset() addIs [constrains_localTo_constrains2]) 1);
paulson@5804
   328
qed "stable_localTo_stable2";
paulson@5804
   329
paulson@5804
   330
paulson@5804
   331
Goal "(UN k. {s. f s = k}) = UNIV";
paulson@5804
   332
by (Blast_tac 1);
paulson@5804
   333
qed "UN_eq_UNIV";
paulson@5804
   334
paulson@5804
   335
Goal "[| F : stable {s. v s <= w s};   \
paulson@5804
   336
\        F Join G : v localTo F;       \
paulson@5804
   337
\        F Join G : Increasing w;      \
paulson@5804
   338
\        Disjoint F G |]               \
paulson@5804
   339
\     ==> F Join G : Stable {s. v s <= w s}";
paulson@5804
   340
by (safe_tac (claset() addSDs [localTo_imp_stable]));
paulson@5804
   341
by (rewrite_goals_tac [stable_def, Stable_def, Increasing_def]);
paulson@6536
   342
by (subgoal_tac "ALL k: UNIV. ?H : ({s. v s = k} Int ?AA) Co ?AA" 1);
paulson@5804
   343
by (dtac ball_Constrains_UN 1);
paulson@5804
   344
by (full_simp_tac (simpset() addsimps [UN_eq_UNIV]) 1);
paulson@5804
   345
by (rtac ballI 1);
paulson@6536
   346
by (subgoal_tac "F Join G : ({s. v s = k} Int {s. v s <= w s}) co \
paulson@5804
   347
\                                      ({s. v s = k} Un {s. v s <= w s})" 1);
paulson@6295
   348
by (asm_simp_tac (simpset() addsimps [constrains_Join]) 2);
paulson@5804
   349
by (blast_tac (claset() addIs [constrains_weaken]) 2);
paulson@5804
   350
by (dtac (constrains_imp_Constrains RS Constrains_Int) 1 THEN etac INT_D 1);
paulson@5804
   351
by (etac Constrains_weaken 2);
paulson@5804
   352
by Auto_tac;
paulson@5804
   353
qed "Increasing_localTo_Stable";