src/ZF/ex/listn.ML
author lcp
Thu, 30 Sep 1993 10:54:01 +0100
changeset 16 0b033d50ca1c
parent 7 268f93ab3bc4
child 56 2caa6f49f06e
permissions -rw-r--r--
ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext ex/prop-log/hyps_thms_if: split up the fast_tac call for more speed called expandshort
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     1
(*  Title: 	ZF/ex/listn
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     3
    Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     4
    Copyright   1993  University of Cambridge
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     6
Inductive definition of lists of n elements
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
See Ch. Paulin-Mohring, Inductive Definitions in the System Coq.
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     9
Research Report 92-49, LIP, ENS Lyon.  Dec 1992.
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    10
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    11
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    12
structure ListN = Inductive_Fun
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    13
 (val thy = ListFn.thy addconsts [(["listn"],"i=>i")];
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    14
  val rec_doms = [("listn", "nat*list(A)")];
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    15
  val sintrs = 
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    16
      ["<0,Nil> : listn(A)",
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    17
       "[| a: A;  <n,l> : listn(A) |] ==> <succ(n), Cons(a,l)> : listn(A)"];
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    18
  val monos = [];
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    19
  val con_defs = [];
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    20
  val type_intrs = nat_typechecks@List.intrs@[SigmaI]
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    21
  val type_elims = [SigmaE2]);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    22
7
268f93ab3bc4 Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents: 0
diff changeset
    23
val listn_induct = standard 
268f93ab3bc4 Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents: 0
diff changeset
    24
    (ListN.mutual_induct RS spec RS spec RSN (2,rev_mp));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    25
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    26
val [major] = goal ListN.thy "l:list(A) ==> <length(l),l> : listn(A)";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    27
by (rtac (major RS List.induct) 1);
7
268f93ab3bc4 Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents: 0
diff changeset
    28
by (ALLGOALS (asm_simp_tac list_ss));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    29
by (REPEAT (ares_tac ListN.intrs 1));
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    30
val list_into_listn = result();
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    31
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    32
goal ListN.thy "<n,l> : listn(A) <-> l:list(A) & length(l)=n";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    33
by (rtac iffI 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    34
by (etac listn_induct 1);
7
268f93ab3bc4 Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents: 0
diff changeset
    35
by (safe_tac (ZF_cs addSIs (list_typechecks @
268f93ab3bc4 Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents: 0
diff changeset
    36
			    [length_Nil, length_Cons, list_into_listn])));
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    37
val listn_iff = result();
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    38
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    39
goal ListN.thy "listn(A)``{n} = {l:list(A). length(l)=n}";
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    40
by (rtac equality_iffI 1);
7
268f93ab3bc4 Installation of new simplifier for ZF/ex. The hom_ss example in misc.ML is
lcp
parents: 0
diff changeset
    41
by (simp_tac (list_ss addsimps [listn_iff,separation,image_singleton_iff]) 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    42
val listn_image_eq = result();
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    43
16
0b033d50ca1c ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 7
diff changeset
    44
goalw ListN.thy ListN.defs "!!A B. A<=B ==> listn(A) <= listn(B)";
0b033d50ca1c ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 7
diff changeset
    45
by (rtac lfp_mono 1);
0b033d50ca1c ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 7
diff changeset
    46
by (REPEAT (rtac ListN.bnd_mono 1));
0b033d50ca1c ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 7
diff changeset
    47
by (REPEAT (ares_tac ([univ_mono,Sigma_mono,list_mono] @ basic_monos) 1));
0b033d50ca1c ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 7
diff changeset
    48
val listn_mono = result();
0b033d50ca1c ex/{bin.ML,comb.ML,prop.ML}: replaced NewSext by Syntax.simple_sext
lcp
parents: 7
diff changeset
    49