src/ZF/Coind/Types.thy
author haftmann
Fri, 17 Jun 2005 16:12:49 +0200
changeset 16417 9bc16273c2d4
parent 13339 0f89104dd377
child 35762 af3ff2ba4c54
permissions -rw-r--r--
migrated theory headers to new format
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     1
(*  Title:      ZF/Coind/Types.thy
916
d03bb9f50b3b New example by Jacob Frost, tidied by lcp
lcp
parents:
diff changeset
     2
    ID:         $Id$
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
     3
    Author:     Jacob Frost, Cambridge University Computer Laboratory
916
d03bb9f50b3b New example by Jacob Frost, tidied by lcp
lcp
parents:
diff changeset
     4
    Copyright   1995  University of Cambridge
d03bb9f50b3b New example by Jacob Frost, tidied by lcp
lcp
parents:
diff changeset
     5
*)
d03bb9f50b3b New example by Jacob Frost, tidied by lcp
lcp
parents:
diff changeset
     6
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 13339
diff changeset
     7
theory Types imports Language begin
916
d03bb9f50b3b New example by Jacob Frost, tidied by lcp
lcp
parents:
diff changeset
     8
d03bb9f50b3b New example by Jacob Frost, tidied by lcp
lcp
parents:
diff changeset
     9
consts
12595
0480d02221b8 conversion to Isar
paulson
parents: 11318
diff changeset
    10
  Ty :: i               (* Datatype of types *)
1478
2b8c2a7547ab expanded tabs
clasohm
parents: 1401
diff changeset
    11
  TyConst :: i          (* Abstract type of type constants *)
6112
5e4871c5136b datatype package improvements
paulson
parents: 6068
diff changeset
    12
5e4871c5136b datatype package improvements
paulson
parents: 6068
diff changeset
    13
datatype
12595
0480d02221b8 conversion to Isar
paulson
parents: 11318
diff changeset
    14
  "Ty" = t_const ("tc \<in> TyConst")
0480d02221b8 conversion to Isar
paulson
parents: 11318
diff changeset
    15
       | t_fun ("t1 \<in> Ty","t2 \<in> Ty")
916
d03bb9f50b3b New example by Jacob Frost, tidied by lcp
lcp
parents:
diff changeset
    16
  
d03bb9f50b3b New example by Jacob Frost, tidied by lcp
lcp
parents:
diff changeset
    17
d03bb9f50b3b New example by Jacob Frost, tidied by lcp
lcp
parents:
diff changeset
    18
(* Definition of type environments and associated operators *)
d03bb9f50b3b New example by Jacob Frost, tidied by lcp
lcp
parents:
diff changeset
    19
d03bb9f50b3b New example by Jacob Frost, tidied by lcp
lcp
parents:
diff changeset
    20
consts
1401
0c439768f45c removed quotes from consts and syntax sections
clasohm
parents: 1155
diff changeset
    21
  TyEnv :: i
6112
5e4871c5136b datatype package improvements
paulson
parents: 6068
diff changeset
    22
5e4871c5136b datatype package improvements
paulson
parents: 6068
diff changeset
    23
datatype
934
2e0203309287 Replaced rules by defs. Also got rid of tyconstU by
lcp
parents: 916
diff changeset
    24
  "TyEnv" = te_emp
12595
0480d02221b8 conversion to Isar
paulson
parents: 11318
diff changeset
    25
          | te_owr ("te \<in> TyEnv","x \<in> ExVar","t \<in> Ty") 
916
d03bb9f50b3b New example by Jacob Frost, tidied by lcp
lcp
parents:
diff changeset
    26
d03bb9f50b3b New example by Jacob Frost, tidied by lcp
lcp
parents:
diff changeset
    27
consts
12595
0480d02221b8 conversion to Isar
paulson
parents: 11318
diff changeset
    28
  te_dom :: "i => i"
0480d02221b8 conversion to Isar
paulson
parents: 11318
diff changeset
    29
  te_app :: "[i,i] => i"
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    30
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    31
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    32
primrec (*domain of the type environment*)
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    33
  "te_dom (te_emp) = 0"
12595
0480d02221b8 conversion to Isar
paulson
parents: 11318
diff changeset
    34
  "te_dom (te_owr(te,x,v)) = te_dom(te) Un {x}"
6046
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    35
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    36
primrec (*lookup up identifiers in the type environment*)
2c8a8be36c94 converted to use new primrec section
paulson
parents: 3840
diff changeset
    37
  "te_app (te_emp,x) = 0"
