| author | paulson | 
| Tue, 18 Apr 2000 15:51:59 +0200 | |
| changeset 8736 | 0bfd6678a5fa | 
| parent 496 | 3fc829fa81d2 | 
| permissions | -rw-r--r-- | 
| 
34
 
747f1aad03cf
changed filenames to lower case name of theory the file contains
 
clasohm 
parents:  
diff
changeset
 | 
1  | 
(* Title: ZF/ex/llist_eq.ML  | 
| 
 
747f1aad03cf
changed filenames to lower case name of theory the file contains
 
clasohm 
parents:  
diff
changeset
 | 
2  | 
ID: $Id$  | 
| 
 
747f1aad03cf
changed filenames to lower case name of theory the file contains
 
clasohm 
parents:  
diff
changeset
 | 
3  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
| 
 
747f1aad03cf
changed filenames to lower case name of theory the file contains
 
clasohm 
parents:  
diff
changeset
 | 
4  | 
Copyright 1993 University of Cambridge  | 
| 
 
747f1aad03cf
changed filenames to lower case name of theory the file contains
 
clasohm 
parents:  
diff
changeset
 | 
5  | 
|
| 120 | 6  | 
Equality for llist(A) as a greatest fixed point  | 
7  | 
***)  | 
|
| 
34
 
747f1aad03cf
changed filenames to lower case name of theory the file contains
 
clasohm 
parents:  
diff
changeset
 | 
8  | 
|
| 120 | 9  | 
(*Previously used <*> in the domain and variant pairs as elements. But  | 
10  | 
standard pairs work just as well. To use variant pairs, must change prefix  | 
|
11  | 
a q/Q to the Sigma, Pair and converse rules.*)  | 
|
| 
34
 
747f1aad03cf
changed filenames to lower case name of theory the file contains
 
clasohm 
parents:  
diff
changeset
 | 
12  | 
|
| 120 | 13  | 
structure LList_Eq = CoInductive_Fun  | 
| 496 | 14  | 
 (val thy 	 = LList.thy |> add_consts [("lleq","i=>i",NoSyn)]
 | 
15  | 
val thy_name = "LList_Eq"  | 
|
| 279 | 16  | 
  val rec_doms   = [("lleq", "llist(A) * llist(A)")]
 | 
17  | 
val sintrs =  | 
|
18  | 
["<LNil, LNil> : lleq(A)",  | 
|
19  | 
"[| a:A; <l,l'>: lleq(A) |] ==> <LCons(a,l), LCons(a,l')>: lleq(A)"]  | 
|
20  | 
val monos = []  | 
|
21  | 
val con_defs = []  | 
|
| 496 | 22  | 
val type_intrs = LList.intrs  | 
23  | 
val type_elims = []);  | 
|
| 
34
 
747f1aad03cf
changed filenames to lower case name of theory the file contains
 
clasohm 
parents:  
diff
changeset
 | 
24  | 
|
| 
 
747f1aad03cf
changed filenames to lower case name of theory the file contains
 
clasohm 
parents:  
diff
changeset
 | 
25  | 
(** Alternatives for above:  | 
| 
 
747f1aad03cf
changed filenames to lower case name of theory the file contains
 
clasohm 
parents:  
diff
changeset
 | 
26  | 
val con_defs = LList.con_defs  | 
| 120 | 27  | 
val type_intrs = codatatype_intrs  | 
| 
34
 
747f1aad03cf
changed filenames to lower case name of theory the file contains
 
clasohm 
parents:  
diff
changeset
 | 
28  | 
val type_elims = [quniv_QPair_E]  | 
| 
 
747f1aad03cf
changed filenames to lower case name of theory the file contains
 
clasohm 
parents:  
diff
changeset
 | 
29  | 
**)  | 
| 
 
747f1aad03cf
changed filenames to lower case name of theory the file contains
 
clasohm 
parents:  
diff
changeset
 | 
30  | 
|
| 
 
747f1aad03cf
changed filenames to lower case name of theory the file contains
 
clasohm 
parents:  
diff
changeset
 | 
31  | 
val lleq_cs = subset_cs  | 
| 173 | 32  | 
addSIs [QPair_Int_Vset_subset_UN RS subset_trans, QPair_mono]  | 
33  | 
addSEs [Ord_in_Ord, Pair_inject];  | 
|
| 
34
 
747f1aad03cf
changed filenames to lower case name of theory the file contains
 
clasohm 
parents:  
diff
changeset
 | 
34  | 
|
| 173 | 35  | 
(*Lemma for proving finality. Unfold the lazy list; use induction hypothesis*)  | 
| 
34
 
747f1aad03cf
changed filenames to lower case name of theory the file contains
 
clasohm 
parents:  
diff
changeset
 | 
36  | 
goal LList_Eq.thy  | 
| 120 | 37  | 
"!!i. Ord(i) ==> ALL l l'. <l,l'> : lleq(A) --> l Int Vset(i) <= l'";  | 
| 
34
 
747f1aad03cf
changed filenames to lower case name of theory the file contains
 
clasohm 
parents:  
diff
changeset
 | 
38  | 
by (etac trans_induct 1);  | 
| 173 | 39  | 
by (REPEAT (resolve_tac [allI, impI] 1));  | 
| 
34
 
747f1aad03cf
changed filenames to lower case name of theory the file contains
 
clasohm 
parents:  
diff
changeset
 | 
