src/HOL/PreList.thy
author wenzelm
Thu, 04 Jul 2002 16:48:21 +0200
changeset 13297 e4ae0732e2be
parent 12919 d6a0d168291e
child 13878 90ca3815e4b2
permissions -rw-r--r--
tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
10519
ade64af4c57c hide many names from Datatype_Universe.
nipkow
parents: 10261
diff changeset
     1
(*  Title:      HOL/PreList.thy
8563
2746bc9a7ef2 comments
nipkow
parents: 8490
diff changeset
     2
    ID:         $Id$
10733
59f82484e000 hide type node item;
wenzelm
parents: 10680
diff changeset
     3
    Author:     Tobias Nipkow and Markus Wenzel
8563
2746bc9a7ef2 comments
nipkow
parents: 8490
diff changeset
     4
    Copyright   2000 TU Muenchen
2746bc9a7ef2 comments
nipkow
parents: 8490
diff changeset
     5
2746bc9a7ef2 comments
nipkow
parents: 8490
diff changeset
     6
A basis for building theory List on. Is defined separately to serve as a
2746bc9a7ef2 comments
nipkow
parents: 8490
diff changeset
     7
basis for theory ToyList in the documentation.
2746bc9a7ef2 comments
nipkow
parents: 8490
diff changeset
     8
*)
8490
6e0f23304061 added HOL/PreLIst.thy;
wenzelm
parents:
diff changeset
     9
6e0f23304061 added HOL/PreLIst.thy;
wenzelm
parents:
diff changeset
    10
theory PreList =
12919
d6a0d168291e removed theory Option;
wenzelm
parents: 12869
diff changeset
    11
  Wellfounded_Relations + NatSimprocs + Recdef + Relation_Power:
12020
a24373086908 theory Calculation move to Set;
wenzelm
parents: 11955
diff changeset
    12
a24373086908 theory Calculation move to Set;
wenzelm
parents: 11955
diff changeset
    13
(*belongs to theory Divides*)
12304
8df202daf55d tuned declarations;
wenzelm
parents: 12020
diff changeset
    14
declare dvdI [intro?]  dvdE [elim?]  dvd_trans [trans]
8490
6e0f23304061 added HOL/PreLIst.thy;
wenzelm
parents:
diff changeset
    15
12397
6766aa05e4eb less_induct, wf_induct_rule;
wenzelm
parents: 12304
diff changeset
    16
(*belongs to theory Nat*)
6766aa05e4eb less_induct, wf_induct_rule;
wenzelm
parents: 12304
diff changeset
    17
lemmas less_induct = nat_less_induct [rule_format, case_names less]
6766aa05e4eb less_induct, wf_induct_rule;
wenzelm
parents: 12304
diff changeset
    18
10261
bb2f1e859177 tuned declarations;
wenzelm
parents: 10212
diff changeset
    19
(*belongs to theory Wellfounded_Recursion*)
12397
6766aa05e4eb less_induct, wf_induct_rule;
wenzelm
parents: 12304
diff changeset
    20
lemmas wf_induct_rule = wf_induct [rule_format, case_names less, induct set: wf]
9066
b1e874e38dab theorems [cases type: bool] = case_split;
wenzelm
parents: 8862
diff changeset
    21
8490
6e0f23304061 added HOL/PreLIst.thy;
wenzelm
parents:
diff changeset
    22
end