src/ZF/ex/LList_Eq.ML
author lcp
Tue, 30 Nov 1993 11:08:18 +0100
changeset 173 85071e6ad295
parent 120 09287f26bfb8
child 279 7738aed3f84d
permissions -rw-r--r--
ZF/ex/llist_eq/lleq_Int_Vset_subset_lemma, ZF/ex/counit/counit2_Int_Vset_subset_lemma: now uses QPair_Int_Vset_subset_UN ZF/ex/llistfn/flip_llist_quniv_lemma: now uses transfinite induction and QPair_Int_Vset_subset_UN ZF/ex/llist/llist_quniv_lemma: now uses transfinite induction and QPair_Int_Vset_subset_UN
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
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
09287f26bfb8 changed all co- and co_ to co
lcp
parents: 95
diff changeset
     6
Equality for llist(A) as a greatest fixed point
09287f26bfb8 changed all co- and co_ to co
lcp
parents: 95
diff changeset
     7
***)
34
747f1aad03cf changed filenames to lower case name of theory the file contains
clasohm
parents:
diff changeset
     8
120
09287f26bfb8 changed all co- and co_ to co
lcp
parents: 95
diff changeset
     9
(*Previously used <*> in the domain and variant pairs as elements.  But
09287f26bfb8 changed all co- and co_ to co
lcp
parents: 95
diff changeset
    10
  standard pairs work just as well.  To use variant pairs, must change prefix
09287f26bfb8 changed all co- and co_ to co
lcp
parents: 95
diff changeset
    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
09287f26bfb8 changed all co- and co_ to co
lcp
parents: 95
diff changeset
    13
structure LList_Eq = CoInductive_Fun
34
747f1aad03cf changed filenames to lower case name of theory the file contains
clasohm
parents:
diff changeset
    14
 (val thy = LList.thy addconsts [(["lleq"],"i=>i")];
120
09287f26bfb8 changed all co- and co_ to co
lcp
parents: 95
diff changeset
    15
  val rec_doms = [("lleq", "llist(A) * llist(A)")];
34
747f1aad03cf changed filenames to lower case name of theory the file contains
clasohm
parents:
diff changeset
    16
  val sintrs = 
120
09287f26bfb8 changed all co- and co_ to co
lcp
parents: 95
diff changeset
    17
      ["<LNil, LNil> : lleq(A)",
09287f26bfb8 changed all co- and co_ to co
lcp
parents: 95
diff changeset
    18
       "[| a:A;  <l, l'>: lleq(A) |] ==> <LCons(a,l), LCons(a,l')> : lleq(A)"];
34
747f1aad03cf changed filenames to lower case name of theory the file contains
clasohm
parents:
diff changeset
    19
  val monos = [];
747f1aad03cf changed filenames to lower case name of theory the file contains
clasohm
parents:
diff changeset
    20
  val con_defs = [];
120
09287f26bfb8 changed all co- and co_ to co
lcp
parents: 95
diff changeset
    21
  val type_intrs = LList.intrs@[SigmaI];
09287f26bfb8 changed all co- and co_ to co
lcp
parents: 95
diff changeset
    22
  val type_elims = [SigmaE2]);
34
747f1aad03cf changed filenames to lower case name of theory the file contains
clasohm
parents:
diff changeset
    23
747f1aad03cf changed filenames to lower case name of theory the file contains
clasohm
parents:
diff changeset
    24
(** Alternatives for above:
747f1aad03cf changed filenames to lower case name of theory the file contains
clasohm
parents:
diff changeset
    25
  val con_defs = LList.con_defs
120
09287f26bfb8 changed all co- and co_ to co
lcp
parents: 95
diff changeset
    26
  val type_intrs = codatatype_intrs
34
747f1aad03cf changed filenames to lower case name of theory the file contains
clasohm
parents:
diff changeset
    27
  val type_elims = [quniv_QPair_E]
747f1aad03cf changed filenames to lower case name of theory the file contains
clasohm
parents:
diff changeset
    28
**)
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
val lleq_cs = subset_cs
173
85071e6ad295 ZF/ex/llist_eq/lleq_Int_Vset_subset_lemma,
lcp
parents: 120
diff changeset
    31
	addSIs [QPair_Int_Vset_subset_UN RS subset_trans, QPair_mono]
85071e6ad295 ZF/ex/llist_eq/lleq_Int_Vset_subset_lemma,
lcp
parents: 120
diff changeset
    32
        addSEs [Ord_in_Ord, Pair_inject];
34
747f1aad03cf changed filenames to lower case name of theory the file contains
clasohm
parents:
diff changeset
    33
173
85071e6ad295 ZF/ex/llist_eq/lleq_Int_Vset_subset_lemma,
lcp
parents: 120
diff changeset
    34
(*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
    35
goal LList_Eq.thy
120
09287f26bfb8 changed all co- and co_ to co
lcp
parents: 95
diff changeset
    36
   "!!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
    37
by (etac trans_induct 1);
173
85071e6ad295 ZF/ex/llist_eq/lleq_Int_Vset_subset_lemma,
lcp
parents: 120
diff changeset
    38
by (REPEAT (resolve_tac [allI, impI] 1));
34
747f1aad03cf changed filenames to lower case name of theory the file contains
clasohm
parents:
diff changeset
    39
by (etac LList_Eq.elim 1);
173
85071e6ad295 ZF/ex/llist_eq/lleq_Int_Vset_subset_lemma,
lcp
parents: 120
diff changeset
    40
by (rewrite_goals_tac (QInr_def::LList.con_defs));
85071e6ad295 ZF/ex/llist_eq/lleq_Int_Vset_subset_lemma,
lcp
parents: 120
diff changeset
    41
by (safe_tac lleq_cs);
85071e6ad295 ZF/ex/llist_eq/lleq_Int_Vset_subset_lemma,
lcp
parents: 120
diff changeset
    42
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
    43
val lleq_Int_Vset_subset_lemma = result();
747f1aad03cf changed filenames to lower case name of theory the file contains
clasohm
parents:
diff changeset
    44
747f1aad03cf changed filenames to lower case name of theory the file contains
clasohm
parents:
diff changeset
    45
val lleq_Int_Vset_subset = standard
747f1aad03cf changed filenames to lower case name of theory the file contains
clasohm
parents:
diff changeset
    46
	(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
    47
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
(*lleq(A) is a symmetric relation because qconverse(lleq(A)) is a fixedpoint*)
120
09287f26bfb8 changed all co- and co_ to co
lcp
parents: 95
diff changeset
    50
val [prem] = goal LList_Eq.thy "<l,l'> : lleq(A) ==> <l',l> : lleq(A)";
09287f26bfb8 changed all co- and co_ to co
lcp
parents: 95
diff changeset
    51
by (rtac (prem RS converseI RS LList_Eq.coinduct) 1);
09287f26bfb8 changed all co- and co_ to co
lcp
parents: 95
diff changeset
    52
by (rtac (LList_Eq.dom_subset RS converse_type) 1);
09287f26bfb8 changed all co- and co_ to co
lcp
parents: 95
diff changeset
    53
by (safe_tac converse_cs);
34
747f1aad03cf changed filenames to lower case name of theory the file contains
clasohm
parents:
diff changeset
    54
by (etac LList_Eq.elim 1);
747f1aad03cf changed filenames to lower case name of theory the file contains
clasohm
parents:
diff changeset
    55
by (ALLGOALS (fast_tac qconverse_cs));
747f1aad03cf changed filenames to lower case name of theory the file contains
clasohm
parents:
diff changeset
    56
val lleq_symmetric = result();
747f1aad03cf changed filenames to lower case name of theory the file contains
clasohm
parents:
diff changeset
    57
120
09287f26bfb8 changed all co- and co_ to co
lcp
parents: 95
diff changeset
    58
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
    59
by (rtac equalityI 1);
747f1aad03cf changed filenames to lower case name of theory the file contains
clasohm
parents:
diff changeset
    60
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
    61
     ORELSE etac lleq_symmetric 1));
747f1aad03cf changed filenames to lower case name of theory the file contains
clasohm
parents:
diff changeset
    62
val lleq_implies_equal = result();
747f1aad03cf changed filenames to lower case name of theory the file contains
clasohm
parents:
diff changeset
    63
747f1aad03cf changed filenames to lower case name of theory the file contains
clasohm
parents:
diff changeset
    64
val [eqprem,lprem] = goal LList_Eq.thy
120
09287f26bfb8 changed all co- and co_ to co
lcp
parents: 95
diff changeset
    65
    "[| l=l';  l: llist(A) |] ==> <l,l'> : lleq(A)";
09287f26bfb8 changed all co- and co_ to co
lcp
parents: 95
diff changeset
    66
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
    67
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
    68
by (safe_tac qpair_cs);
747f1aad03cf changed filenames to lower case name of theory the file contains
clasohm
parents:
diff changeset
    69
by (etac LList.elim 1);
120
09287f26bfb8 changed all co- and co_ to co
lcp
parents: 95
diff changeset
    70
by (ALLGOALS (fast_tac pair_cs));
34
747f1aad03cf changed filenames to lower case name of theory the file contains
clasohm
parents:
diff changeset
    71
val equal_llist_implies_leq = result();
747f1aad03cf changed filenames to lower case name of theory the file contains
clasohm
parents:
diff changeset
    72