| author | noschinl | 
| Fri, 01 Jul 2011 13:54:23 +0200 | |
| changeset 43615 | 8e0f6cfa8eb2 | 
| parent 32960 | 69916a850301 | 
| child 46823 | 57bf0cecb366 | 
| permissions | -rw-r--r-- | 
| 9284 | 1  | 
(* Title: ZF/Resid/Redex.thy  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
22808 
diff
changeset
 | 
2  | 
Author: Ole Rasmussen, University of Cambridge  | 
| 1048 | 3  | 
*)  | 
4  | 
||
| 16417 | 5  | 
theory Redex imports Main begin  | 
| 1048 | 6  | 
consts  | 
| 1401 | 7  | 
redexes :: i  | 
| 1048 | 8  | 
|
9  | 
datatype  | 
|
| 12593 | 10  | 
  "redexes" = Var ("n \<in> nat")            
 | 
11  | 
            | Fun ("t \<in> redexes")
 | 
|
| 
13339
 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 
paulson 
parents: 
12610 
diff
changeset
 | 
12  | 
            | App ("b \<in> bool","f \<in> redexes", "a \<in> redexes")
 | 
| 1048 | 13  | 
|
| 6046 | 14  | 
|
| 1048 | 15  | 
consts  | 
| 12593 | 16  | 
Ssub :: "i"  | 
17  | 
Scomp :: "i"  | 
|
18  | 
Sreg :: "i"  | 
|
| 22808 | 19  | 
|
20  | 
abbreviation  | 
|
21  | 
Ssub_rel (infixl "<==" 70) where  | 
|
22  | 
"a<==b == <a,b> \<in> Ssub"  | 
|
| 12593 | 23  | 
|
| 22808 | 24  | 
abbreviation  | 
25  | 
Scomp_rel (infixl "~" 70) where  | 
|
26  | 
"a ~ b == <a,b> \<in> Scomp"  | 
|
| 12593 | 27  | 
|
| 22808 | 28  | 
abbreviation  | 
29  | 
"regular(a) == a \<in> Sreg"  | 
|
| 9284 | 30  | 
|
| 22808 | 31  | 
consts union_aux :: "i=>i"  | 
| 6046 | 32  | 
primrec (*explicit lambda is required because both arguments of "un" vary*)  | 
33  | 
"union_aux(Var(n)) =  | 
|
| 12593 | 34  | 
(\<lambda>t \<in> redexes. redexes_case(%j. Var(n), %x. 0, %b x y.0, t))"  | 
| 6046 | 35  | 
|
36  | 
"union_aux(Fun(u)) =  | 
|
| 12593 | 37  | 
(\<lambda>t \<in> redexes. redexes_case(%j. 0, %y. Fun(union_aux(u)`y),  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
22808 
diff
changeset
 | 
38  | 
%b y z. 0, t))"  | 
| 6046 | 39  | 
|
40  | 
"union_aux(App(b,f,a)) =  | 
|
| 12593 | 41  | 
(\<lambda>t \<in> redexes.  | 
| 6046 | 42  | 
redexes_case(%j. 0, %y. 0,  | 
| 
32960
 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 
wenzelm 
parents: 
22808 
diff
changeset
 | 
43  | 
%c z u. App(b or c, union_aux(f)`z, union_aux(a)`u), t))"  | 
| 6046 | 44  | 
|
| 22808 | 45  | 
definition  | 
46  | 
union (infixl "un" 70) where  | 
|
47  | 
"u un v == union_aux(u)`v"  | 
|
| 6046 | 48  | 
|
| 22808 | 49  | 
notation (xsymbols)  | 
50  | 
union (infixl "\<squnion>" 70) and  | 
|
51  | 
Ssub_rel (infixl "\<Longleftarrow>" 70) and  | 
|
52  | 
Scomp_rel (infixl "\<sim>" 70)  | 
|
53  | 
||
| 6046 | 54  | 
|
55  | 
inductive  | 
|
56  | 
domains "Ssub" <= "redexes*redexes"  | 
|
| 12593 | 57  | 
intros  | 
58  | 
Sub_Var: "n \<in> nat ==> Var(n)<== Var(n)"  | 
|
59  | 
Sub_Fun: "[|u<== v|]==> Fun(u)<== Fun(v)"  | 
|
60  | 
Sub_App1: "[|u1<== v1; u2<== v2; b \<in> bool|]==>  | 
|
| 6046 | 61  | 
App(0,u1,u2)<== App(b,v1,v2)"  | 
| 12593 | 62  | 
Sub_App2: "[|u1<== v1; u2<== v2|]==> App(1,u1,u2)<== App(1,v1,v2)"  | 
63  | 
type_intros redexes.intros bool_typechecks  | 
|
| 6046 | 64  | 
|
65  | 
inductive  | 
|
66  | 
domains "Scomp" <= "redexes*redexes"  | 
|
| 12593 | 67  | 
intros  | 
68  | 
Comp_Var: "n \<in> nat ==> Var(n) ~ Var(n)"  | 
|
69  | 
Comp_Fun: "[|u ~ v|]==> Fun(u) ~ Fun(v)"  | 
|
70  | 
Comp_App: "[|u1 ~ v1; u2 ~ v2; b1 \<in> bool; b2 \<in> bool|]  | 
|
71  | 
==> App(b1,u1,u2) ~ App(b2,v1,v2)"  | 
|
72  | 
type_intros redexes.intros bool_typechecks  | 
|
| 6046 | 73  | 
|
74  | 
inductive  | 
|
| 
13339
 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 
paulson 
parents: 
12610 
diff
changeset
 | 
75  | 
domains "Sreg" <= redexes  | 
| 12593 | 76  | 
intros  | 
77  | 
Reg_Var: "n \<in> nat ==> regular(Var(n))"  | 
|
78  | 
Reg_Fun: "[|regular(u)|]==> regular(Fun(u))"  | 
|
79  | 
Reg_App1: "[|regular(Fun(u)); regular(v) |]==>regular(App(1,Fun(u),v))"  | 
|
80  | 
Reg_App2: "[|regular(u); regular(v) |]==>regular(App(0,u,v))"  | 
|
81  | 
type_intros redexes.intros bool_typechecks  | 
|
82  | 
||
83  | 
||
84  | 
declare redexes.intros [simp]  | 
|
85  | 
||
86  | 
(* ------------------------------------------------------------------------- *)  | 
|
87  | 
(* Specialisation of comp-rules *)  | 
|
88  | 
(* ------------------------------------------------------------------------- *)  | 
|
89  | 
||
90  | 
lemmas compD1 [simp] = Scomp.dom_subset [THEN subsetD, THEN SigmaD1]  | 
|
91  | 
lemmas compD2 [simp] = Scomp.dom_subset [THEN subsetD, THEN SigmaD2]  | 
|
92  | 
lemmas regD [simp] = Sreg.dom_subset [THEN subsetD]  | 
|
93  | 
||
94  | 
(* ------------------------------------------------------------------------- *)  | 
|
95  | 
(* Equality rules for union *)  | 
|
96  | 
(* ------------------------------------------------------------------------- *)  | 
|
97  | 
||
98  | 
lemma union_Var [simp]: "n \<in> nat ==> Var(n) un Var(n)=Var(n)"  | 
|
99  | 
by (simp add: union_def)  | 
|
100  | 
||
101  | 
lemma union_Fun [simp]: "v \<in> redexes ==> Fun(u) un Fun(v) = Fun(u un v)"  | 
|
102  | 
by (simp add: union_def)  | 
|
103  | 
||
104  | 
lemma union_App [simp]:  | 
|
105  | 
"[|b2 \<in> bool; u2 \<in> redexes; v2 \<in> redexes|]  | 
|
106  | 
==> App(b1,u1,v1) un App(b2,u2,v2)=App(b1 or b2,u1 un u2,v1 un v2)"  | 
|
107  | 
by (simp add: union_def)  | 
|
108  | 
||
109  | 
||
110  | 
lemma or_1_right [simp]: "a or 1 = 1"  | 
|
111  | 
by (simp add: or_def cond_def)  | 
|
112  | 
||
113  | 
lemma or_0_right [simp]: "a \<in> bool \<Longrightarrow> a or 0 = a"  | 
|
114  | 
by (simp add: or_def cond_def bool_def, auto)  | 
|
115  | 
||
116  | 
||
117  | 
||
118  | 
declare Ssub.intros [simp]  | 
|
119  | 
declare bool_typechecks [simp]  | 
|
120  | 
declare Sreg.intros [simp]  | 
|
121  | 
declare Scomp.intros [simp]  | 
|
122  | 
||
123  | 
declare Scomp.intros [intro]  | 
|
| 6046 | 124  | 
|
| 12610 | 125  | 
inductive_cases [elim!]:  | 
126  | 
"regular(App(b,f,a))"  | 
|
127  | 
"regular(Fun(b))"  | 
|
128  | 
"regular(Var(b))"  | 
|
129  | 
"Fun(u) ~ Fun(t)"  | 
|
130  | 
"u ~ Fun(t)"  | 
|
131  | 
"u ~ Var(n)"  | 
|
132  | 
"u ~ App(b,t,a)"  | 
|
133  | 
"Fun(t) ~ v"  | 
|
134  | 
"App(b,f,a) ~ v"  | 
|
135  | 
"Var(n) ~ u"  | 
|
| 12593 | 136  | 
|
137  | 
||
138  | 
||
139  | 
(* ------------------------------------------------------------------------- *)  | 
|
140  | 
(* comp proofs *)  | 
|
141  | 
(* ------------------------------------------------------------------------- *)  | 
|
142  | 
||
143  | 
lemma comp_refl [simp]: "u \<in> redexes ==> u ~ u"  | 
|
144  | 
by (erule redexes.induct, blast+)  | 
|
145  | 
||
146  | 
lemma comp_sym: "u ~ v ==> v ~ u"  | 
|
147  | 
by (erule Scomp.induct, blast+)  | 
|
148  | 
||
149  | 
lemma comp_sym_iff: "u ~ v <-> v ~ u"  | 
|
150  | 
by (blast intro: comp_sym)  | 
|
151  | 
||
152  | 
lemma comp_trans [rule_format]: "u ~ v ==> \<forall>w. v ~ w-->u ~ w"  | 
|
153  | 
by (erule Scomp.induct, blast+)  | 
|
154  | 
||
155  | 
(* ------------------------------------------------------------------------- *)  | 
|
156  | 
(* union proofs *)  | 
|
157  | 
(* ------------------------------------------------------------------------- *)  | 
|
158  | 
||
159  | 
lemma union_l: "u ~ v ==> u <== (u un v)"  | 
|
160  | 
apply (erule Scomp.induct)  | 
|
| 
13339
 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 
paulson 
parents: 
12610 
diff
changeset
 | 
161  | 
apply (erule_tac [3] boolE, simp_all)  | 
| 12593 | 162  | 
done  | 
163  | 
||
164  | 
lemma union_r: "u ~ v ==> v <== (u un v)"  | 
|
165  | 
apply (erule Scomp.induct)  | 
|
| 
13339
 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 
paulson 
parents: 
12610 
diff
changeset
 | 
166  | 
apply (erule_tac [3] c = b2 in boolE, simp_all)  | 
| 12593 | 167  | 
done  | 
168  | 
||
169  | 
lemma union_sym: "u ~ v ==> u un v = v un u"  | 
|
170  | 
by (erule Scomp.induct, simp_all add: or_commute)  | 
|
171  | 
||
172  | 
(* ------------------------------------------------------------------------- *)  | 
|
173  | 
(* regular proofs *)  | 
|
174  | 
(* ------------------------------------------------------------------------- *)  | 
|
175  | 
||
176  | 
lemma union_preserve_regular [rule_format]:  | 
|
177  | 
"u ~ v ==> regular(u)-->regular(v)-->regular(u un v)"  | 
|
| 13612 | 178  | 
by (erule Scomp.induct, auto)  | 
| 6046 | 179  | 
|
| 1048 | 180  | 
end  | 
181  |