src/HOL/UNITY/Extend.thy
author wenzelm
Sat, 04 Nov 2006 19:25:43 +0100
changeset 21177 e8228486aa03
parent 16417 9bc16273c2d4
child 32960 69916a850301
child 32988 d1d4d7a08a66
permissions -rw-r--r--
removed checkpoint interface; moved back to copy/checkpoint instead of checkpoint/checkpoint (NB 1: checkpoint is idempotent, i.e. impure data is being shared, which is inappropriate; NB 2: copying a checkpoint always produces a related theory); present_proof: proper treatment of history; tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
6297
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
     1
(*  Title:      HOL/UNITY/Extend.thy
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
     2
    ID:         $Id$
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
     4
    Copyright   1998  University of Cambridge
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
     5
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13790
diff changeset
     6
Extending of state setsExtending of state sets
6297
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
     7
  function f (forget)    maps the extended state to the original state
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
     8
  function g (forgotten) maps the extended state to the "extending part"
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
     9
*)
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    10
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13790
diff changeset
    11
header{*Extending State Sets*}
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13790
diff changeset
    12
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 13819
diff changeset
    13
theory Extend imports Guar begin
6297
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    14
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    15
constdefs
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    16
8948
b797cfa3548d restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
paulson
parents: 8703
diff changeset
    17
  (*MOVE to Relation.thy?*)
b797cfa3548d restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
paulson
parents: 8703
diff changeset
    18
  Restrict :: "[ 'a set, ('a*'b) set] => ('a*'b) set"
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
    19
    "Restrict A r == r \<inter> (A <*> UNIV)"
8948
b797cfa3548d restructuring: LessThan.ML mostly moved to HOL/SetInterval.ML
paulson
parents: 8703
diff changeset
    20
7482
7badd511844d working snapshot
paulson
parents: 7399
diff changeset
    21
  good_map :: "['a*'b => 'c] => bool"
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
    22
    "good_map h == surj h & (\<forall>x y. fst (inv h (h (x,y))) = x)"
7482
7badd511844d working snapshot
paulson
parents: 7399
diff changeset
    23
     (*Using the locale constant "f", this is  f (h (x,y))) = x*)
7badd511844d working snapshot
paulson
parents: 7399
diff changeset
    24
  
6297
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    25
  extend_set :: "['a*'b => 'c, 'a set] => 'c set"
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10064
diff changeset
    26
    "extend_set h A == h ` (A <*> UNIV)"
6297
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    27
7342
532841541d73 project constants
paulson
parents: 6677
diff changeset
    28
  project_set :: "['a*'b => 'c, 'c set] => 'a set"
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
    29
    "project_set h C == {x. \<exists>y. h(x,y) \<in> C}"
7342
532841541d73 project constants
paulson
parents: 6677
diff changeset
    30
532841541d73 project constants
paulson
parents: 6677
diff changeset
    31
  extend_act :: "['a*'b => 'c, ('a*'a) set] => ('c*'c) set"
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
    32
    "extend_act h == %act. \<Union>(s,s') \<in> act. \<Union>y. {(h(s,y), h(s',y))}"
6297
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    33
7878
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7826
diff changeset
    34
  project_act :: "['a*'b => 'c, ('c*'c) set] => ('a*'a) set"
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
    35
    "project_act h act == {(x,x'). \<exists>y y'. (h(x,y), h(x',y')) \<in> act}"
7342
532841541d73 project constants
paulson
parents: 6677
diff changeset
    36
6297
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    37
  extend :: "['a*'b => 'c, 'a program] => 'c program"
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    38
    "extend h F == mk_program (extend_set h (Init F),
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10064
diff changeset
    39
			       extend_act h ` Acts F,
a7897aebbffc *** empty log message ***
nipkow
parents: 10064
diff changeset
    40
			       project_act h -` AllowedActs F)"
6297
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    41
7878
43b03d412b82 working version with localTo[C] instead of localTo
paulson
parents: 7826
diff changeset
    42
  (*Argument C allows weak safety laws to be projected*)
7880
62fb24e28e5e exchanged the first two args of "project" and "drop_prog"
paulson
parents: 7878
diff changeset
    43
  project :: "['a*'b => 'c, 'c set, 'c program] => 'a program"
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 8948
diff changeset
    44
    "project h C F ==
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 8948
diff changeset
    45
       mk_program (project_set h (Init F),
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10064
diff changeset
    46
		   project_act h ` Restrict C ` Acts F,
10064
1a77667b21ef added compatibility relation: AllowedActs, Allowed, ok,
paulson
parents: 8948
diff changeset
    47
		   {act. Restrict (project_set h C) act :
10834
a7897aebbffc *** empty log message ***
nipkow
parents: 10064
diff changeset
    48
		         project_act h ` Restrict C ` AllowedActs F})"
7342
532841541d73 project constants
paulson
parents: 6677
diff changeset
    49
