src/HOL/UNITY/SubstAx.ML
author paulson
Thu, 02 Jul 1998 16:53:55 +0200
changeset 5111 8f4b72f0c15d
parent 5069 3ea049f7979d
child 5232 e5a7cdd07ea5
permissions -rw-r--r--
Uncurried functions LeadsTo and reach Deleted leading parameters thanks to new Goal command
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     1
(*  Title:      HOL/UNITY/SubstAx
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     2
    ID:         $Id$
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     4
    Copyright   1998  University of Cambridge
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     5
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     6
Weak Fairness versions of transient, ensures, LeadsTo.
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     7
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     8
From Misra, "A Logic for Concurrent Programming", 1994
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
     9
*)
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    10
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    11
open SubstAx;
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    12
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    13
(*constrains Acts B B' ==> constrains Acts (reachable(Init,Acts) Int B)
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    14
                                           (reachable(Init,Acts) Int B') *)
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    15
bind_thm ("constrains_reachable",
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    16
	  rewrite_rule [stable_def] stable_reachable RS constrains_Int);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    17
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    18
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    19
(*** Introduction rules: Basis, Trans, Union ***)
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    20
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    21
Goal "leadsTo Acts A B ==> LeadsTo(Init,Acts) A B";
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    22
by (simp_tac (simpset() addsimps [LeadsTo_def]) 1);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    23
by (blast_tac (claset() addIs [PSP_stable2, stable_reachable]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    24
qed "leadsTo_imp_LeadsTo";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    25
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    26
Goal "[| constrains Acts (reachable(Init,Acts) Int (A - A'))   \
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    27
\                               (A Un A'); \
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    28
\               transient  Acts (reachable(Init,Acts) Int (A-A')) |]   \
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    29
\           ==> LeadsTo(Init,Acts) A A'";
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    30
by (simp_tac (simpset() addsimps [LeadsTo_def]) 1);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    31
by (rtac (stable_reachable RS stable_ensures_Int RS leadsTo_Basis) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    32
by (assume_tac 2);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    33
by (asm_simp_tac 
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    34
    (simpset() addsimps [Int_Un_distrib RS sym, Diff_Int_distrib RS sym,
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    35
			 stable_constrains_Int]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    36
qed "LeadsTo_Basis";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    37
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    38
Goal "[| LeadsTo(Init,Acts) A B;  LeadsTo(Init,Acts) B C |] \
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    39
\            ==> LeadsTo(Init,Acts) A C";
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    40
by (full_simp_tac (simpset() addsimps [LeadsTo_def]) 1);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    41
by (blast_tac (claset() addIs [leadsTo_Trans]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    42
qed "LeadsTo_Trans";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    43
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    44
val [prem] = goalw thy [LeadsTo_def]
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    45
 "(!!A. A : S ==> LeadsTo(Init,Acts) A B) ==> LeadsTo(Init,Acts) (Union S) B";
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    46
by (Simp_tac 1);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    47
by (stac Int_Union 1);
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    48
by (blast_tac (claset() addIs [leadsTo_UN,
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    49
			        simplify (simpset()) prem]) 1);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    50
qed "LeadsTo_Union";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    51
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    52
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    53
(*** Derived rules ***)
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    54
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    55
Goal "id: Acts ==> LeadsTo(Init,Acts) A UNIV";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    56
by (asm_simp_tac (simpset() addsimps [LeadsTo_def, 
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    57
				      Int_lower1 RS subset_imp_leadsTo]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    58
qed "LeadsTo_UNIV";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    59
Addsimps [LeadsTo_UNIV];
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    60
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    61
(*Useful with cancellation, disjunction*)
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    62
Goal "LeadsTo(Init,Acts) A (A' Un A') ==> LeadsTo(Init,Acts) A A'";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    63
by (asm_full_simp_tac (simpset() addsimps Un_ac) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    64
qed "LeadsTo_Un_duplicate";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    65
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    66
Goal "LeadsTo(Init,Acts) A (A' Un C Un C) ==> LeadsTo(Init,Acts) A (A' Un C)";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    67
by (asm_full_simp_tac (simpset() addsimps Un_ac) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    68
qed "LeadsTo_Un_duplicate2";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    69
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    70
val prems = goal thy
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    71
   "(!!i. i : I ==> LeadsTo(Init,Acts) (A i) B) \
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    72
\   ==> LeadsTo(Init,Acts) (UN i:I. A i) B";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    73
by (simp_tac (simpset() addsimps [Union_image_eq RS sym]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    74
by (blast_tac (claset() addIs (LeadsTo_Union::prems)) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    75
qed "LeadsTo_UN";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    76
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    77
(*Binary union introduction rule*)
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
    78
Goal
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    79
  "[| LeadsTo(Init,Acts) A C; LeadsTo(Init,Acts) B C |] ==> LeadsTo(Init,Acts) (A Un B) C";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    80
by (stac Un_eq_Union 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    81
by (blast_tac (claset() addIs [LeadsTo_Union]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    82
qed "LeadsTo_Un";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    83
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    84
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    85
Goal "[| reachable(Init,Acts) Int A <= B;  id: Acts |] \
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    86
\            ==> LeadsTo(Init,Acts) A B";
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    87
by (simp_tac (simpset() addsimps [LeadsTo_def]) 1);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    88
by (blast_tac (claset() addIs [subset_imp_leadsTo]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    89
qed "Int_subset_imp_LeadsTo";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    90
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    91
Goal "[| A <= B;  id: Acts |] ==> LeadsTo(Init,Acts) A B";
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    92
by (simp_tac (simpset() addsimps [LeadsTo_def]) 1);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    93
by (blast_tac (claset() addIs [subset_imp_leadsTo]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    94
qed "subset_imp_LeadsTo";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    95
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    96
bind_thm ("empty_LeadsTo", empty_subsetI RS subset_imp_LeadsTo);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    97
Addsimps [empty_LeadsTo];
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
    98
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
    99
Goal "[| reachable(Init,Acts) Int A = {};  id: Acts |] \
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   100
\            ==> LeadsTo(Init,Acts) A B";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   101
by (asm_simp_tac (simpset() addsimps [Int_subset_imp_LeadsTo]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   102
qed "Int_empty_LeadsTo";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   103
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   104
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   105
Goal "[| LeadsTo(Init,Acts) A A';   \
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   106
\                reachable(Init,Acts) Int A' <= B' |] \
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   107
\             ==> LeadsTo(Init,Acts) A B'";
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   108
by (full_simp_tac (simpset() addsimps [LeadsTo_def]) 1);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   109
by (blast_tac (claset() addIs [leadsTo_weaken_R]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   110
qed_spec_mp "LeadsTo_weaken_R";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   111
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   112
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   113
Goal "[| LeadsTo(Init,Acts) A A'; \
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   114
\                reachable(Init,Acts) Int B <= A; id: Acts |]  \
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   115
\             ==> LeadsTo(Init,Acts) B A'";
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   116
by (full_simp_tac (simpset() addsimps [LeadsTo_def]) 1);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   117
by (blast_tac (claset() addIs [leadsTo_weaken_L]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   118
qed_spec_mp "LeadsTo_weaken_L";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   119
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   120
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   121
(*Distributes over binary unions*)
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
   122
Goal
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   123
  "id: Acts ==> \
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   124
\       LeadsTo(Init,Acts) (A Un B) C  =  \
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   125
\       (LeadsTo(Init,Acts) A C & LeadsTo(Init,Acts) B C)";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   126
by (blast_tac (claset() addIs [LeadsTo_Un, LeadsTo_weaken_L]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   127
qed "LeadsTo_Un_distrib";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   128
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
   129
Goal
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   130
  "id: Acts ==> \
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   131
\       LeadsTo(Init,Acts) (UN i:I. A i) B  =  \
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   132
\       (ALL i : I. LeadsTo(Init,Acts) (A i) B)";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   133
by (blast_tac (claset() addIs [LeadsTo_UN, LeadsTo_weaken_L]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   134
qed "LeadsTo_UN_distrib";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   135
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
   136
Goal
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   137
  "id: Acts ==> \
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   138
\       LeadsTo(Init,Acts) (Union S) B  =  \
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   139
\       (ALL A : S. LeadsTo(Init,Acts) A B)";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   140
by (blast_tac (claset() addIs [LeadsTo_Union, LeadsTo_weaken_L]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   141
qed "LeadsTo_Union_distrib";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   142
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   143
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
   144
Goal 
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   145
   "[| LeadsTo(Init,Acts) A A'; id: Acts;   \
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   146
\               reachable(Init,Acts) Int B  <= A;     \
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   147
\               reachable(Init,Acts) Int A' <= B' |] \
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   148
\           ==> LeadsTo(Init,Acts) B B'";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   149
(*PROOF FAILED: why?*)
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   150
by (blast_tac (claset() addIs [LeadsTo_Trans, LeadsTo_weaken_R,
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   151
			       LeadsTo_weaken_L]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   152
qed "LeadsTo_weaken";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   153
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   154
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   155
(*Set difference: maybe combine with leadsTo_weaken_L??*)
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
   156
Goal
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   157
  "[| LeadsTo(Init,Acts) (A-B) C; LeadsTo(Init,Acts) B C; id: Acts |] \
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   158
\       ==> LeadsTo(Init,Acts) A C";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   159
by (blast_tac (claset() addIs [LeadsTo_Un, LeadsTo_weaken]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   160
qed "LeadsTo_Diff";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   161
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   162
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   163
(** Meta or object quantifier ???????????????????
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   164
    see ball_constrains_UN in UNITY.ML***)
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   165
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   166
val prems = goal thy
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   167
   "(!! i. i:I ==> LeadsTo(Init,Acts) (A i) (A' i)) \
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   168
\   ==> LeadsTo(Init,Acts) (UN i:I. A i) (UN i:I. A' i)";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   169
by (simp_tac (simpset() addsimps [Union_image_eq RS sym]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   170
by (blast_tac (claset() addIs [LeadsTo_Union, LeadsTo_weaken_R] 
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   171
                        addIs prems) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   172
qed "LeadsTo_UN_UN";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   173
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   174
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   175
(*Version with no index set*)
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   176
val prems = goal thy
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   177
   "(!! i. LeadsTo(Init,Acts) (A i) (A' i)) \
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   178
\   ==> LeadsTo(Init,Acts) (UN i. A i) (UN i. A' i)";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   179
by (blast_tac (claset() addIs [LeadsTo_UN_UN] 
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   180
                        addIs prems) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   181
qed "LeadsTo_UN_UN_noindex";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   182
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   183
(*Version with no index set*)
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
   184
Goal
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   185
   "ALL i. LeadsTo(Init,Acts) (A i) (A' i) \
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   186
\           ==> LeadsTo(Init,Acts) (UN i. A i) (UN i. A' i)";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   187
by (blast_tac (claset() addIs [LeadsTo_UN_UN]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   188
qed "all_LeadsTo_UN_UN";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   189
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   190
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   191
(*Binary union version*)
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   192
Goal "[| LeadsTo(Init,Acts) A A'; LeadsTo(Init,Acts) B B' |] \
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   193
\                 ==> LeadsTo(Init,Acts) (A Un B) (A' Un B')";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   194
by (blast_tac (claset() addIs [LeadsTo_Un, 
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   195
			       LeadsTo_weaken_R]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   196
qed "LeadsTo_Un_Un";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   197
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   198
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   199
(** The cancellation law **)
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   200
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
   201
Goal
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   202
   "[| LeadsTo(Init,Acts) A (A' Un B); LeadsTo(Init,Acts) B B'; \
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   203
\              id: Acts |]    \
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   204
\           ==> LeadsTo(Init,Acts) A (A' Un B')";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   205
by (blast_tac (claset() addIs [LeadsTo_Un_Un, 
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   206
			       subset_imp_LeadsTo, LeadsTo_Trans]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   207
qed "LeadsTo_cancel2";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   208
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
   209
Goal
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   210
   "[| LeadsTo(Init,Acts) A (A' Un B); LeadsTo(Init,Acts) (B-A') B'; id: Acts |] \
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   211
\           ==> LeadsTo(Init,Acts) A (A' Un B')";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   212
by (rtac LeadsTo_cancel2 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   213
by (assume_tac 2);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   214
by (ALLGOALS Asm_simp_tac);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   215
qed "LeadsTo_cancel_Diff2";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   216
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
   217
Goal
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   218
   "[| LeadsTo(Init,Acts) A (B Un A'); LeadsTo(Init,Acts) B B'; id: Acts |] \
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   219
\           ==> LeadsTo(Init,Acts) A (B' Un A')";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   220
by (asm_full_simp_tac (simpset() addsimps [Un_commute]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   221
by (blast_tac (claset() addSIs [LeadsTo_cancel2]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   222
qed "LeadsTo_cancel1";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   223
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
   224
Goal
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   225
   "[| LeadsTo(Init,Acts) A (B Un A'); LeadsTo(Init,Acts) (B-A') B'; id: Acts |] \
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   226
\           ==> LeadsTo(Init,Acts) A (B' Un A')";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   227
by (rtac LeadsTo_cancel1 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   228
by (assume_tac 2);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   229
by (ALLGOALS Asm_simp_tac);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   230
qed "LeadsTo_cancel_Diff1";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   231
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   232
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   233
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   234
(** The impossibility law **)
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   235
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   236
Goal "LeadsTo(Init,Acts) A {} ==> reachable(Init,Acts) Int A  = {}";
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   237
by (full_simp_tac (simpset() addsimps [LeadsTo_def]) 1);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   238
by (etac leadsTo_empty 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   239
qed "LeadsTo_empty";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   240
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   241
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   242
(** PSP: Progress-Safety-Progress **)
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   243
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   244
(*Special case of PSP: Misra's "stable conjunction".  Doesn't need id:Acts. *)
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   245
Goal "[| LeadsTo(Init,Acts) A A'; stable Acts B |] \
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   246
\           ==> LeadsTo(Init,Acts) (A Int B) (A' Int B)";
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   247
by (asm_full_simp_tac (simpset() addsimps [LeadsTo_def, Int_assoc RS sym, 
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   248
					   PSP_stable]) 1);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   249
qed "R_PSP_stable";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   250
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   251
Goal "[| LeadsTo(Init,Acts) A A'; stable Acts B |] \
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   252
\             ==> LeadsTo(Init,Acts) (B Int A) (B Int A')";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   253
by (asm_simp_tac (simpset() addsimps (R_PSP_stable::Int_ac)) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   254
qed "R_PSP_stable2";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   255
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   256
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   257
Goal "[| LeadsTo(Init,Acts) A A'; constrains Acts B B'; id: Acts |] \
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   258
\           ==> LeadsTo(Init,Acts) (A Int B) ((A' Int B) Un (B' - B))";
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   259
by (full_simp_tac (simpset() addsimps [LeadsTo_def]) 1);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   260
by (dtac PSP 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   261
by (etac constrains_reachable 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   262
by (etac leadsTo_weaken 2);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   263
by (ALLGOALS Blast_tac);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   264
qed "R_PSP";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   265
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
   266
Goal
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   267
   "[| LeadsTo(Init,Acts) A A'; constrains Acts B B'; id: Acts |] \
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   268
\           ==> LeadsTo(Init,Acts) (B Int A) ((B Int A') Un (B' - B))";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   269
by (asm_simp_tac (simpset() addsimps (R_PSP::Int_ac)) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   270
qed "R_PSP2";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   271
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
   272
Goalw [unless_def]
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   273
   "[| LeadsTo(Init,Acts) A A'; unless Acts B B'; id: Acts |] \
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   274
\           ==> LeadsTo(Init,Acts) (A Int B) ((A' Int B) Un B')";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   275
by (dtac R_PSP 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   276
by (assume_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   277
by (asm_full_simp_tac (simpset() addsimps [Un_Diff_Diff, Int_Diff_Un]) 2);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   278
by (asm_full_simp_tac (simpset() addsimps [Diff_Int_distrib]) 2);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   279
by (etac LeadsTo_Diff 2);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   280
by (blast_tac (claset() addIs [subset_imp_LeadsTo]) 2);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   281
by Auto_tac;
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   282
qed "R_PSP_unless";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   283
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   284
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   285
(*** Induction rules ***)
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   286
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   287
(** Meta or object quantifier ????? **)
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   288
Goal
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   289
   "[| wf r;     \
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   290
\              ALL m. LeadsTo(Init,Acts) (A Int f-``{m})                     \
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   291
\                                       ((A Int f-``(r^-1 ^^ {m})) Un B);   \
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   292
\              id: Acts |] \
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   293
\           ==> LeadsTo(Init,Acts) A B";
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   294
by (full_simp_tac (simpset() addsimps [LeadsTo_def]) 1);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   295
by (etac leadsTo_wf_induct 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   296
by (assume_tac 2);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   297
by (blast_tac (claset() addIs [leadsTo_weaken]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   298
qed "LeadsTo_wf_induct";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   299
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   300
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
   301
Goal
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   302
   "[| wf r;     \
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   303
\              ALL m:I. LeadsTo(Init,Acts) (A Int f-``{m})                   \
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   304
\                                  ((A Int f-``(r^-1 ^^ {m})) Un B);   \
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   305
\              id: Acts |] \
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   306
\           ==> LeadsTo(Init,Acts) A ((A - (f-``I)) Un B)";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   307
by (etac LeadsTo_wf_induct 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   308
by Safe_tac;
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   309
by (case_tac "m:I" 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   310
by (blast_tac (claset() addIs [LeadsTo_weaken]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   311
by (blast_tac (claset() addIs [subset_imp_LeadsTo]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   312
qed "R_bounded_induct";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   313
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   314
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
   315
Goal
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   316
   "[| ALL m. LeadsTo(Init,Acts) (A Int f-``{m})                     \
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   317
\                                  ((A Int f-``(lessThan m)) Un B);   \
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   318
\              id: Acts |] \
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   319
\           ==> LeadsTo(Init,Acts) A B";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   320
by (rtac (wf_less_than RS LeadsTo_wf_induct) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   321
by (assume_tac 2);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   322
by (Asm_simp_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   323
qed "R_lessThan_induct";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   324
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
   325
Goal
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   326
   "[| ALL m:(greaterThan l). LeadsTo(Init,Acts) (A Int f-``{m})   \
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   327
\                                        ((A Int f-``(lessThan m)) Un B);   \
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   328
\              id: Acts |] \
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   329
\           ==> LeadsTo(Init,Acts) A ((A Int (f-``(atMost l))) Un B)";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   330
by (simp_tac (HOL_ss addsimps [Diff_eq RS sym, vimage_Compl, Compl_greaterThan RS sym]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   331
by (rtac (wf_less_than RS R_bounded_induct) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   332
by (assume_tac 2);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   333
by (Asm_simp_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   334
qed "R_lessThan_bounded_induct";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   335
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
   336
Goal
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   337
   "[| ALL m:(lessThan l). LeadsTo(Init,Acts) (A Int f-``{m})   \
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   338
\                                    ((A Int f-``(greaterThan m)) Un B);   \
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   339
\              id: Acts |] \
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   340
\           ==> LeadsTo(Init,Acts) A ((A Int (f-``(atLeast l))) Un B)";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   341
by (res_inst_tac [("f","f"),("f1", "%k. l - k")]
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   342
    (wf_less_than RS wf_inv_image RS LeadsTo_wf_induct) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   343
by (assume_tac 2);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   344
by (simp_tac (simpset() addsimps [inv_image_def, Image_singleton]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   345
by (Clarify_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   346
by (case_tac "m<l" 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   347
by (blast_tac (claset() addIs [not_leE, subset_imp_LeadsTo]) 2);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   348
by (blast_tac (claset() addIs [LeadsTo_weaken_R, diff_less_mono2]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   349
qed "R_greaterThan_bounded_induct";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   350
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   351
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   352
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   353
(*** Completion: Binary and General Finite versions ***)
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   354
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   355
Goal
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   356
   "[| LeadsTo(Init,Acts) A A';  stable Acts A';   \
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   357
\              LeadsTo(Init,Acts) B B';  stable Acts B';  id: Acts |] \
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   358
\           ==> LeadsTo(Init,Acts) (A Int B) (A' Int B')";
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   359
by (full_simp_tac (simpset() addsimps [LeadsTo_def]) 1);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   360
by (blast_tac (claset() addIs [stable_completion RS leadsTo_weaken] 
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   361
                        addSIs [stable_Int, stable_reachable]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   362
qed "R_stable_completion";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   363
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   364
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   365
Goal "[| finite I;  id: Acts |]                     \
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   366
\           ==> (ALL i:I. LeadsTo(Init,Acts) (A i) (A' i)) -->  \
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   367
\               (ALL i:I. stable Acts (A' i)) -->         \
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   368
\               LeadsTo(Init,Acts) (INT i:I. A i) (INT i:I. A' i)";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   369
by (etac finite_induct 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   370
by (Asm_simp_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   371
by (asm_simp_tac 
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   372
    (simpset() addsimps [R_stable_completion, stable_def, 
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   373
			 ball_constrains_INT]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   374
qed_spec_mp "R_finite_stable_completion";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   375
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   376
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   377
Goal
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   378
 "[| LeadsTo(Init,Acts) A (A' Un C);  constrains Acts A' (A' Un C); \
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   379
\            LeadsTo(Init,Acts) B (B' Un C);  constrains Acts B' (B' Un C); \
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   380
\            id: Acts |] \
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   381
\         ==> LeadsTo(Init,Acts) (A Int B) ((A' Int B') Un C)";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   382
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   383
by (full_simp_tac (simpset() addsimps [LeadsTo_def, Int_Un_distrib]) 1);
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   384
by (dtac completion 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   385
by (assume_tac 2);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   386
by (ALLGOALS 
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   387
    (asm_simp_tac 
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   388
     (simpset() addsimps [constrains_reachable, Int_Un_distrib RS sym])));
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   389
by (blast_tac (claset() addIs [leadsTo_weaken]) 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   390
qed "R_completion";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   391
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   392
5069
3ea049f7979d isatool fixgoal;
wenzelm
parents: 4776
diff changeset
   393
Goal
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   394
   "[| finite I;  id: Acts |] \
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   395
\           ==> (ALL i:I. LeadsTo(Init,Acts) (A i) (A' i Un C)) -->  \
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   396
\               (ALL i:I. constrains Acts (A' i) (A' i Un C)) --> \
5111
8f4b72f0c15d Uncurried functions LeadsTo and reach
paulson
parents: 5069
diff changeset
   397
\               LeadsTo(Init,Acts) (INT i:I. A i) ((INT i:I. A' i) Un C)";
4776
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   398
by (etac finite_induct 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   399
by (ALLGOALS Asm_simp_tac);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   400
by (Clarify_tac 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   401
by (dtac ball_constrains_INT 1);
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   402
by (asm_full_simp_tac (simpset() addsimps [R_completion]) 1); 
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   403
qed "R_finite_completion";
1f9362e769c1 New UNITY theory
paulson
parents:
diff changeset
   404