40  | 
by (etac LList_Eq.elim 1);  | 
| 173 | 41  | 
by (rewrite_goals_tac (QInr_def::LList.con_defs));  | 
42  | 
by (safe_tac lleq_cs);  | 
|
43  | 
by (fast_tac (subset_cs addSEs [Ord_trans, make_elim bspec]) 1);  | 
|
| 
34
 
747f1aad03cf
changed filenames to lower case name of theory the file contains
 
clasohm 
parents:  
diff
changeset
 | 
44  | 
val lleq_Int_Vset_subset_lemma = result();  | 
| 
 
747f1aad03cf
changed filenames to lower case name of theory the file contains
 
clasohm 
parents:  
diff
changeset
 | 
45  | 
|
| 
 
747f1aad03cf
changed filenames to lower case name of theory the file contains
 
clasohm 
parents:  
diff
changeset
 | 
46  | 
val lleq_Int_Vset_subset = standard  | 
| 
 
747f1aad03cf
changed filenames to lower case name of theory the file contains
 
clasohm 
parents:  
diff
changeset
 | 
47  | 
(lleq_Int_Vset_subset_lemma RS spec RS spec RS mp);  | 
| 
 
747f1aad03cf
changed filenames to lower case name of theory the file contains
 
clasohm 
parents:  
diff
changeset
 | 
48  | 
|
| 
 
747f1aad03cf
changed filenames to lower case name of theory the file contains
 
clasohm 
parents:  
diff
changeset
 | 
49  | 
|
| 
 
747f1aad03cf
changed filenames to lower case name of theory the file contains
 
clasohm 
parents:  
diff
changeset
 | 
50  | 
(*lleq(A) is a symmetric relation because qconverse(lleq(A)) is a fixedpoint*)  | 
| 120 | 51  | 
val [prem] = goal LList_Eq.thy "<l,l'> : lleq(A) ==> <l',l> : lleq(A)";  | 
52  | 
by (rtac (prem RS converseI RS LList_Eq.coinduct) 1);  | 
|
53  | 
by (rtac (LList_Eq.dom_subset RS converse_type) 1);  | 
|
54  | 
by (safe_tac converse_cs);  | 
|
| 
34
 
747f1aad03cf
changed filenames to lower case name of theory the file contains
 
clasohm 
parents:  
diff
changeset
 | 
55  | 
by (etac LList_Eq.elim 1);  | 
| 
 
747f1aad03cf
changed filenames to lower case name of theory the file contains
 
clasohm 
parents:  
diff
changeset
 | 
56  | 
by (ALLGOALS (fast_tac qconverse_cs));  | 
| 
 
747f1aad03cf
changed filenames to lower case name of theory the file contains
 
clasohm 
parents:  
diff
changeset
 | 
57  | 
val lleq_symmetric = result();  | 
| 
 
747f1aad03cf
changed filenames to lower case name of theory the file contains
 
clasohm 
parents:  
diff
changeset
 | 
58  | 
|
| 120 | 59  | 
goal LList_Eq.thy "!!l l'. <l,l'> : lleq(A) ==> l=l'";  | 
| 
34
 
747f1aad03cf
changed filenames to lower case name of theory the file contains
 
clasohm 
parents:  
diff
changeset
 | 
60  | 
by (rtac equalityI 1);  | 
| 
 
747f1aad03cf
changed filenames to lower case name of theory the file contains
 
clasohm 
parents:  
diff
changeset
 | 
61  | 
by (REPEAT (ares_tac [lleq_Int_Vset_subset RS Int_Vset_subset] 1  | 
| 
 
747f1aad03cf
changed filenames to lower case name of theory the file contains
 
clasohm 
parents:  
diff
changeset
 | 
62  | 
ORELSE etac lleq_symmetric 1));  | 
| 
 
747f1aad03cf
changed filenames to lower case name of theory the file contains
 
clasohm 
parents:  
diff
changeset
 | 
63  | 
val lleq_implies_equal = result();  | 
| 
 
747f1aad03cf
changed filenames to lower case name of theory the file contains
 
clasohm 
parents:  
diff
changeset
 | 
64  | 
|
| 
 
747f1aad03cf
changed filenames to lower case name of theory the file contains
 
clasohm 
parents:  
diff
changeset
 | 
65  | 
val [eqprem,lprem] = goal LList_Eq.thy  | 
| 120 | 66  | 
"[| l=l'; l: llist(A) |] ==> <l,l'> : lleq(A)";  | 
67  | 
by (res_inst_tac [("X", "{<l,l>. l: llist(A)}")] LList_Eq.coinduct 1);
 | 
|
| 
34
 
747f1aad03cf
changed filenames to lower case name of theory the file contains
 
clasohm 
parents:  
diff
changeset
 | 
68  | 
by (rtac (lprem RS RepFunI RS (eqprem RS subst)) 1);  | 
| 
 
747f1aad03cf
changed filenames to lower case name of theory the file contains
 
clasohm 
parents:  
diff
changeset
 | 
69  | 
by (safe_tac qpair_cs);  | 
| 
 
747f1aad03cf
changed filenames to lower case name of theory the file contains
 
clasohm 
parents:  
diff
changeset
 | 
70  | 
by (etac LList.elim 1);  | 
| 120 | 71  | 
by (ALLGOALS (fast_tac pair_cs));  | 
| 
34
 
747f1aad03cf
changed filenames to lower case name of theory the file contains
 
clasohm 
parents:  
diff
changeset
 | 
72  | 
val equal_llist_implies_leq = result();  | 
| 
 
747f1aad03cf
changed filenames to lower case name of theory the file contains
 
clasohm 
parents:  
diff
changeset
 | 
73  |