src/HOL/UNITY/Constrains.ML
author paulson
Fri, 06 Nov 1998 13:20:29 +0100
changeset 5804 8e0a4c4fd67b
parent 5784 54276fba8420
child 6012 1894bfc4aee9
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.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
     1
(*  Title:      HOL/UNITY/Constrains
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
     2
    ID:         $Id$
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
     4
    Copyright   1998  University of Cambridge
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
     5
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
     6
Safety relations: restricted to the set of reachable states.
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
     7
*)
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
     8
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
     9
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    10
(*** Constrains ***)
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    11
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
    12
overload_1st_set "Constrains.Constrains";
5340
d75c03cf77b5 Misc changes
paulson
parents: 5319
diff changeset
    13
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
    14
(*F : constrains B B'
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
    15
  ==> F : constrains (reachable F Int B) (reachable F Int B')*)
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    16
bind_thm ("constrains_reachable_Int",
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    17
	  subset_refl RS
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    18
	  rewrite_rule [stable_def] stable_reachable RS 
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    19
	  constrains_Int);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    20
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
    21
Goalw [Constrains_def] "F : constrains A A' ==> F : Constrains A A'";
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
    22
by (blast_tac (claset() addIs [constrains_reachable_Int]) 1);
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    23
qed "constrains_imp_Constrains";
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    24
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
    25
Goalw [stable_def, Stable_def] "F : stable A ==> F : Stable A";
5631
707f30bc7fe7 new theorems
paulson
parents: 5620
diff changeset
    26
by (etac constrains_imp_Constrains 1);
707f30bc7fe7 new theorems
paulson
parents: 5620
diff changeset
    27
qed "stable_imp_Stable";
707f30bc7fe7 new theorems
paulson
parents: 5620
diff changeset
    28
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    29
val prems = Goal
5620
3ac11c4af76a tidying and renaming
paulson
parents: 5584
diff changeset
    30
    "(!!act s s'. [| act: Acts F;  (s,s') : act;  s: A |] ==> s': A') \
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
    31
\    ==> F : Constrains A A'";
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    32
by (rtac constrains_imp_Constrains 1);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    33
by (blast_tac (claset() addIs (constrainsI::prems)) 1);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    34
qed "ConstrainsI";
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    35
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
    36
Goalw [Constrains_def, constrains_def] "F : Constrains {} B";
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    37
by (Blast_tac 1);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    38
qed "Constrains_empty";
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    39
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
    40
Goal "F : Constrains A UNIV";
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    41
by (blast_tac (claset() addIs [ConstrainsI]) 1);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    42
qed "Constrains_UNIV";
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    43
AddIffs [Constrains_empty, Constrains_UNIV];
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    44
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    45
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    46
Goalw [Constrains_def]
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
    47
    "[| F : Constrains A A'; A'<=B' |] ==> F : Constrains A B'";
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    48
by (blast_tac (claset() addIs [constrains_weaken_R]) 1);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    49
qed "Constrains_weaken_R";
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    50
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    51
Goalw [Constrains_def]
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
    52
    "[| F : Constrains A A'; B<=A |] ==> F : Constrains B A'";
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    53
by (blast_tac (claset() addIs [constrains_weaken_L]) 1);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    54
qed "Constrains_weaken_L";
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    55
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    56
Goalw [Constrains_def]
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
    57
   "[| F : Constrains A A'; B<=A; A'<=B' |] ==> F : Constrains B B'";
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    58
by (blast_tac (claset() addIs [constrains_weaken]) 1);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    59
qed "Constrains_weaken";
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    60
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    61
(** Union **)
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    62
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    63
Goalw [Constrains_def]
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
    64
    "[| F : Constrains A A'; F : Constrains B B' |]   \
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
    65
\    ==> F : Constrains (A Un B) (A' Un B')";
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    66
by (blast_tac (claset() addIs [constrains_Un RS constrains_weaken]) 1);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    67
qed "Constrains_Un";
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    68
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
    69
Goal "ALL i:I. F : Constrains (A i) (A' i) \
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
    70
\     ==> F : Constrains (UN i:I. A i) (UN i:I. A' i)";
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
    71
by (asm_full_simp_tac (simpset() addsimps [Constrains_def]) 1);
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    72
by (dtac ball_constrains_UN 1);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    73
by (blast_tac (claset() addIs [constrains_weaken]) 1);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    74
qed "ball_Constrains_UN";
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    75
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    76
(** Intersection **)
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    77
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    78
Goalw [Constrains_def]
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
    79
    "[| F : Constrains A A'; F : Constrains B B' |]   \
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
    80
\    ==> F : Constrains (A Int B) (A' Int B')";
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    81
by (blast_tac (claset() addIs [constrains_Int RS constrains_weaken]) 1);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    82
qed "Constrains_Int";
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    83
5804
8e0a4c4fd67b Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents: 5784
diff changeset
    84
Goal "ALL i:I. F : Constrains (A i) (A' i)   \
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
    85
\     ==> F : Constrains (INT i:I. A i) (INT i:I. A' i)";
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
    86
by (asm_full_simp_tac (simpset() addsimps [Constrains_def]) 1);
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    87
by (dtac ball_constrains_INT 1);
5340
d75c03cf77b5 Misc changes
paulson
parents: 5319
diff changeset
    88
by (dtac constrains_reachable_Int 1);
d75c03cf77b5 Misc changes
paulson
parents: 5319
diff changeset
    89
by (blast_tac (claset() addIs [constrains_weaken]) 1);
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    90
qed "ball_Constrains_INT";
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    91
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
    92
Goal "F : Constrains A A' ==> reachable F Int A <= A'";
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
    93
by (asm_full_simp_tac (simpset() addsimps [Constrains_def]) 1);
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    94
by (dtac constrains_imp_subset 1);
5584
aad639e56d4e Now id:(Acts prg) is implicit
paulson
parents: 5536
diff changeset
    95
by (ALLGOALS
aad639e56d4e Now id:(Acts prg) is implicit
paulson
parents: 5536
diff changeset
    96
    (full_simp_tac (simpset() addsimps [Int_subset_iff, Int_lower1])));
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    97
qed "Constrains_imp_subset";
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    98
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
    99
Goalw [Constrains_def]
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   100
    "[| F : Constrains A B; F : Constrains B C |]   \
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   101
\    ==> F : Constrains A C";
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   102
by (blast_tac (claset() addIs [constrains_trans, constrains_weaken]) 1);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   103
qed "Constrains_trans";
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   104
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   105
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   106
(*** Stable ***)
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   107
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   108
Goal "(F : Stable A) = (F : stable (reachable F Int A))";
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   109
by (simp_tac (simpset() addsimps [Stable_def, Constrains_def, stable_def]) 1);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   110
qed "Stable_eq_stable";
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   111
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   112
Goalw [Stable_def] "F : Constrains A A ==> F : Stable A";
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   113
by (assume_tac 1);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   114
qed "StableI";
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   115
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   116
Goalw [Stable_def] "F : Stable A ==> F : Constrains A A";
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   117
by (assume_tac 1);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   118
qed "StableD";
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   119
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   120
Goalw [Stable_def]
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   121
    "[| F : Stable A; F : Stable A' |] ==> F : Stable (A Un A')";
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   122
by (blast_tac (claset() addIs [Constrains_Un]) 1);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   123
qed "Stable_Un";
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   124
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   125
Goalw [Stable_def]
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   126
    "[| F : Stable A; F : Stable A' |] ==> F : Stable (A Int A')";
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   127
by (blast_tac (claset() addIs [Constrains_Int]) 1);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   128
qed "Stable_Int";
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   129
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   130
Goalw [Stable_def]
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   131
    "[| F : Stable C; F : Constrains A (C Un A') |]   \
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   132
\    ==> F : Constrains (C Un A) (C Un A')";
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   133
by (blast_tac (claset() addIs [Constrains_Un RS Constrains_weaken]) 1);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   134
qed "Stable_Constrains_Un";
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   135
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   136
Goalw [Stable_def]
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   137
    "[| F : Stable C; F : Constrains (C Int A) A' |]   \
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   138
\    ==> F : Constrains (C Int A) (C Int A')";
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   139
by (blast_tac (claset() addIs [Constrains_Int RS Constrains_weaken]) 1);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   140
qed "Stable_Constrains_Int";
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   141
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   142
Goalw [Stable_def]
5804
8e0a4c4fd67b Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents: 5784
diff changeset
   143
    "(ALL i:I. F : Stable (A i)) ==> F : Stable (UN i:I. A i)";
8e0a4c4fd67b Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents: 5784
diff changeset
   144
by (etac ball_Constrains_UN 1);
8e0a4c4fd67b Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents: 5784
diff changeset
   145
qed "ball_Stable_UN";
8e0a4c4fd67b Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents: 5784
diff changeset
   146
8e0a4c4fd67b Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents: 5784
diff changeset
   147
Goalw [Stable_def]
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   148
    "(ALL i:I. F : Stable (A i)) ==> F : Stable (INT i:I. A i)";
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   149
by (etac ball_Constrains_INT 1);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   150
qed "ball_Stable_INT";
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   151
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   152
Goal "F : Stable (reachable F)";
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   153
by (simp_tac (simpset() addsimps [Stable_eq_stable, stable_reachable]) 1);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   154
qed "Stable_reachable";
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   155
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   156
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   157
5784
54276fba8420 the Increasing operator
paulson
parents: 5669
diff changeset
   158
(*** Increasing ***)
54276fba8420 the Increasing operator
paulson
parents: 5669
diff changeset
   159
54276fba8420 the Increasing operator
paulson
parents: 5669
diff changeset
   160
Goalw [Increasing_def, Stable_def, Constrains_def, stable_def, constrains_def]
54276fba8420 the Increasing operator
paulson
parents: 5669
diff changeset
   161
     "Increasing f <= Increasing (length o f)";
54276fba8420 the Increasing operator
paulson
parents: 5669
diff changeset
   162
by Auto_tac;
54276fba8420 the Increasing operator
paulson
parents: 5669
diff changeset
   163
by (blast_tac (claset() addIs [prefix_length_le, le_trans]) 1);
5804
8e0a4c4fd67b Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents: 5784
diff changeset
   164
qed "Increasing_size";
5784
54276fba8420 the Increasing operator
paulson
parents: 5669
diff changeset
   165
54276fba8420 the Increasing operator
paulson
parents: 5669
diff changeset
   166
Goalw [Increasing_def]
54276fba8420 the Increasing operator
paulson
parents: 5669
diff changeset
   167
     "Increasing f <= {F. ALL z::nat. F: Stable {s. z < f s}}";
54276fba8420 the Increasing operator
paulson
parents: 5669
diff changeset
   168
by (simp_tac (simpset() addsimps [Suc_le_eq RS sym]) 1);
54276fba8420 the Increasing operator
paulson
parents: 5669
diff changeset
   169
by (Blast_tac 1);
5804
8e0a4c4fd67b Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents: 5784
diff changeset
   170
qed "Increasing_Stable_less";
5784
54276fba8420 the Increasing operator
paulson
parents: 5669
diff changeset
   171
54276fba8420 the Increasing operator
paulson
parents: 5669
diff changeset
   172
Goalw [increasing_def, Increasing_def]
54276fba8420 the Increasing operator
paulson
parents: 5669
diff changeset
   173
     "F : increasing f ==> F : Increasing f";
54276fba8420 the Increasing operator
paulson
parents: 5669
diff changeset
   174
by (blast_tac (claset() addIs [stable_imp_Stable]) 1);
54276fba8420 the Increasing operator
paulson
parents: 5669
diff changeset
   175
qed "increasing_imp_Increasing";
54276fba8420 the Increasing operator
paulson
parents: 5669
diff changeset
   176
54276fba8420 the Increasing operator
paulson
parents: 5669
diff changeset
   177
54276fba8420 the Increasing operator
paulson
parents: 5669
diff changeset
   178
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   179
(*** The Elimination Theorem.  The "free" m has become universally quantified!
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   180
     Should the premise be !!m instead of ALL m ?  Would make it harder to use
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   181
     in forward proof. ***)
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   182
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   183
Goalw [Constrains_def, constrains_def]
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   184
    "[| ALL m. F : Constrains {s. s x = m} (B m) |] \
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   185
\    ==> F : Constrains {s. s x : M} (UN m:M. B m)";
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   186
by (Blast_tac 1);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   187
qed "Elimination";
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   188
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   189
(*As above, but for the trivial case of a one-variable state, in which the
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   190
  state is identified with its one variable.*)
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   191
Goalw [Constrains_def, constrains_def]
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   192
    "(ALL m. F : Constrains {m} (B m)) ==> F : Constrains M (UN m:M. B m)";
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   193
by (Blast_tac 1);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   194
qed "Elimination_sing";
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   195
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   196
Goalw [Constrains_def, constrains_def]
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   197
   "[| F : Constrains A (A' Un B); F : Constrains B B' |] \
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   198
\   ==> F : Constrains A (A' Un B')";
5669
f5d9caafc3bd added Clarify_tac to speed up proofs
paulson
parents: 5648
diff changeset
   199
by (Clarify_tac 1);
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   200
by (Blast_tac 1);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   201
qed "Constrains_cancel";
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   202
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   203
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   204
(*** Specialized laws for handling Invariants ***)
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   205
5620
3ac11c4af76a tidying and renaming
paulson
parents: 5584
diff changeset
   206
(** Natural deduction rules for "Invariant F A" **)
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   207
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   208
Goal "[| Init F<=A;  F : Stable A |] ==> F : Invariant A";
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   209
by (asm_simp_tac (simpset() addsimps [Invariant_def]) 1);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   210
qed "InvariantI";
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   211
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   212
Goal "F : Invariant A ==> Init F<=A & F : Stable A";
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   213
by (asm_full_simp_tac (simpset() addsimps [Invariant_def]) 1);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   214
qed "InvariantD";
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   215
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   216
bind_thm ("InvariantE", InvariantD RS conjE);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   217
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   218
5804
8e0a4c4fd67b Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents: 5784
diff changeset
   219
(*The set of all reachable states is the strongest Invariant*)
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   220
Goal "F : Invariant A ==> reachable F <= A";
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   221
by (full_simp_tac 
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   222
    (simpset() addsimps [Stable_def, Constrains_def, constrains_def, 
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   223
			 Invariant_def]) 1);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   224
by (rtac subsetI 1);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   225
by (etac reachable.induct 1);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   226
by (REPEAT (blast_tac (claset() addIs reachable.intrs) 1));
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   227
qed "Invariant_includes_reachable";
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   228
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   229
Goalw [Invariant_def, invariant_def, Stable_def, Constrains_def, stable_def]
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   230
     "F : invariant A ==> F : Invariant A";
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   231
by (blast_tac (claset() addIs [constrains_reachable_Int]) 1);
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   232
qed "invariant_imp_Invariant";
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   233
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   234
Goalw [Invariant_def, invariant_def, Stable_def, Constrains_def, stable_def]
5804
8e0a4c4fd67b Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents: 5784
diff changeset
   235
     "Invariant A = {F. F : invariant (reachable F Int A)}";
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   236
by (blast_tac (claset() addIs reachable.intrs) 1);
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   237
qed "Invariant_eq_invariant_reachable";
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   238
5804
8e0a4c4fd67b Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents: 5784
diff changeset
   239
(*Invariant is the "always" notion*)
8e0a4c4fd67b Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents: 5784
diff changeset
   240
Goal "Invariant A = {F. reachable F <= A}";
8e0a4c4fd67b Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents: 5784
diff changeset
   241
by (auto_tac (claset() addDs [invariant_includes_reachable],
8e0a4c4fd67b Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents: 5784
diff changeset
   242
	      simpset() addsimps [Int_absorb2, invariant_reachable,
8e0a4c4fd67b Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents: 5784
diff changeset
   243
				  Invariant_eq_invariant_reachable]));
8e0a4c4fd67b Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents: 5784
diff changeset
   244
qed "Invariant_eq_always";
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   245
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   246
5804
8e0a4c4fd67b Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents: 5784
diff changeset
   247
Goal "Invariant A = (UN I: Pow A. invariant I)";
8e0a4c4fd67b Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents: 5784
diff changeset
   248
by (simp_tac (simpset() addsimps [Invariant_eq_always]) 1);
8e0a4c4fd67b Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents: 5784
diff changeset
   249
by (blast_tac (claset() addIs [invariantI, impOfSubs Init_subset_reachable, 
8e0a4c4fd67b Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents: 5784
diff changeset
   250
                               stable_reachable,
8e0a4c4fd67b Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents: 5784
diff changeset
   251
			       impOfSubs invariant_includes_reachable]) 1);
8e0a4c4fd67b Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents: 5784
diff changeset
   252
qed "Invariant_eq_UN_invariant";
8e0a4c4fd67b Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents: 5784
diff changeset
   253
8e0a4c4fd67b Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents: 5784
diff changeset
   254
8e0a4c4fd67b Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents: 5784
diff changeset
   255
(*** "Constrains" rules involving Invariant ***)
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   256
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   257
Goal "[| F : Invariant INV;  F : Constrains (INV Int A) A' |]   \
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   258
\     ==> F : Constrains A A'";
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   259
by (asm_full_simp_tac
5804
8e0a4c4fd67b Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents: 5784
diff changeset
   260
    (simpset() addsimps [Invariant_includes_reachable RS Int_absorb2,
8e0a4c4fd67b Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents: 5784
diff changeset
   261
			 Constrains_def, Int_assoc RS sym]) 1);
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   262
qed "Invariant_ConstrainsI";
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   263
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   264
(* [| F : Invariant INV; F : Constrains (INV Int A) A |]
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   265
   ==> F : Stable A *)
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   266
bind_thm ("Invariant_StableI", Invariant_ConstrainsI RS StableI);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   267
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   268
Goal "[| F : Invariant INV;  F : Constrains A A' |]   \
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   269
\     ==> F : Constrains A (INV Int A')";
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   270
by (asm_full_simp_tac
5804
8e0a4c4fd67b Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents: 5784
diff changeset
   271
    (simpset() addsimps [Invariant_includes_reachable RS Int_absorb2,
8e0a4c4fd67b Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents: 5784
diff changeset
   272
			 Constrains_def, Int_assoc RS sym]) 1);
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   273
qed "Invariant_ConstrainsD";
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   274
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   275
bind_thm ("Invariant_StableD", StableD RSN (2,Invariant_ConstrainsD));
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   276
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   277
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   278
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   279
(** Conjoining Invariants **)
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   280
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   281
Goal "[| F : Invariant A;  F : Invariant B |] ==> F : Invariant (A Int B)";
5804
8e0a4c4fd67b Revising the Client proof as suggested by Michel Charpentier. New lemmas
paulson
parents: 5784
diff changeset
   282
by (auto_tac (claset(), simpset() addsimps [Invariant_eq_always]));
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   283
qed "Invariant_Int";
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   284
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   285
(*Delete the nearest invariance assumption (which will be the second one
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   286
  used by Invariant_Int) *)
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   287
val Invariant_thin =
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   288
    read_instantiate_sg (sign_of thy)
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   289
                [("V", "?F : Invariant ?A")] thin_rl;
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   290
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   291
(*Combines two invariance ASSUMPTIONS into one.  USEFUL??*)
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   292
val Invariant_Int_tac = dtac Invariant_Int THEN' 
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   293
                        assume_tac THEN'
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   294
			etac Invariant_thin;
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   295
5319
7356d0c88b1b Moved Un_subset_iff and Int_subset_iff from UNITY to equalities.ML
paulson
parents: 5313
diff changeset
   296
(*Combines a list of invariance THEOREMS into one.*)
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   297
val Invariant_Int_rule = foldr1 (fn (th1,th2) => [th1,th2] MRS Invariant_Int);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   298
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   299
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   300
(*To allow expansion of the program's definition when appropriate*)
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   301
val program_defs_ref = ref ([] : thm list);
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   302
5422
578dc167453f Moved constrains_tac from SubstAx to Constrains.
paulson
parents: 5340
diff changeset
   303
(*proves "constrains" properties when the program is specified*)
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   304
fun constrains_tac i = 
5422
578dc167453f Moved constrains_tac from SubstAx to Constrains.
paulson
parents: 5340
diff changeset
   305
   SELECT_GOAL
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   306
      (EVERY [simp_tac (simpset() addsimps !program_defs_ref) 1,
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   307
	      REPEAT (resolve_tac [StableI, stableI,
5422
578dc167453f Moved constrains_tac from SubstAx to Constrains.
paulson
parents: 5340
diff changeset
   308
				   constrains_imp_Constrains] 1),
578dc167453f Moved constrains_tac from SubstAx to Constrains.
paulson
parents: 5340
diff changeset
   309
	      rtac constrainsI 1,
5426
566f47250bd0 A new approach, using simp_of_act and simp_of_set to activate definitions when
paulson
parents: 5422
diff changeset
   310
	      Full_simp_tac 1,
5620
3ac11c4af76a tidying and renaming
paulson
parents: 5584
diff changeset
   311
	      REPEAT (FIRSTGOAL (etac disjE)),
5422
578dc167453f Moved constrains_tac from SubstAx to Constrains.
paulson
parents: 5340
diff changeset
   312
	      ALLGOALS Clarify_tac,
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   313
	      ALLGOALS Asm_full_simp_tac]) i;
5422
578dc167453f Moved constrains_tac from SubstAx to Constrains.
paulson
parents: 5340
diff changeset
   314
578dc167453f Moved constrains_tac from SubstAx to Constrains.
paulson
parents: 5340
diff changeset
   315