6297
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
    50
locale Extend =
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    51
  fixes f     :: "'c => 'a"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    52
    and g     :: "'c => 'b"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    53
    and h     :: "'a*'b => 'c"    (*isomorphism between 'a * 'b and 'c *)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    54
    and slice :: "['c set, 'b] => 'a set"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    55
  assumes
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    56
    good_h:  "good_map h"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    57
  defines f_def: "f z == fst (inv h z)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    58
      and g_def: "g z == snd (inv h z)"
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
    59
      and slice_def: "slice Z y == {x. h(x,y) \<in> Z}"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    60
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    61
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    62
(** These we prove OUTSIDE the locale. **)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    63
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    64
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13790
diff changeset
    65
subsection{*Restrict*}
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13790
diff changeset
    66
(*MOVE to Relation.thy?*)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    67
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
    68
lemma Restrict_iff [iff]: "((x,y): Restrict A r) = ((x,y): r & x \<in> A)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    69
by (unfold Restrict_def, blast)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    70
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    71
lemma Restrict_UNIV [simp]: "Restrict UNIV = id"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    72
apply (rule ext)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    73
apply (auto simp add: Restrict_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    74
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    75
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    76
lemma Restrict_empty [simp]: "Restrict {} r = {}"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    77
by (auto simp add: Restrict_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    78
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
    79
lemma Restrict_Int [simp]: "Restrict A (Restrict B r) = Restrict (A \<inter> B) r"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    80
by (unfold Restrict_def, blast)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    81
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
    82
lemma Restrict_triv: "Domain r \<subseteq> A ==> Restrict A r = r"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    83
by (unfold Restrict_def, auto)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    84
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
    85
lemma Restrict_subset: "Restrict A r \<subseteq> r"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    86
by (unfold Restrict_def, auto)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    87
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    88
lemma Restrict_eq_mono: 
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
    89
     "[| A \<subseteq> B;  Restrict B r = Restrict B s |]  
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    90
      ==> Restrict A r = Restrict A s"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    91
by (unfold Restrict_def, blast)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    92
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    93
lemma Restrict_imageI: 
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
    94
     "[| s \<in> RR;  Restrict A r = Restrict A s |]  
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
    95
      ==> Restrict A r \<in> Restrict A ` RR"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    96
by (unfold Restrict_def image_def, auto)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    97
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
    98
lemma Domain_Restrict [simp]: "Domain (Restrict A r) = A \<inter> Domain r"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
    99
by blast
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   100
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   101
lemma Image_Restrict [simp]: "(Restrict A r) `` B = r `` (A \<inter> B)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   102
by blast
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   103
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   104
(*Possibly easier than reasoning about "inv h"*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   105
lemma good_mapI: 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   106
     assumes surj_h: "surj h"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   107
	 and prem:   "!! x x' y y'. h(x,y) = h(x',y') ==> x=x'"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   108
     shows "good_map h"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   109
apply (simp add: good_map_def) 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   110
apply (safe intro!: surj_h)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   111
apply (rule prem)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   112
apply (subst surjective_pairing [symmetric])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   113
apply (subst surj_h [THEN surj_f_inv_f])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   114
apply (rule refl)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   115
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   116
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   117
lemma good_map_is_surj: "good_map h ==> surj h"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   118
by (unfold good_map_def, auto)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   119
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   120
(*A convenient way of finding a closed form for inv h*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   121
lemma fst_inv_equalityI: 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   122
     assumes surj_h: "surj h"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   123
	 and prem:   "!! x y. g (h(x,y)) = x"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   124
     shows "fst (inv h z) = g z"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   125
apply (unfold inv_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   126
apply (rule_tac y1 = z in surj_h [THEN surjD, THEN exE])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   127
apply (rule someI2)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   128
apply (drule_tac [2] f = g in arg_cong)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   129
apply (auto simp add: prem)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   130
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   131
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   132
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13790
diff changeset
   133
subsection{*Trivial properties of f, g, h*}
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   134
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   135
lemma (in Extend) f_h_eq [simp]: "f(h(x,y)) = x" 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   136
by (simp add: f_def good_h [unfolded good_map_def, THEN conjunct2])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   137
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   138
lemma (in Extend) h_inject1 [dest]: "h(x,y) = h(x',y') ==> x=x'"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   139
apply (drule_tac f = f in arg_cong)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   140
apply (simp add: f_def good_h [unfolded good_map_def, THEN conjunct2])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   141
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   142
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   143
lemma (in Extend) h_f_g_equiv: "h(f z, g z) == z"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   144
by (simp add: f_def g_def 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   145
            good_h [unfolded good_map_def, THEN conjunct1, THEN surj_f_inv_f])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   146
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   147
lemma (in Extend) h_f_g_eq: "h(f z, g z) = z"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   148
by (simp add: h_f_g_equiv)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   149
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   150
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   151
lemma (in Extend) split_extended_all:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   152
     "(!!z. PROP P z) == (!!u y. PROP P (h (u, y)))"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   153
proof 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   154
   assume allP: "\<And>z. PROP P z"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   155
   fix u y
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   156
   show "PROP P (h (u, y))" by (rule allP)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   157
 next
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   158
   assume allPh: "\<And>u y. PROP P (h(u,y))"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   159
   fix z
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   160
   have Phfgz: "PROP P (h (f z, g z))" by (rule allPh)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   161
   show "PROP P z" by (rule Phfgz [unfolded h_f_g_equiv])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   162
qed 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   163
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   164
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   165
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13790
diff changeset
   166
subsection{*@{term extend_set}: basic properties*}
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   167
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   168
lemma project_set_iff [iff]:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   169
     "(x \<in> project_set h C) = (\<exists>y. h(x,y) \<in> C)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   170
by (simp add: project_set_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   171
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   172
lemma extend_set_mono: "A \<subseteq> B ==> extend_set h A \<subseteq> extend_set h B"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   173
by (unfold extend_set_def, blast)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   174
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   175
lemma (in Extend) mem_extend_set_iff [iff]: "z \<in> extend_set h A = (f z \<in> A)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   176
apply (unfold extend_set_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   177
apply (force intro: h_f_g_eq [symmetric])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   178
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   179
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   180
lemma (in Extend) extend_set_strict_mono [iff]:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   181
     "(extend_set h A \<subseteq> extend_set h B) = (A \<subseteq> B)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   182
by (unfold extend_set_def, force)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   183
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   184
lemma extend_set_empty [simp]: "extend_set h {} = {}"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   185
by (unfold extend_set_def, auto)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   186
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   187
lemma (in Extend) extend_set_eq_Collect: "extend_set h {s. P s} = {s. P(f s)}"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   188
by auto
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   189
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   190
lemma (in Extend) extend_set_sing: "extend_set h {x} = {s. f s = x}"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   191
by auto
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   192
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   193
lemma (in Extend) extend_set_inverse [simp]:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   194
     "project_set h (extend_set h C) = C"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   195
by (unfold extend_set_def, auto)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   196
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   197
lemma (in Extend) extend_set_project_set:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   198
     "C \<subseteq> extend_set h (project_set h C)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   199
apply (unfold extend_set_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   200
apply (auto simp add: split_extended_all, blast)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   201
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   202
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   203
lemma (in Extend) inj_extend_set: "inj (extend_set h)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   204
apply (rule inj_on_inverseI)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   205
apply (rule extend_set_inverse)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   206
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   207
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   208
lemma (in Extend) extend_set_UNIV_eq [simp]: "extend_set h UNIV = UNIV"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   209
apply (unfold extend_set_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   210
apply (auto simp add: split_extended_all)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   211
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   212
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13790
diff changeset
   213
subsection{*@{term project_set}: basic properties*}
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   214
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   215
(*project_set is simply image!*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   216
lemma (in Extend) project_set_eq: "project_set h C = f ` C"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   217
by (auto intro: f_h_eq [symmetric] simp add: split_extended_all)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   218
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   219
(*Converse appears to fail*)
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   220
lemma (in Extend) project_set_I: "!!z. z \<in> C ==> f z \<in> project_set h C"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   221
by (auto simp add: split_extended_all)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   222
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   223
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13790
diff changeset
   224
subsection{*More laws*}
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   225
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   226
(*Because A and B could differ on the "other" part of the state, 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   227
   cannot generalize to 
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   228
      project_set h (A \<inter> B) = project_set h A \<inter> project_set h B
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   229
*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   230
lemma (in Extend) project_set_extend_set_Int:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   231
     "project_set h ((extend_set h A) \<inter> B) = A \<inter> (project_set h B)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   232
by auto
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   233
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   234
(*Unused, but interesting?*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   235
lemma (in Extend) project_set_extend_set_Un:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   236
     "project_set h ((extend_set h A) \<union> B) = A \<union> (project_set h B)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   237
by auto
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   238
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   239
lemma project_set_Int_subset:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   240
     "project_set h (A \<inter> B) \<subseteq> (project_set h A) \<inter> (project_set h B)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   241
by auto
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   242
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   243
lemma (in Extend) extend_set_Un_distrib:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   244
     "extend_set h (A \<union> B) = extend_set h A \<union> extend_set h B"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   245
by auto
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   246
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   247
lemma (in Extend) extend_set_Int_distrib:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   248
     "extend_set h (A \<inter> B) = extend_set h A \<inter> extend_set h B"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   249
by auto
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   250
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   251
lemma (in Extend) extend_set_INT_distrib:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   252
     "extend_set h (INTER A B) = (\<Inter>x \<in> A. extend_set h (B x))"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   253
by auto
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   254
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   255
lemma (in Extend) extend_set_Diff_distrib:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   256
     "extend_set h (A - B) = extend_set h A - extend_set h B"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   257
by auto
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   258
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   259
lemma (in Extend) extend_set_Union:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   260
     "extend_set h (Union A) = (\<Union>X \<in> A. extend_set h X)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   261
by blast
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   262
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   263
lemma (in Extend) extend_set_subset_Compl_eq:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   264
     "(extend_set h A \<subseteq> - extend_set h B) = (A \<subseteq> - B)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   265
by (unfold extend_set_def, auto)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   266
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   267
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13790
diff changeset
   268
subsection{*@{term extend_act}*}
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   269
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   270
(*Can't strengthen it to
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   271
  ((h(s,y), h(s',y')) \<in> extend_act h act) = ((s, s') \<in> act & y=y')
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   272
  because h doesn't have to be injective in the 2nd argument*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   273
lemma (in Extend) mem_extend_act_iff [iff]: 
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   274
     "((h(s,y), h(s',y)) \<in> extend_act h act) = ((s, s') \<in> act)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   275
by (unfold extend_act_def, auto)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   276
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   277
(*Converse fails: (z,z') would include actions that changed the g-part*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   278
lemma (in Extend) extend_act_D: 
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   279
     "(z, z') \<in> extend_act h act ==> (f z, f z') \<in> act"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   280
by (unfold extend_act_def, auto)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   281
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   282
lemma (in Extend) extend_act_inverse [simp]: 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   283
     "project_act h (extend_act h act) = act"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   284
by (unfold extend_act_def project_act_def, blast)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   285
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   286
lemma (in Extend) project_act_extend_act_restrict [simp]: 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   287
     "project_act h (Restrict C (extend_act h act)) =  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   288
      Restrict (project_set h C) act"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   289
by (unfold extend_act_def project_act_def, blast)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   290
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   291
lemma (in Extend) subset_extend_act_D: 
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   292
     "act' \<subseteq> extend_act h act ==> project_act h act' \<subseteq> act"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   293
by (unfold extend_act_def project_act_def, force)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   294
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   295
lemma (in Extend) inj_extend_act: "inj (extend_act h)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   296
apply (rule inj_on_inverseI)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   297
apply (rule extend_act_inverse)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   298
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   299
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   300
lemma (in Extend) extend_act_Image [simp]: 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   301
     "extend_act h act `` (extend_set h A) = extend_set h (act `` A)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   302
by (unfold extend_set_def extend_act_def, force)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   303
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   304
lemma (in Extend) extend_act_strict_mono [iff]:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   305
     "(extend_act h act' \<subseteq> extend_act h act) = (act'<=act)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   306
by (unfold extend_act_def, auto)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   307
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   308
declare (in Extend) inj_extend_act [THEN inj_eq, iff]
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   309
(*This theorem is  (extend_act h act' = extend_act h act) = (act'=act) *)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   310
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   311
lemma Domain_extend_act: 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   312
    "Domain (extend_act h act) = extend_set h (Domain act)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   313
by (unfold extend_set_def extend_act_def, force)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   314
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   315
lemma (in Extend) extend_act_Id [simp]: 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   316
    "extend_act h Id = Id"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   317
apply (unfold extend_act_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   318
apply (force intro: h_f_g_eq [symmetric])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   319
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   320
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   321
lemma (in Extend) project_act_I: 
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   322
     "!!z z'. (z, z') \<in> act ==> (f z, f z') \<in> project_act h act"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   323
apply (unfold project_act_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   324
apply (force simp add: split_extended_all)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   325
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   326
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   327
lemma (in Extend) project_act_Id [simp]: "project_act h Id = Id"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   328
by (unfold project_act_def, force)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   329
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   330
lemma (in Extend) Domain_project_act: 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   331
  "Domain (project_act h act) = project_set h (Domain act)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   332
apply (unfold project_act_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   333
apply (force simp add: split_extended_all)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   334
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   335
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   336
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   337
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
   338
subsection{*extend*}
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   339
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
   340
text{*Basic properties*}
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   341
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   342
lemma Init_extend [simp]:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   343
     "Init (extend h F) = extend_set h (Init F)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   344
by (unfold extend_def, auto)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   345
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   346
lemma Init_project [simp]:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   347
     "Init (project h C F) = project_set h (Init F)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   348
by (unfold project_def, auto)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   349
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   350
lemma (in Extend) Acts_extend [simp]:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   351
     "Acts (extend h F) = (extend_act h ` Acts F)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   352
by (simp add: extend_def insert_Id_image_Acts)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   353
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   354
lemma (in Extend) AllowedActs_extend [simp]:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   355
     "AllowedActs (extend h F) = project_act h -` AllowedActs F"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   356
by (simp add: extend_def insert_absorb)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   357
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   358
lemma Acts_project [simp]:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   359
     "Acts(project h C F) = insert Id (project_act h ` Restrict C ` Acts F)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   360
by (auto simp add: project_def image_iff)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   361
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   362
lemma (in Extend) AllowedActs_project [simp]:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   363
     "AllowedActs(project h C F) =  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   364
        {act. Restrict (project_set h C) act  
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   365
               \<in> project_act h ` Restrict C ` AllowedActs F}"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   366
apply (simp (no_asm) add: project_def image_iff)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   367
apply (subst insert_absorb)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   368
apply (auto intro!: bexI [of _ Id] simp add: project_act_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   369
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   370
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   371
lemma (in Extend) Allowed_extend:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   372
     "Allowed (extend h F) = project h UNIV -` Allowed F"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   373
apply (simp (no_asm) add: AllowedActs_extend Allowed_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   374
apply blast
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   375
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   376
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   377
lemma (in Extend) extend_SKIP [simp]: "extend h SKIP = SKIP"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   378
apply (unfold SKIP_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   379
apply (rule program_equalityI, auto)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   380
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   381
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   382
lemma project_set_UNIV [simp]: "project_set h UNIV = UNIV"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   383
by auto
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   384
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   385
lemma project_set_Union:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   386
     "project_set h (Union A) = (\<Union>X \<in> A. project_set h X)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   387
by blast
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   388
6297
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
   389
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   390
(*Converse FAILS: the extended state contributing to project_set h C
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   391
  may not coincide with the one contributing to project_act h act*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   392
lemma (in Extend) project_act_Restrict_subset:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   393
     "project_act h (Restrict C act) \<subseteq>  
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   394
      Restrict (project_set h C) (project_act h act)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   395
by (auto simp add: project_act_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   396
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   397
lemma (in Extend) project_act_Restrict_Id_eq:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   398
     "project_act h (Restrict C Id) = Restrict (project_set h C) Id"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   399
by (auto simp add: project_act_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   400
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   401
lemma (in Extend) project_extend_eq:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   402
     "project h C (extend h F) =  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   403
      mk_program (Init F, Restrict (project_set h C) ` Acts F,  
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   404
                  {act. Restrict (project_set h C) act 
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   405
                          \<in> project_act h ` Restrict C ` 
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   406
                                     (project_act h -` AllowedActs F)})"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   407
apply (rule program_equalityI)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   408
  apply simp
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   409
 apply (simp add: image_eq_UN)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   410
apply (simp add: project_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   411
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   412
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   413
lemma (in Extend) extend_inverse [simp]:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   414
     "project h UNIV (extend h F) = F"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   415
apply (simp (no_asm_simp) add: project_extend_eq image_eq_UN
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   416
          subset_UNIV [THEN subset_trans, THEN Restrict_triv])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   417
apply (rule program_equalityI)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   418
apply (simp_all (no_asm))
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   419
apply (subst insert_absorb)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   420
apply (simp (no_asm) add: bexI [of _ Id])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   421
apply auto
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   422
apply (rename_tac "act")
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   423
apply (rule_tac x = "extend_act h act" in bexI, auto)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   424
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   425
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   426
lemma (in Extend) inj_extend: "inj (extend h)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   427
apply (rule inj_on_inverseI)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   428
apply (rule extend_inverse)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   429
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   430
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   431
lemma (in Extend) extend_Join [simp]:
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   432
     "extend h (F\<squnion>G) = extend h F\<squnion>extend h G"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   433
apply (rule program_equalityI)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   434
apply (simp (no_asm) add: extend_set_Int_distrib)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   435
apply (simp add: image_Un, auto)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   436
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   437
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   438
lemma (in Extend) extend_JN [simp]:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   439
     "extend h (JOIN I F) = (\<Squnion>i \<in> I. extend h (F i))"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   440
apply (rule program_equalityI)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   441
  apply (simp (no_asm) add: extend_set_INT_distrib)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   442
 apply (simp add: image_UN, auto)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   443
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   444
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   445
(** These monotonicity results look natural but are UNUSED **)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   446
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   447
lemma (in Extend) extend_mono: "F \<le> G ==> extend h F \<le> extend h G"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   448
by (force simp add: component_eq_subset)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   449
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   450
lemma (in Extend) project_mono: "F \<le> G ==> project h C F \<le> project h C G"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   451
by (simp add: component_eq_subset, blast)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   452
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
   453
lemma (in Extend) all_total_extend: "all_total F ==> all_total (extend h F)"
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
   454
by (simp add: all_total_def Domain_extend_act)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   455
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13790
diff changeset
   456
subsection{*Safety: co, stable*}
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   457
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   458
lemma (in Extend) extend_constrains:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   459
     "(extend h F \<in> (extend_set h A) co (extend_set h B)) =  
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   460
      (F \<in> A co B)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   461
by (simp add: constrains_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   462
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   463
lemma (in Extend) extend_stable:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   464
     "(extend h F \<in> stable (extend_set h A)) = (F \<in> stable A)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   465
by (simp add: stable_def extend_constrains)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   466
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   467
lemma (in Extend) extend_invariant:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   468
     "(extend h F \<in> invariant (extend_set h A)) = (F \<in> invariant A)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   469
by (simp add: invariant_def extend_stable)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   470
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   471
(*Projects the state predicates in the property satisfied by  extend h F.
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   472
  Converse fails: A and B may differ in their extra variables*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   473
lemma (in Extend) extend_constrains_project_set:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   474
     "extend h F \<in> A co B ==> F \<in> (project_set h A) co (project_set h B)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   475
by (auto simp add: constrains_def, force)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   476
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   477
lemma (in Extend) extend_stable_project_set:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   478
     "extend h F \<in> stable A ==> F \<in> stable (project_set h A)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   479
by (simp add: stable_def extend_constrains_project_set)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   480
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   481
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13790
diff changeset
   482
subsection{*Weak safety primitives: Co, Stable*}
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   483
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   484
lemma (in Extend) reachable_extend_f:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   485
     "p \<in> reachable (extend h F) ==> f p \<in> reachable F"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   486
apply (erule reachable.induct)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   487
apply (auto intro: reachable.intros simp add: extend_act_def image_iff)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   488
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   489
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   490
lemma (in Extend) h_reachable_extend:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   491
     "h(s,y) \<in> reachable (extend h F) ==> s \<in> reachable F"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   492
by (force dest!: reachable_extend_f)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   493
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   494
lemma (in Extend) reachable_extend_eq: 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   495
     "reachable (extend h F) = extend_set h (reachable F)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   496
apply (unfold extend_set_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   497
apply (rule equalityI)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   498
apply (force intro: h_f_g_eq [symmetric] dest!: reachable_extend_f, clarify)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   499
apply (erule reachable.induct)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   500
apply (force intro: reachable.intros)+
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   501
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   502
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   503
lemma (in Extend) extend_Constrains:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   504
     "(extend h F \<in> (extend_set h A) Co (extend_set h B)) =   
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   505
      (F \<in> A Co B)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   506
by (simp add: Constrains_def reachable_extend_eq extend_constrains 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   507
              extend_set_Int_distrib [symmetric])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   508
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   509
lemma (in Extend) extend_Stable:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   510
     "(extend h F \<in> Stable (extend_set h A)) = (F \<in> Stable A)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   511
by (simp add: Stable_def extend_Constrains)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   512
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   513
lemma (in Extend) extend_Always:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   514
     "(extend h F \<in> Always (extend_set h A)) = (F \<in> Always A)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   515
by (simp (no_asm_simp) add: Always_def extend_Stable)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   516
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   517
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   518
(** Safety and "project" **)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   519
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   520
(** projection: monotonicity for safety **)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   521
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   522
lemma project_act_mono:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   523
     "D \<subseteq> C ==>  
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   524
      project_act h (Restrict D act) \<subseteq> project_act h (Restrict C act)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   525
by (auto simp add: project_act_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   526
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   527
lemma (in Extend) project_constrains_mono:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   528
     "[| D \<subseteq> C; project h C F \<in> A co B |] ==> project h D F \<in> A co B"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   529
apply (auto simp add: constrains_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   530
apply (drule project_act_mono, blast)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   531
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   532
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   533
lemma (in Extend) project_stable_mono:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   534
     "[| D \<subseteq> C;  project h C F \<in> stable A |] ==> project h D F \<in> stable A"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   535
by (simp add: stable_def project_constrains_mono)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   536
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   537
(*Key lemma used in several proofs about project and co*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   538
lemma (in Extend) project_constrains: 
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   539
     "(project h C F \<in> A co B)  =   
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   540
      (F \<in> (C \<inter> extend_set h A) co (extend_set h B) & A \<subseteq> B)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   541
apply (unfold constrains_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   542
apply (auto intro!: project_act_I simp add: ball_Un)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   543
apply (force intro!: project_act_I dest!: subsetD)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   544
(*the <== direction*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   545
apply (unfold project_act_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   546
apply (force dest!: subsetD)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   547
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   548
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   549
lemma (in Extend) project_stable: 
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   550
     "(project h UNIV F \<in> stable A) = (F \<in> stable (extend_set h A))"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   551
apply (unfold stable_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   552
apply (simp (no_asm) add: project_constrains)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   553
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   554
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   555
lemma (in Extend) project_stable_I:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   556
     "F \<in> stable (extend_set h A) ==> project h C F \<in> stable A"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   557
apply (drule project_stable [THEN iffD2])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   558
apply (blast intro: project_stable_mono)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   559
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   560
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   561
lemma (in Extend) Int_extend_set_lemma:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   562
     "A \<inter> extend_set h ((project_set h A) \<inter> B) = A \<inter> extend_set h B"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   563
by (auto simp add: split_extended_all)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   564
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   565
(*Strange (look at occurrences of C) but used in leadsETo proofs*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   566
lemma project_constrains_project_set:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   567
     "G \<in> C co B ==> project h C G \<in> project_set h C co project_set h B"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   568
by (simp add: constrains_def project_def project_act_def, blast)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   569
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   570
lemma project_stable_project_set:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   571
     "G \<in> stable C ==> project h C G \<in> stable (project_set h C)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   572
by (simp add: stable_def project_constrains_project_set)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   573
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   574
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13790
diff changeset
   575
subsection{*Progress: transient, ensures*}
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   576
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   577
lemma (in Extend) extend_transient:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   578
     "(extend h F \<in> transient (extend_set h A)) = (F \<in> transient A)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   579
by (auto simp add: transient_def extend_set_subset_Compl_eq Domain_extend_act)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   580
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   581
lemma (in Extend) extend_ensures:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   582
     "(extend h F \<in> (extend_set h A) ensures (extend_set h B)) =  
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   583
      (F \<in> A ensures B)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   584
by (simp add: ensures_def extend_constrains extend_transient 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   585
        extend_set_Un_distrib [symmetric] extend_set_Diff_distrib [symmetric])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   586
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   587
lemma (in Extend) leadsTo_imp_extend_leadsTo:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   588
     "F \<in> A leadsTo B  
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   589
      ==> extend h F \<in> (extend_set h A) leadsTo (extend_set h B)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   590
apply (erule leadsTo_induct)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   591
  apply (simp add: leadsTo_Basis extend_ensures)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   592
 apply (blast intro: leadsTo_Trans)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   593
apply (simp add: leadsTo_UN extend_set_Union)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   594
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   595
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13790
diff changeset
   596
subsection{*Proving the converse takes some doing!*}
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   597
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   598
lemma (in Extend) slice_iff [iff]: "(x \<in> slice C y) = (h(x,y) \<in> C)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   599
by (simp (no_asm) add: slice_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   600
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   601
lemma (in Extend) slice_Union: "slice (Union S) y = (\<Union>x \<in> S. slice x y)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   602
by auto
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   603
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   604
lemma (in Extend) slice_extend_set: "slice (extend_set h A) y = A"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   605
by auto
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   606
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   607
lemma (in Extend) project_set_is_UN_slice:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   608
     "project_set h A = (\<Union>y. slice A y)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   609
by auto
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   610
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   611
lemma (in Extend) extend_transient_slice:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   612
     "extend h F \<in> transient A ==> F \<in> transient (slice A y)"
13812
91713a1915ee converting HOL/UNITY to use unconditional fairness
paulson
parents: 13805
diff changeset
   613
by (unfold transient_def, auto)
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   614
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   615
(*Converse?*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   616
lemma (in Extend) extend_constrains_slice:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   617
     "extend h F \<in> A co B ==> F \<in> (slice A y) co (slice B y)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   618
by (auto simp add: constrains_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   619
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   620
lemma (in Extend) extend_ensures_slice:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   621
     "extend h F \<in> A ensures B ==> F \<in> (slice A y) ensures (project_set h B)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   622
apply (auto simp add: ensures_def extend_constrains extend_transient)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   623
apply (erule_tac [2] extend_transient_slice [THEN transient_strengthen])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   624
apply (erule extend_constrains_slice [THEN constrains_weaken], auto)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   625
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   626
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   627
lemma (in Extend) leadsTo_slice_project_set:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   628
     "\<forall>y. F \<in> (slice B y) leadsTo CU ==> F \<in> (project_set h B) leadsTo CU"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   629
apply (simp (no_asm) add: project_set_is_UN_slice)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   630
apply (blast intro: leadsTo_UN)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   631
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   632
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13790
diff changeset
   633
lemma (in Extend) extend_leadsTo_slice [rule_format]:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   634
     "extend h F \<in> AU leadsTo BU  
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   635
      ==> \<forall>y. F \<in> (slice AU y) leadsTo (project_set h BU)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   636
apply (erule leadsTo_induct)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   637
  apply (blast intro: extend_ensures_slice leadsTo_Basis)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   638
 apply (blast intro: leadsTo_slice_project_set leadsTo_Trans)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   639
apply (simp add: leadsTo_UN slice_Union)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   640
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   641
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   642
lemma (in Extend) extend_leadsTo:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   643
     "(extend h F \<in> (extend_set h A) leadsTo (extend_set h B)) =  
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   644
      (F \<in> A leadsTo B)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   645
apply safe
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   646
apply (erule_tac [2] leadsTo_imp_extend_leadsTo)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   647
apply (drule extend_leadsTo_slice)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   648
apply (simp add: slice_extend_set)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   649
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   650
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   651
lemma (in Extend) extend_LeadsTo:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   652
     "(extend h F \<in> (extend_set h A) LeadsTo (extend_set h B)) =   
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   653
      (F \<in> A LeadsTo B)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   654
by (simp add: LeadsTo_def reachable_extend_eq extend_leadsTo
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   655
              extend_set_Int_distrib [symmetric])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   656
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   657
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13790
diff changeset
   658
subsection{*preserves*}
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   659
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   660
lemma (in Extend) project_preserves_I:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   661
     "G \<in> preserves (v o f) ==> project h C G \<in> preserves v"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   662
by (auto simp add: preserves_def project_stable_I extend_set_eq_Collect)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   663
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   664
(*to preserve f is to preserve the whole original state*)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   665
lemma (in Extend) project_preserves_id_I:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   666
     "G \<in> preserves f ==> project h C G \<in> preserves id"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   667
by (simp add: project_preserves_I)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   668
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   669
lemma (in Extend) extend_preserves:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   670
     "(extend h G \<in> preserves (v o f)) = (G \<in> preserves v)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   671
by (auto simp add: preserves_def extend_stable [symmetric] 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   672
                   extend_set_eq_Collect)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   673
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   674
lemma (in Extend) inj_extend_preserves: "inj h ==> (extend h G \<in> preserves g)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   675
by (auto simp add: preserves_def extend_def extend_act_def stable_def 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   676
                   constrains_def g_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   677
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   678
13798
4c1a53627500 conversion to new-style theories and tidying
paulson
parents: 13790
diff changeset
   679
subsection{*Guarantees*}
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   680
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   681
lemma (in Extend) project_extend_Join:
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   682
     "project h UNIV ((extend h F)\<squnion>G) = F\<squnion>(project h UNIV G)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   683
apply (rule program_equalityI)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   684
  apply (simp add: project_set_extend_set_Int)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   685
 apply (simp add: image_eq_UN UN_Un, auto)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   686
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   687
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   688
lemma (in Extend) extend_Join_eq_extend_D:
13819
78f5885b76a9 minor revisions
paulson
parents: 13812
diff changeset
   689
     "(extend h F)\<squnion>G = extend h H ==> H = F\<squnion>(project h UNIV G)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   690
apply (drule_tac f = "project h UNIV" in arg_cong)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   691
apply (simp add: project_extend_Join)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   692
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   693
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   694
(** Strong precondition and postcondition; only useful when
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   695
    the old and new state sets are in bijection **)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   696
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   697
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   698
lemma (in Extend) ok_extend_imp_ok_project:
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   699
     "extend h F ok G ==> F ok project h UNIV G"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   700
apply (auto simp add: ok_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   701
apply (drule subsetD)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   702
apply (auto intro!: rev_image_eqI)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   703
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   704
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   705
lemma (in Extend) ok_extend_iff: "(extend h F ok extend h G) = (F ok G)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   706
apply (simp add: ok_def, safe)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   707
apply (force+)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   708
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   709
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   710
lemma (in Extend) OK_extend_iff: "OK I (%i. extend h (F i)) = (OK I F)"
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   711
apply (unfold OK_def, safe)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   712
apply (drule_tac x = i in bspec)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   713
apply (drule_tac [2] x = j in bspec)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   714
apply (force+)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   715
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   716
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   717
lemma (in Extend) guarantees_imp_extend_guarantees:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   718
     "F \<in> X guarantees Y ==>  
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   719
      extend h F \<in> (extend h ` X) guarantees (extend h ` Y)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   720
apply (rule guaranteesI, clarify)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   721
apply (blast dest: ok_extend_imp_ok_project extend_Join_eq_extend_D 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   722
                   guaranteesD)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   723
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   724
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   725
lemma (in Extend) extend_guarantees_imp_guarantees:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   726
     "extend h F \<in> (extend h ` X) guarantees (extend h ` Y)  
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   727
      ==> F \<in> X guarantees Y"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   728
apply (auto simp add: guar_def)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   729
apply (drule_tac x = "extend h G" in spec)
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   730
apply (simp del: extend_Join 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   731
            add: extend_Join [symmetric] ok_extend_iff 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   732
                 inj_extend [THEN inj_image_mem_iff])
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   733
done
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   734
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   735
lemma (in Extend) extend_guarantees_eq:
13805
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   736
     "(extend h F \<in> (extend h ` X) guarantees (extend h ` Y)) =  
3786b2fd6808 some x-symbols
paulson
parents: 13798
diff changeset
   737
      (F \<in> X guarantees Y)"
13790
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   738
by (blast intro: guarantees_imp_extend_guarantees 
8d7e9fce8c50 converting UNITY to new-style theories
paulson
parents: 10834
diff changeset
   739
                 extend_guarantees_imp_guarantees)
6297
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
   740
5b9fbdfe22b7 new theory of extending the state space
paulson
parents:
diff changeset
   741
end