src/FOL/ex/List.ML
author wenzelm
Thu, 18 Jun 1998 18:28:45 +0200
changeset 5050 e925308df78b
parent 3922 ca23ee574faa
child 17245 1c519a3cca59
permissions -rw-r--r--
isatool fixgoal;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1459
d12da312eff4 expanded tabs
clasohm
parents: 755
diff changeset
     1
(*  Title:      FOL/ex/list
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     2
    ID:         $Id$
1459
d12da312eff4 expanded tabs
clasohm
parents: 755
diff changeset
     3
    Author:     Tobias Nipkow
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     4
    Copyright   1991  University of Cambridge
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     5
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     6
For ex/list.thy.  Examples of simplification and induction on lists
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     7
*)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     8
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
     9
open List;
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    10
3922
ca23ee574faa tuned "." syntax;
wenzelm
parents: 3835
diff changeset
    11
val prems = goal List.thy "[| P([]);  !!x l. P(x . l) |] ==> All(P)";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    12
by (rtac list_ind 1);
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    13
by (REPEAT (resolve_tac (prems@[allI,impI]) 1));
725
d9c629fbacc6 replaced 'val ... = result();' by 'qed "...";'
clasohm
parents: 0
diff changeset
    14
qed "list_exh";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    15
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1459
diff changeset
    16
Addsimps [list_distinct1, list_distinct2, app_nil, app_cons,
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1459
diff changeset
    17
	  hd_eq, tl_eq, forall_nil, forall_cons, list_free,
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1459
diff changeset
    18
	  len_nil, len_cons, at_0, at_succ];
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    19
5050
e925308df78b isatool fixgoal;
wenzelm
parents: 3922
diff changeset
    20
Goal "~l=[] --> hd(l).tl(l) = l";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1459
diff changeset
    21
by (IND_TAC list_exh Simp_tac "l" 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    22
result();
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    23
5050
e925308df78b isatool fixgoal;
wenzelm
parents: 3922
diff changeset
    24
Goal "(l1++l2)++l3 = l1++(l2++l3)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1459
diff changeset
    25
by (IND_TAC list_ind Simp_tac "l1" 1);
755
dfb3894d78e4 replaced "val ... = result()" by "qed ..."
clasohm
parents: 725
diff changeset
    26
qed "append_assoc";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    27
5050
e925308df78b isatool fixgoal;
wenzelm
parents: 3922
diff changeset
    28
Goal "l++[] = l";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1459
diff changeset
    29
by (IND_TAC list_ind Simp_tac "l" 1);
755
dfb3894d78e4 replaced "val ... = result()" by "qed ..."
clasohm
parents: 725
diff changeset
    30
qed "app_nil_right";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    31
5050
e925308df78b isatool fixgoal;
wenzelm
parents: 3922
diff changeset
    32
Goal "l1++l2=[] <-> l1=[] & l2=[]";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1459
diff changeset
    33
by (IND_TAC list_exh Simp_tac "l1" 1);
755
dfb3894d78e4 replaced "val ... = result()" by "qed ..."
clasohm
parents: 725
diff changeset
    34
qed "app_eq_nil_iff";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    35
5050
e925308df78b isatool fixgoal;
wenzelm
parents: 3922
diff changeset
    36
Goal "forall(l++l',P) <-> forall(l,P) & forall(l',P)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1459
diff changeset
    37
by (IND_TAC list_ind Simp_tac "l" 1);
755
dfb3894d78e4 replaced "val ... = result()" by "qed ..."
clasohm
parents: 725
diff changeset
    38
qed "forall_app";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    39
5050
e925308df78b isatool fixgoal;
wenzelm
parents: 3922
diff changeset
    40
Goal "forall(l,%x. P(x)&Q(x)) <-> forall(l,P) & forall(l,Q)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1459
diff changeset
    41
by (IND_TAC list_ind Simp_tac "l" 1);
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1459
diff changeset
    42
by (Fast_tac 1);
755
dfb3894d78e4 replaced "val ... = result()" by "qed ..."
clasohm
parents: 725
diff changeset
    43
qed "forall_conj";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    44
5050
e925308df78b isatool fixgoal;
wenzelm
parents: 3922
diff changeset
    45
Goal "~l=[] --> forall(l,P) <-> P(hd(l)) & forall(tl(l),P)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1459
diff changeset
    46
by (IND_TAC list_ind Simp_tac "l" 1);
755
dfb3894d78e4 replaced "val ... = result()" by "qed ..."
clasohm
parents: 725
diff changeset
    47
qed "forall_ne";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    48
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    49
(*** Lists with natural numbers ***)
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    50
5050
e925308df78b isatool fixgoal;
wenzelm
parents: 3922
diff changeset
    51
Goal "len(l1++l2) = len(l1)+len(l2)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1459
diff changeset
    52
by (IND_TAC list_ind Simp_tac "l1" 1);
755
dfb3894d78e4 replaced "val ... = result()" by "qed ..."
clasohm
parents: 725
diff changeset
    53
qed "len_app";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    54
5050
e925308df78b isatool fixgoal;
wenzelm
parents: 3922
diff changeset
    55
Goal "i<len(l1) --> at(l1++l2,i) = at(l1,i)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1459
diff changeset
    56
by (IND_TAC list_ind Simp_tac "l1" 1);
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    57
by (REPEAT (rtac allI 1));
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    58
by (rtac impI 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1459
diff changeset
    59
by (ALL_IND_TAC nat_exh Asm_simp_tac 1);
755
dfb3894d78e4 replaced "val ... = result()" by "qed ..."
clasohm
parents: 725
diff changeset
    60
qed "at_app1";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    61
5050
e925308df78b isatool fixgoal;
wenzelm
parents: 3922
diff changeset
    62
Goal "at(l1++(x . l2), len(l1)) = x";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1459
diff changeset
    63
by (IND_TAC list_ind Simp_tac "l1" 1);
755
dfb3894d78e4 replaced "val ... = result()" by "qed ..."
clasohm
parents: 725
diff changeset
    64
qed "at_app_hd2";
0
a5a9c433f639 Initial revision
clasohm
parents:
diff changeset
    65