author | obua |
Thu, 16 Feb 2006 04:17:19 +0100 | |
changeset 19067 | c0321d7d6b3d |
parent 16417 | 9bc16273c2d4 |
child 24893 | b8ef7afe3a6b |
permissions | -rw-r--r-- |
11479 | 1 |
(* Title: HOL/UNITY/State.thy |
2 |
ID: $Id$ |
|
3 |
Author: Sidi O Ehmety, Computer Laboratory |
|
4 |
Copyright 2001 University of Cambridge |
|
5 |
||
12195 | 6 |
Formalizes UNITY-program states using dependent types so that: |
11479 | 7 |
- variables are typed. |
8 |
- the state space is uniform, common to all defined programs. |
|
9 |
- variables can be quantified over. |
|
10 |
*) |
|
11 |
||
14084
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
12 |
header{*UNITY Program States*} |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
13 |
|
16417 | 14 |
theory State imports Main begin |
11479 | 15 |
|
14084
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
16 |
consts var :: i |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
17 |
datatype var = Var("i \<in> list(nat)") |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
18 |
type_intros nat_subset_univ [THEN list_subset_univ, THEN subsetD] |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
19 |
|
11479 | 20 |
consts |
14084
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
21 |
type_of :: "i=>i" |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
22 |
default_val :: "i=>i" |
11479 | 23 |
|
14046 | 24 |
constdefs |
25 |
state :: i |
|
14171
0cab06e3bbd0
Extended the notion of letter and digit, such that now one may use greek,
skalberg
parents:
14084
diff
changeset
|
26 |
"state == \<Pi> x \<in> var. cons(default_val(x), type_of(x))" |
14084
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
27 |
|
14046 | 28 |
st0 :: i |
14084
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
29 |
"st0 == \<lambda>x \<in> var. default_val(x)" |
14046 | 30 |
|
14084
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
31 |
st_set :: "i=>o" |
14046 | 32 |
(* To prevent typing conditions like `A<=state' from |
33 |
being used in combination with the rules `constrains_weaken', etc. *) |
|
34 |
"st_set(A) == A<=state" |
|
11479 | 35 |
|
14084
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
36 |
st_compl :: "i=>i" |
14046 | 37 |
"st_compl(A) == state-A" |
14084
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
38 |
|
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
39 |
|
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
40 |
lemma st0_in_state [simp,TC]: "st0 \<in> state" |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
41 |
by (simp add: state_def st0_def) |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
42 |
|
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
43 |
lemma st_set_Collect [iff]: "st_set({x \<in> state. P(x)})" |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
44 |
by (simp add: st_set_def, auto) |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
45 |
|
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
46 |
lemma st_set_0 [iff]: "st_set(0)" |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
47 |
by (simp add: st_set_def) |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
48 |
|
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
49 |
lemma st_set_state [iff]: "st_set(state)" |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
50 |
by (simp add: st_set_def) |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
51 |
|
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
52 |
(* Union *) |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
53 |
|
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
54 |
lemma st_set_Un_iff [iff]: "st_set(A Un B) <-> st_set(A) & st_set(B)" |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
55 |
by (simp add: st_set_def, auto) |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
56 |
|
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
57 |
lemma st_set_Union_iff [iff]: "st_set(Union(S)) <-> (\<forall>A \<in> S. st_set(A))" |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
58 |
by (simp add: st_set_def, auto) |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
59 |
|
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
60 |
(* Intersection *) |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
61 |
|
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
62 |
lemma st_set_Int [intro!]: "st_set(A) | st_set(B) ==> st_set(A Int B)" |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
63 |
by (simp add: st_set_def, auto) |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
64 |
|
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
65 |
lemma st_set_Inter [intro!]: |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
66 |
"(S=0) | (\<exists>A \<in> S. st_set(A)) ==> st_set(Inter(S))" |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
67 |
apply (simp add: st_set_def Inter_def, auto) |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
68 |
done |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
69 |
|
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
70 |
(* Diff *) |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
71 |
lemma st_set_DiffI [intro!]: "st_set(A) ==> st_set(A - B)" |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
72 |
by (simp add: st_set_def, auto) |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
73 |
|
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
74 |
lemma Collect_Int_state [simp]: "Collect(state,P) Int state = Collect(state,P)" |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
75 |
by auto |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
76 |
|
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
77 |
lemma state_Int_Collect [simp]: "state Int Collect(state,P) = Collect(state,P)" |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
78 |
by auto |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
79 |
|
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
80 |
|
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
81 |
(* Introduction and destruction rules for st_set *) |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
82 |
|
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
83 |
lemma st_setI: "A <= state ==> st_set(A)" |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
84 |
by (simp add: st_set_def) |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
85 |
|
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
86 |
lemma st_setD: "st_set(A) ==> A<=state" |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
87 |
by (simp add: st_set_def) |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
88 |
|
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
89 |
lemma st_set_subset: "[| st_set(A); B<=A |] ==> st_set(B)" |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
90 |
by (simp add: st_set_def, auto) |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
91 |
|
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
92 |
|
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
93 |
lemma state_update_type: |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
94 |
"[| s \<in> state; x \<in> var; y \<in> type_of(x) |] ==> s(x:=y):state" |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
95 |
apply (simp add: state_def) |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
96 |
apply (blast intro: update_type) |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
97 |
done |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
98 |
|
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
99 |
lemma st_set_compl [simp]: "st_set(st_compl(A))" |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
100 |
by (simp add: st_compl_def, auto) |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
101 |
|
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
102 |
lemma st_compl_iff [simp]: "x \<in> st_compl(A) <-> x \<in> state & x \<notin> A" |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
103 |
by (simp add: st_compl_def) |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
104 |
|
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
105 |
lemma st_compl_Collect [simp]: |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
106 |
"st_compl({s \<in> state. P(s)}) = {s \<in> state. ~P(s)}" |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
107 |
by (simp add: st_compl_def, auto) |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
108 |
|
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
109 |
(*For using "disjunction" (union over an index set) to eliminate a variable.*) |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
110 |
lemma UN_conj_eq: |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
111 |
"\<forall>d\<in>D. f(d) \<in> A ==> (\<Union>k\<in>A. {d\<in>D. P(d) & f(d) = k}) = {d\<in>D. P(d)}" |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
112 |
by blast |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
113 |
|
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
114 |
|
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
115 |
ML |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
116 |
{* |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
117 |
val st_set_def = thm "st_set_def"; |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
118 |
val state_def = thm "state_def"; |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
119 |
|
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
120 |
val st0_in_state = thm "st0_in_state"; |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
121 |
val st_set_Collect = thm "st_set_Collect"; |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
122 |
val st_set_0 = thm "st_set_0"; |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
123 |
val st_set_state = thm "st_set_state"; |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
124 |
val st_set_Un_iff = thm "st_set_Un_iff"; |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
125 |
val st_set_Union_iff = thm "st_set_Union_iff"; |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
126 |
val st_set_Int = thm "st_set_Int"; |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
127 |
val st_set_Inter = thm "st_set_Inter"; |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
128 |
val st_set_DiffI = thm "st_set_DiffI"; |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
129 |
val Collect_Int_state = thm "Collect_Int_state"; |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
130 |
val state_Int_Collect = thm "state_Int_Collect"; |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
131 |
val st_setI = thm "st_setI"; |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
132 |
val st_setD = thm "st_setD"; |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
133 |
val st_set_subset = thm "st_set_subset"; |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
134 |
val state_update_type = thm "state_update_type"; |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
135 |
val st_set_compl = thm "st_set_compl"; |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
136 |
val st_compl_iff = thm "st_compl_iff"; |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
137 |
val st_compl_Collect = thm "st_compl_Collect"; |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
138 |
*} |
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
139 |
|
ccb48f3239f7
Removal of UNITY/UNITYMisc, moving its theorems elsewhere.
paulson
parents:
14046
diff
changeset
|
140 |
end |