6068
2d8f3e1f1151 if-then-else syntax for ZF
paulson
parents: 6046
diff changeset
    38
  "te_app (te_owr(te,y,t),x) = (if x=y then t else te_app(te,x))"
916
d03bb9f50b3b New example by Jacob Frost, tidied by lcp
lcp
parents:
diff changeset
    39
12595
0480d02221b8 conversion to Isar
paulson
parents: 11318
diff changeset
    40
inductive_cases te_owrE [elim!]: "te_owr(te,f,t) \<in> TyEnv"
0480d02221b8 conversion to Isar
paulson
parents: 11318
diff changeset
    41
0480d02221b8 conversion to Isar
paulson
parents: 11318
diff changeset
    42
(*redundant??*)
0480d02221b8 conversion to Isar
paulson
parents: 11318
diff changeset
    43
lemma te_app_owr1: "te_app(te_owr(te,x,t),x) = t"
0480d02221b8 conversion to Isar
paulson
parents: 11318
diff changeset
    44
by simp
0480d02221b8 conversion to Isar
paulson
parents: 11318
diff changeset
    45
0480d02221b8 conversion to Isar
paulson
parents: 11318
diff changeset
    46
(*redundant??*)
0480d02221b8 conversion to Isar
paulson
parents: 11318
diff changeset
    47
lemma te_app_owr2: "x \<noteq> y ==> te_app(te_owr(te,x,t),y) = te_app(te,y)"
0480d02221b8 conversion to Isar
paulson
parents: 11318
diff changeset
    48
by auto
0480d02221b8 conversion to Isar
paulson
parents: 11318
diff changeset
    49
0480d02221b8 conversion to Isar
paulson
parents: 11318
diff changeset
    50
lemma te_app_owr [simp]:
0480d02221b8 conversion to Isar
paulson
parents: 11318
diff changeset
    51
     "te_app(te_owr(te,x,t),y) = (if x=y then t else te_app(te,y))"
0480d02221b8 conversion to Isar
paulson
parents: 11318
diff changeset
    52
by auto
0480d02221b8 conversion to Isar
paulson
parents: 11318
diff changeset
    53
0480d02221b8 conversion to Isar
paulson
parents: 11318
diff changeset
    54
lemma te_appI:
0480d02221b8 conversion to Isar
paulson
parents: 11318
diff changeset
    55
     "[| te \<in> TyEnv; x \<in> ExVar; x \<in> te_dom(te) |] ==> te_app(te,x) \<in> Ty"
0480d02221b8 conversion to Isar
paulson
parents: 11318
diff changeset
    56
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
    57
apply (erule TyEnv.induct, auto)
12595
0480d02221b8 conversion to Isar
paulson
parents: 11318
diff changeset
    58
done
916
d03bb9f50b3b New example by Jacob Frost, tidied by lcp
lcp
parents:
diff changeset
    59
d03bb9f50b3b New example by Jacob Frost, tidied by lcp
lcp
parents:
diff changeset
    60
d03bb9f50b3b New example by Jacob Frost, tidied by lcp
lcp
parents:
diff changeset
    61
d03bb9f50b3b New example by Jacob Frost, tidied by lcp
lcp
parents:
diff changeset
    62
d03bb9f50b3b New example by Jacob Frost, tidied by lcp
lcp
parents:
diff changeset
    63
12595
0480d02221b8 conversion to Isar
paulson
parents: 11318
diff changeset
    64
0480d02221b8 conversion to Isar
paulson
parents: 11318
diff changeset
    65
0480d02221b8 conversion to Isar
paulson
parents: 11318
diff changeset
    66
0480d02221b8 conversion to Isar
paulson
parents: 11318
diff changeset
    67
0480d02221b8 conversion to Isar
paulson
parents: 11318
diff changeset
    68
0480d02221b8 conversion to Isar
paulson
parents: 11318
diff changeset
    69
0480d02221b8 conversion to Isar
paulson
parents: 11318
diff changeset
    70
0480d02221b8 conversion to Isar
paulson
parents: 11318
diff changeset
    71
0480d02221b8 conversion to Isar
paulson
parents: 11318
diff changeset
    72
0480d02221b8 conversion to Isar
paulson
parents: 11318
diff changeset
    73
0480d02221b8 conversion to Isar
paulson
parents: 11318
diff changeset
    74
0480d02221b8 conversion to Isar
paulson
parents: 11318
diff changeset
    75
0480d02221b8 conversion to Isar
paulson
parents: 11318
diff changeset
    76
0480d02221b8 conversion to Isar
paulson
parents: 11318
diff changeset
    77
end