author | paulson |
Tue, 08 Jan 2002 16:09:09 +0100 | |
changeset 12667 | 7e6eaaa125f2 |
parent 12620 | 4e6626725e21 |
child 12776 | 249600a63ba9 |
permissions | -rw-r--r-- |
12552
d2d2ab3f1f37
separation of the AC part of Main into Main_ZFC, plus a few new lemmas
paulson
parents:
12426
diff
changeset
|
1 |
(*$Id$ |
d2d2ab3f1f37
separation of the AC part of Main into Main_ZFC, plus a few new lemmas
paulson
parents:
12426
diff
changeset
|
2 |
theory Main includes everything except AC*) |
5157 | 3 |
|
12552
d2d2ab3f1f37
separation of the AC part of Main into Main_ZFC, plus a few new lemmas
paulson
parents:
12426
diff
changeset
|
4 |
theory Main = Update + List + EquivClass + IntDiv + CardinalArith: |
12426 | 5 |
|
6 |
(* belongs to theory Trancl *) |
|
7 |
lemmas rtrancl_induct = rtrancl_induct [case_names initial step, induct set: rtrancl] |
|
8 |
and trancl_induct = trancl_induct [case_names initial step, induct set: trancl] |
|
9 |
and converse_trancl_induct = converse_trancl_induct [case_names initial step, consumes 1] |
|
10 |
and rtrancl_full_induct = rtrancl_full_induct [case_names initial step, consumes 1] |
|
11 |
||
12 |
(* belongs to theory WF *) |
|
13 |
lemmas wf_induct = wf_induct [induct set: wf] |
|
14 |
and wf_induct_rule = wf_induct [rule_format, induct set: wf] |
|
15 |
and wf_on_induct = wf_on_induct [consumes 2, induct set: wf_on] |
|
16 |
and wf_on_induct_rule = wf_on_induct [rule_format, consumes 2, induct set: wf_on] |
|
17 |
||
18 |
(* belongs to theory Ordinal *) |
|
12667 | 19 |
declare Ord_Least [intro,simp,TC] |
12426 | 20 |
lemmas Ord_induct = Ord_induct [consumes 2] |
21 |
and Ord_induct_rule = Ord_induct [rule_format, consumes 2] |
|
22 |
and trans_induct = trans_induct [consumes 1] |
|
23 |
and trans_induct_rule = trans_induct [rule_format, consumes 1] |
|
24 |
and trans_induct3 = trans_induct3 [case_names 0 succ limit, consumes 1] |
|
25 |
and trans_induct3_rule = trans_induct3 [rule_format, case_names 0 succ limit, consumes 1] |
|
9175 | 26 |
|
12426 | 27 |
(* belongs to theory Nat *) |
28 |
lemmas nat_induct = nat_induct [case_names 0 succ, induct set: nat] |
|
29 |
and complete_induct = complete_induct [case_names less, consumes 1] |
|
30 |
and complete_induct_rule = complete_induct [rule_format, case_names less, consumes 1] |
|
31 |
and diff_induct = diff_induct [case_names 0 0_succ succ_succ, consumes 2] |
|
32 |
||
33 |
(* belongs to theory Epsilon *) |
|
34 |
lemmas eclose_induct = eclose_induct [induct set: eclose] |
|
35 |
and eclose_induct_down = eclose_induct_down [consumes 1] |
|
36 |
||
37 |
(* belongs to theory Finite *) |
|
38 |
lemmas Fin_induct = Fin_induct [case_names 0 cons, induct set: Fin] |
|
39 |
||
40 |
(* belongs to theory CardinalArith *) |
|
41 |
lemmas Finite_induct = Finite_induct [case_names 0 cons, induct set: Finite] |
|
42 |
||
43 |
(* belongs to theory List *) |
|
44 |
lemmas list_append_induct = list_append_induct [case_names Nil snoc, consumes 1] |
|
45 |
||
46 |
(* belongs to theory IntDiv *) |
|
47 |
lemmas posDivAlg_induct = posDivAlg_induct [consumes 2] |
|
48 |
and negDivAlg_induct = negDivAlg_induct [consumes 2] |
|
49 |
||
12620 | 50 |
|
51 |
(* belongs to theory Epsilon *) |
|
52 |
||
53 |
lemma def_transrec2: |
|
54 |
"(!!x. f(x)==transrec2(x,a,b)) |
|
55 |
==> f(0) = a & |
|
56 |
f(succ(i)) = b(i, f(i)) & |
|
57 |
(Limit(K) --> f(K) = (UN j<K. f(j)))" |
|
58 |
by (simp add: transrec2_Limit) |
|
59 |
||
12426 | 60 |
end |