| author | blanchet | 
| Sat, 21 Apr 2012 11:15:49 +0200 | |
| changeset 47643 | e33c2be488fe | 
| parent 46822 | 95f1e700b712 | 
| child 76213 | e44d86131648 | 
| permissions | -rw-r--r-- | 
| 1478 | 1  | 
(* Title: ZF/Coind/Types.thy  | 
2  | 
Author: Jacob Frost, Cambridge University Computer Laboratory  | 
|
| 916 | 3  | 
Copyright 1995 University of Cambridge  | 
4  | 
*)  | 
|
5  | 
||
| 16417 | 6  | 
theory Types imports Language begin  | 
| 916 | 7  | 
|
8  | 
consts  | 
|
| 12595 | 9  | 
Ty :: i (* Datatype of types *)  | 
| 1478 | 10  | 
TyConst :: i (* Abstract type of type constants *)  | 
| 6112 | 11  | 
|
12  | 
datatype  | 
|
| 12595 | 13  | 
  "Ty" = t_const ("tc \<in> TyConst")
 | 
14  | 
       | t_fun ("t1 \<in> Ty","t2 \<in> Ty")
 | 
|
| 916 | 15  | 
|
16  | 
||
17  | 
(* Definition of type environments and associated operators *)  | 
|
18  | 
||
19  | 
consts  | 
|
| 1401 | 20  | 
TyEnv :: i  | 
| 6112 | 21  | 
|
22  | 
datatype  | 
|
| 934 | 23  | 
"TyEnv" = te_emp  | 
| 12595 | 24  | 
          | te_owr ("te \<in> TyEnv","x \<in> ExVar","t \<in> Ty") 
 | 
| 916 | 25  | 
|
26  | 
consts  | 
|
| 12595 | 27  | 
te_dom :: "i => i"  | 
28  | 
te_app :: "[i,i] => i"  | 
|
| 6046 | 29  | 
|
30  | 
||
31  | 
primrec (*domain of the type environment*)  | 
|
32  | 
"te_dom (te_emp) = 0"  | 
|
| 
46822
 
95f1e700b712
 mathematical symbols for Isabelle/ZF example theories
 
paulson 
parents: 
35762 
diff
changeset
 | 
33  | 
  "te_dom (te_owr(te,x,v)) = te_dom(te) \<union> {x}"
 | 
| 6046 | 34  | 
|
35  | 
primrec (*lookup up identifiers in the type environment*)  | 
|
36  | 
"te_app (te_emp,x) = 0"  | 
|
| 6068 | 37  | 
"te_app (te_owr(te,y,t),x) = (if x=y then t else te_app(te,x))"  | 
| 916 | 38  | 
|
| 12595 | 39  | 
inductive_cases te_owrE [elim!]: "te_owr(te,f,t) \<in> TyEnv"  | 
40  | 
||
41  | 
(*redundant??*)  | 
|
42  | 
lemma te_app_owr1: "te_app(te_owr(te,x,t),x) = t"  | 
|
43  | 
by simp  | 
|
44  | 
||
45  | 
(*redundant??*)  | 
|
46  | 
lemma te_app_owr2: "x \<noteq> y ==> te_app(te_owr(te,x,t),y) = te_app(te,y)"  | 
|
47  | 
by auto  | 
|
48  | 
||
49  | 
lemma te_app_owr [simp]:  | 
|
50  | 
"te_app(te_owr(te,x,t),y) = (if x=y then t else te_app(te,y))"  | 
|
51  | 
by auto  | 
|
52  | 
||
53  | 
lemma te_appI:  | 
|
54  | 
"[| te \<in> TyEnv; x \<in> ExVar; x \<in> te_dom(te) |] ==> te_app(te,x) \<in> Ty"  | 
|
55  | 
apply (erule_tac P = "x \<in> te_dom (te) " in rev_mp)  | 
|
| 
13339
 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 
paulson 
parents: 
12595 
diff
changeset
 | 
56  | 
apply (erule TyEnv.induct, auto)  | 
| 12595 | 57  | 
done  | 
| 916 | 58  | 
|
59  | 
||
60  | 
||
61  | 
||
62  | 
||
| 12595 | 63  | 
|
64  | 
||
65  | 
||
66  | 
||
67  | 
||
68  | 
||
69  | 
||
70  | 
||
71  | 
||
72  | 
||
73  | 
||
74  | 
||
75  | 
||
76  | 
end  |