src/ZF/ex/counit.ML
changeset 95 2246a80b1cb5
child 120 09287f26bfb8
equal deleted inserted replaced
94:40f292719398 95:2246a80b1cb5
       
     1 (*  Title: 	ZF/ex/counit.ML
       
     2     ID:         $Id$
       
     3     Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
       
     4     Copyright   1993  University of Cambridge
       
     5 
       
     6 Trivial co-datatype definitions, one of which goes wrong!
       
     7 
       
     8 Need to find sufficient conditions for co-datatypes to work correctly!
       
     9 *)
       
    10 
       
    11 (*This degenerate definition does not work well because the one constructor's
       
    12   definition is trivial!
       
    13 *)
       
    14 structure CoUnit = Co_Datatype_Fun
       
    15  (val thy = QUniv.thy;
       
    16   val rec_specs = 
       
    17       [("counit", "quniv(0)",
       
    18 	  [(["Con"],	"i=>i")])];
       
    19   val rec_styp = "i";
       
    20   val ext = None
       
    21   val sintrs = ["x: counit ==> Con(x) : counit"];
       
    22   val monos = [];
       
    23   val type_intrs = co_datatype_intrs
       
    24   val type_elims = co_datatype_elims);
       
    25   
       
    26 val [ConI] = CoUnit.intrs;
       
    27 
       
    28 (*USELESS because folding on Con(?xa) == ?xa fails*)
       
    29 val ConE = CoUnit.mk_cases CoUnit.con_defs "Con(x) : counit";
       
    30 
       
    31 (*Proving freeness results*)
       
    32 val Con_iff = CoUnit.mk_free "Con(x)=Con(y) <-> x=y";
       
    33 
       
    34 (*Should be a singleton, not everything!*)
       
    35 goal CoUnit.thy "counit = quniv(0)";
       
    36 by (rtac (CoUnit.dom_subset RS equalityI) 1);
       
    37 by (rtac subsetI 1);
       
    38 by (etac CoUnit.co_induct 1);
       
    39 by (rtac subset_refl 1);
       
    40 by (rewrite_goals_tac CoUnit.con_defs);
       
    41 by (fast_tac ZF_cs 1);
       
    42 val counit_eq_univ = result();
       
    43 
       
    44 
       
    45 (*****************************************************************)
       
    46 
       
    47 (*A similar example, but the constructor is non-degenerate and it works!
       
    48   The resulting set is a singleton.
       
    49 *)
       
    50 
       
    51 structure CoUnit2 = Co_Datatype_Fun
       
    52  (val thy = QUniv.thy;
       
    53   val rec_specs = 
       
    54       [("counit2", "quniv(0)",
       
    55 	  [(["Con2"],	"[i,i]=>i")])];
       
    56   val rec_styp = "i";
       
    57   val ext = None
       
    58   val sintrs = ["[| x: counit2;  y: counit2 |] ==> Con2(x,y) : counit2"];
       
    59   val monos = [];
       
    60   val type_intrs = co_datatype_intrs
       
    61   val type_elims = co_datatype_elims);
       
    62 
       
    63 val [Con2I] = CoUnit2.intrs;
       
    64 
       
    65 val Con2E = CoUnit2.mk_cases CoUnit2.con_defs "Con2(x,y) : counit2";
       
    66 
       
    67 (*Proving freeness results*)
       
    68 val Con2_iff = CoUnit2.mk_free "Con2(x,y)=Con2(x',y') <-> x=x' & y=y'";
       
    69 
       
    70 goalw CoUnit2.thy CoUnit2.con_defs "bnd_mono(univ(0), %x. Con2(x,x))";
       
    71 by (rtac bnd_monoI 1);
       
    72 by (REPEAT (ares_tac [subset_refl, QPair_subset_univ, QPair_mono] 1));
       
    73 val Con2_bnd_mono = result();
       
    74 
       
    75 goal CoUnit2.thy "lfp(univ(0), %x. Con2(x,x)) : counit2";
       
    76 by (rtac (singletonI RS CoUnit2.co_induct) 1);
       
    77 by (rtac (qunivI RS singleton_subsetI) 1);
       
    78 by (rtac ([lfp_subset, empty_subsetI RS univ_mono] MRS subset_trans) 1);
       
    79 by (fast_tac (ZF_cs addSIs [Con2_bnd_mono RS lfp_Tarski]) 1);
       
    80 val lfp_Con2_in_counit2 = result();
       
    81 
       
    82 (*borrowed from ex/llist_eq.ML! the proofs are almost identical!*)
       
    83 val lleq_cs = subset_cs
       
    84 	addSIs [succI1, Int_Vset_0_subset,
       
    85 		QPair_Int_Vset_succ_subset_trans,
       
    86 		QPair_Int_Vset_subset_trans];
       
    87 
       
    88 goal CoUnit2.thy
       
    89     "!!i. Ord(i) ==> ALL x y. x: counit2 & y: counit2 --> x Int Vset(i) <= y";
       
    90 by (etac trans_induct 1);
       
    91 by (safe_tac subset_cs);
       
    92 by (etac CoUnit2.elim 1);
       
    93 by (etac CoUnit2.elim 1);
       
    94 by (safe_tac subset_cs);
       
    95 by (rewrite_goals_tac CoUnit2.con_defs);
       
    96 by (etac Ord_cases 1 THEN REPEAT_FIRST hyp_subst_tac);
       
    97 (*0 case*)
       
    98 by (fast_tac lleq_cs 1);
       
    99 (*succ(j) case*)
       
   100 by (fast_tac lleq_cs 1);
       
   101 (*Limit(i) case*)
       
   102 by (etac (Limit_Vfrom_eq RS ssubst) 1);
       
   103 by (rtac (Int_UN_distrib RS ssubst) 1);
       
   104 by (fast_tac lleq_cs 1);
       
   105 val counit2_Int_Vset_subset_lemma = result();
       
   106 
       
   107 val counit2_Int_Vset_subset = standard
       
   108 	(counit2_Int_Vset_subset_lemma RS spec RS spec RS mp);
       
   109 
       
   110 goal CoUnit2.thy "!!x y. [| x: counit2;  y: counit2 |] ==> x=y";
       
   111 by (rtac equalityI 1);
       
   112 by (REPEAT (ares_tac [conjI, counit2_Int_Vset_subset RS Int_Vset_subset] 1));
       
   113 val counit2_implies_equal = result();
       
   114 
       
   115 goal CoUnit2.thy "counit2 = {lfp(univ(0), %x. Con2(x,x))}";
       
   116 by (rtac equalityI 1);
       
   117 by (rtac (lfp_Con2_in_counit2 RS singleton_subsetI) 2);
       
   118 by (rtac subsetI 1);
       
   119 by (dtac (lfp_Con2_in_counit2 RS counit2_implies_equal) 1);
       
   120 by (etac subst 1);
       
   121 by (rtac singletonI 1);
       
   122 val counit2_eq_univ = result();