src/HOL/UNITY/Constrains.ML
author berghofe
Fri, 23 Oct 1998 22:36:15 +0200
changeset 5760 7e2cf2820684
parent 5669 f5d9caafc3bd
child 5784 54276fba8420
permissions -rw-r--r--
Added theorems True_not_False and False_not_True (for rep_datatype).
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
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
    84
Goal "[| ALL i:I. F : Constrains (A i) (A' i) |]   \
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]
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   143
    "(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
   144
by (etac ball_Constrains_INT 1);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   145
qed "ball_Stable_INT";
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   146
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   147
Goal "F : Stable (reachable F)";
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   148
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
   149
qed "Stable_reachable";
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   150
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   151
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   152
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   153
(*** 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
   154
     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
   155
     in forward proof. ***)
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
Goalw [Constrains_def, constrains_def]
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   158
    "[| ALL m. F : Constrains {s. s x = m} (B m) |] \
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   159
\    ==> 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
   160
by (Blast_tac 1);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   161
qed "Elimination";
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   162
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   163
(*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
   164
  state is identified with its one variable.*)
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   165
Goalw [Constrains_def, constrains_def]
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   166
    "(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
   167
by (Blast_tac 1);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   168
qed "Elimination_sing";
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   169
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   170
Goalw [Constrains_def, constrains_def]
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   171
   "[| F : Constrains A (A' Un B); F : Constrains B B' |] \
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   172
\   ==> F : Constrains A (A' Un B')";
5669
f5d9caafc3bd added Clarify_tac to speed up proofs
paulson
parents: 5648
diff changeset
   173
by (Clarify_tac 1);
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   174
by (Blast_tac 1);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   175
qed "Constrains_cancel";
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   176
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   177
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   178
(*** Specialized laws for handling Invariants ***)
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   179
5620
3ac11c4af76a tidying and renaming
paulson
parents: 5584
diff changeset
   180
(** Natural deduction rules for "Invariant F A" **)
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   181
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   182
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
   183
by (asm_simp_tac (simpset() addsimps [Invariant_def]) 1);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   184
qed "InvariantI";
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   185
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   186
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
   187
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
   188
qed "InvariantD";
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   189
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   190
bind_thm ("InvariantE", InvariantD RS conjE);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   191
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   192
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   193
(*The set of all reachable states is an Invariant...*)
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   194
Goal "F : Invariant (reachable F)";
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   195
by (simp_tac (simpset() addsimps [Invariant_def]) 1);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   196
by (blast_tac (claset() addIs (Stable_reachable::reachable.intrs)) 1);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   197
qed "Invariant_reachable";
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   198
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   199
(*...in fact the strongest Invariant!*)
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   200
Goal "F : Invariant A ==> reachable F <= A";
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   201
by (full_simp_tac 
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   202
    (simpset() addsimps [Stable_def, Constrains_def, constrains_def, 
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   203
			 Invariant_def]) 1);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   204
by (rtac subsetI 1);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   205
by (etac reachable.induct 1);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   206
by (REPEAT (blast_tac (claset() addIs reachable.intrs) 1));
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   207
qed "Invariant_includes_reachable";
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   208
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   209
Goalw [Invariant_def, invariant_def, Stable_def, Constrains_def, stable_def]
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   210
     "F : invariant A ==> F : Invariant A";
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   211
by (blast_tac (claset() addIs [constrains_reachable_Int]) 1);
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   212
qed "invariant_imp_Invariant";
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   213
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   214
Goalw [Invariant_def, invariant_def, Stable_def, Constrains_def, stable_def]
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   215
     "(F : Invariant A) = (F : invariant (reachable F Int A))";
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   216
by (blast_tac (claset() addIs reachable.intrs) 1);
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   217
qed "Invariant_eq_invariant_reachable";
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   218
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   219
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   220
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   221
Goal "F : Invariant INV ==> reachable F Int INV = reachable F";
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   222
by (dtac Invariant_includes_reachable 1);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   223
by (Blast_tac 1);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   224
qed "reachable_Int_INV";
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   225
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   226
Goal "[| F : Invariant INV;  F : Constrains (INV Int A) A' |]   \
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   227
\     ==> F : Constrains A A'";
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   228
by (asm_full_simp_tac
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   229
    (simpset() addsimps [Constrains_def, reachable_Int_INV,
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   230
			 Int_assoc RS sym]) 1);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   231
