| author | paulson | 
| Wed, 06 Dec 2000 10:24:44 +0100 | |
| changeset 10600 | 322475c2cb75 | 
| parent 71 | 729fe026c5f3 | 
| permissions | -rw-r--r-- | 
| 0 | 1  | 
(* Title: ZF/ex/term.ML  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
|
4  | 
Copyright 1993 University of Cambridge  | 
|
5  | 
||
6  | 
Datatype definition of terms over an alphabet.  | 
|
7  | 
Illustrates the list functor (essentially the same type as in Trees & Forests)  | 
|
8  | 
*)  | 
|
9  | 
||
10  | 
structure Term = Datatype_Fun  | 
|
11  | 
(val thy = List.thy;  | 
|
12  | 
val rec_specs =  | 
|
13  | 
      [("term", "univ(A)",
 | 
|
14  | 
[(["Apply"], "[i,i]=>i")])];  | 
|
15  | 
val rec_styp = "i=>i";  | 
|
16  | 
val ext = None  | 
|
17  | 
val sintrs = ["[| a: A; l: list(term(A)) |] ==> Apply(a,l) : term(A)"];  | 
|
18  | 
val monos = [list_mono];  | 
|
| 
71
 
729fe026c5f3
sample datatype defs now use datatype_intrs, datatype_elims
 
lcp 
parents: 
56 
diff
changeset
 | 
19  | 
val type_intrs = datatype_intrs;  | 
| 56 | 20  | 
val type_elims = [make_elim (list_univ RS subsetD)]);  | 
| 0 | 21  | 
|
22  | 
val [ApplyI] = Term.intrs;  | 
|
23  | 
||
24  | 
(*Induction on term(A) followed by induction on List *)  | 
|
25  | 
val major::prems = goal Term.thy  | 
|
26  | 
"[| t: term(A); \  | 
|
27  | 
\ !!x. [| x: A |] ==> P(Apply(x,Nil)); \  | 
|
28  | 
\ !!x z zs. [| x: A; z: term(A); zs: list(term(A)); P(Apply(x,zs)) \  | 
|
29  | 
\ |] ==> P(Apply(x, Cons(z,zs))) \  | 
|
30  | 
\ |] ==> P(t)";  | 
|
31  | 
by (rtac (major RS Term.induct) 1);  | 
|
32  | 
by (etac List.induct 1);  | 
|
33  | 
by (etac CollectE 2);  | 
|
34  | 
by (REPEAT (ares_tac (prems@[list_CollectD]) 1));  | 
|
35  | 
val term_induct2 = result();  | 
|
36  | 
||
37  | 
(*Induction on term(A) to prove an equation*)  | 
|
38  | 
val major::prems = goal (merge_theories(Term.thy,ListFn.thy))  | 
|
39  | 
"[| t: term(A); \  | 
|
40  | 
\ !!x zs. [| x: A; zs: list(term(A)); map(f,zs) = map(g,zs) |] ==> \  | 
|
41  | 
\ f(Apply(x,zs)) = g(Apply(x,zs)) \  | 
|
42  | 
\ |] ==> f(t)=g(t)";  | 
|
43  | 
by (rtac (major RS Term.induct) 1);  | 
|
44  | 
by (resolve_tac prems 1);  | 
|
45  | 
by (REPEAT (eresolve_tac [asm_rl, map_list_Collect, list_CollectD] 1));  | 
|
46  | 
val term_induct_eqn = result();  | 
|
47  | 
||
48  | 
(** Lemmas to justify using "term" in other recursive type definitions **)  | 
|
49  | 
||
50  | 
goalw Term.thy Term.defs "!!A B. A<=B ==> term(A) <= term(B)";  | 
|
51  | 
by (rtac lfp_mono 1);  | 
|
52  | 
by (REPEAT (rtac Term.bnd_mono 1));  | 
|
53  | 
by (REPEAT (ares_tac (univ_mono::basic_monos) 1));  | 
|
54  | 
val term_mono = result();  | 
|
55  | 
||
56  | 
(*Easily provable by induction also*)  | 
|
57  | 
goalw Term.thy (Term.defs@Term.con_defs) "term(univ(A)) <= univ(A)";  | 
|
58  | 
by (rtac lfp_lowerbound 1);  | 
|
59  | 
by (rtac (A_subset_univ RS univ_mono) 2);  | 
|
60  | 
by (safe_tac ZF_cs);  | 
|
61  | 
by (REPEAT (ares_tac [Pair_in_univ, list_univ RS subsetD] 1));  | 
|
62  | 
val term_univ = result();  | 
|
63  | 
||
64  | 
val term_subset_univ = standard  | 
|
65  | 
(term_mono RS (term_univ RSN (2,subset_trans)));  | 
|
66  |