qed "Invariant_ConstrainsI";
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   232
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   233
(* [| F : Invariant INV; F : Constrains (INV Int A) A |]
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   234
   ==> F : Stable A *)
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   235
bind_thm ("Invariant_StableI", Invariant_ConstrainsI RS StableI);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   236
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   237
Goal "[| F : Invariant INV;  F : Constrains A A' |]   \
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   238
\     ==> F : Constrains A (INV Int A')";
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   239
by (asm_full_simp_tac
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   240
    (simpset() addsimps [Constrains_def, reachable_Int_INV,
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   241
			 Int_assoc RS sym]) 1);
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   242
qed "Invariant_ConstrainsD";
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   243
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   244
bind_thm ("Invariant_StableD", StableD RSN (2,Invariant_ConstrainsD));
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   245
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   246
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   247
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   248
(** Conjoining Invariants **)
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   249
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   250
Goal "[| F : Invariant A;  F : Invariant B |] ==> F : Invariant (A Int B)";
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   251
by (auto_tac (claset(),
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   252
	      simpset() addsimps [Invariant_def, Stable_Int]));
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   253
qed "Invariant_Int";
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   254
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   255
(*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
   256
  used by Invariant_Int) *)
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   257
val Invariant_thin =
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   258
    read_instantiate_sg (sign_of thy)
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   259
                [("V", "?F : Invariant ?A")] thin_rl;
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   260
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   261
(*Combines two invariance ASSUMPTIONS into one.  USEFUL??*)
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   262
val Invariant_Int_tac = dtac Invariant_Int THEN' 
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   263
                        assume_tac THEN'
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   264
			etac Invariant_thin;
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   265
5319
7356d0c88b1b Moved Un_subset_iff and Int_subset_iff from UNITY to equalities.ML
paulson
parents: 5313
diff changeset
   266
(*Combines a list of invariance THEOREMS into one.*)
5313
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   267
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
   268
1861a564d7e2 Constrains, Stable, Invariant...more of the substitution axiom, but Union
paulson
parents:
diff changeset
   269
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   270
(*To allow expansion of the program's definition when appropriate*)
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   271
val program_defs_ref = ref ([] : thm list);
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   272
5422
578dc167453f Moved constrains_tac from SubstAx to Constrains.
paulson
parents: 5340
diff changeset
   273
(*proves "constrains" properties when the program is specified*)
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   274
fun constrains_tac i = 
5422
578dc167453f Moved constrains_tac from SubstAx to Constrains.
paulson
parents: 5340
diff changeset
   275
   SELECT_GOAL
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   276
      (EVERY [simp_tac (simpset() addsimps !program_defs_ref) 1,
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   277
	      REPEAT (resolve_tac [StableI, stableI,
5422
578dc167453f Moved constrains_tac from SubstAx to Constrains.
paulson
parents: 5340
diff changeset
   278
				   constrains_imp_Constrains] 1),
578dc167453f Moved constrains_tac from SubstAx to Constrains.
paulson
parents: 5340
diff changeset
   279
	      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
   280
	      Full_simp_tac 1,
5620
3ac11c4af76a tidying and renaming
paulson
parents: 5584
diff changeset
   281
	      REPEAT (FIRSTGOAL (etac disjE)),
5422
578dc167453f Moved constrains_tac from SubstAx to Constrains.
paulson
parents: 5340
diff changeset
   282
	      ALLGOALS Clarify_tac,
5648
fe887910e32e specifications as sets of programs
paulson
parents: 5631
diff changeset
   283
	      ALLGOALS Asm_full_simp_tac]) i;
5422
578dc167453f Moved constrains_tac from SubstAx to Constrains.
paulson
parents: 5340
diff changeset
   284
578dc167453f Moved constrains_tac from SubstAx to Constrains.
paulson
parents: 5340
diff changeset
   285