src/ZF/Main.thy
author paulson
Tue, 08 Jan 2002 16:09:09 +0100
changeset 12667 7e6eaaa125f2
parent 12620 4e6626725e21
child 12776 249600a63ba9
permissions -rw-r--r--
Added some simprules proofs. Converted theories CardinalArith and OrdQuant to Isar style
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
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
6e03de8ec2b4 as in HOL
paulson
parents:
diff changeset
     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
9032bdbc2125 new-style theory;
wenzelm
parents: 9578
diff changeset
     5
9032bdbc2125 new-style theory;
wenzelm
parents: 9578
diff changeset
     6
(* belongs to theory Trancl *)
9032bdbc2125 new-style theory;
wenzelm
parents: 9578
diff changeset
     7
lemmas rtrancl_induct = rtrancl_induct [case_names initial step, induct set: rtrancl]
9032bdbc2125 new-style theory;
wenzelm
parents: 9578
diff changeset
     8
  and trancl_induct = trancl_induct [case_names initial step, induct set: trancl]
9032bdbc2125 new-style theory;
wenzelm
parents: 9578
diff changeset
     9
  and converse_trancl_induct = converse_trancl_induct [case_names initial step, consumes 1]
9032bdbc2125 new-style theory;
wenzelm
parents: 9578
diff changeset
    10
  and rtrancl_full_induct = rtrancl_full_induct [case_names initial step, consumes 1]
9032bdbc2125 new-style theory;
wenzelm
parents: 9578
diff changeset
    11
9032bdbc2125 new-style theory;
wenzelm
parents: 9578
diff changeset
    12
(* belongs to theory WF *)
9032bdbc2125 new-style theory;
wenzelm
parents: 9578
diff changeset
    13
lemmas wf_induct = wf_induct [induct set: wf]
9032bdbc2125 new-style theory;
wenzelm
parents: 9578
diff changeset
    14
  and wf_induct_rule = wf_induct [rule_format, induct set: wf]
9032bdbc2125 new-style theory;
wenzelm
parents: 9578
diff changeset
    15
  and wf_on_induct = wf_on_induct [consumes 2, induct set: wf_on]
9032bdbc2125 new-style theory;
wenzelm
parents: 9578
diff changeset
    16
  and wf_on_induct_rule = wf_on_induct [rule_format, consumes 2, induct set: wf_on]
9032bdbc2125 new-style theory;
wenzelm
parents: 9578
diff changeset
    17
9032bdbc2125 new-style theory;
wenzelm
parents: 9578
diff changeset
    18
(* belongs to theory Ordinal *)
12667
7e6eaaa125f2 Added some simprules proofs.
paulson
parents: 12620
diff changeset
    19
declare Ord_Least [intro,simp,TC]
12426
9032bdbc2125 new-style theory;
wenzelm
parents: 9578
diff changeset
    20
lemmas Ord_induct = Ord_induct [consumes 2]
9032bdbc2125 new-style theory;
wenzelm
parents: 9578
diff changeset
    21
  and Ord_induct_rule = Ord_induct [rule_format, consumes 2]
9032bdbc2125 new-style theory;
wenzelm
parents: 9578
diff changeset
    22
  and trans_induct = trans_induct [consumes 1]
9032bdbc2125 new-style theory;
wenzelm
parents: 9578
diff changeset
    23
  and trans_induct_rule = trans_induct [rule_format, consumes 1]
9032bdbc2125 new-style theory;
wenzelm
parents: 9578
diff changeset
    24
  and trans_induct3 = trans_induct3 [case_names 0 succ limit, consumes 1]
9032bdbc2125 new-style theory;
wenzelm
parents: 9578
diff changeset
    25
  and trans_induct3_rule = trans_induct3 [rule_format, case_names 0 succ limit, consumes 1]
9175
6f8499d86d4f finally theory Bin (the integers) is included
paulson
parents: 5531
diff changeset
    26
12426
9032bdbc2125 new-style theory;
wenzelm
parents: 9578
diff changeset
    27
(* belongs to theory Nat *)
9032bdbc2125 new-style theory;
wenzelm
parents: 9578
diff changeset
    28
lemmas nat_induct = nat_induct [case_names 0 succ, induct set: nat]
9032bdbc2125 new-style theory;
wenzelm
parents: 9578
diff changeset
    29
  and complete_induct = complete_induct [case_names less, consumes 1]
9032bdbc2125 new-style theory;
wenzelm
parents: 9578
diff changeset
    30
  and complete_induct_rule = complete_induct [rule_format, case_names less, consumes 1]
9032bdbc2125 new-style theory;
wenzelm
parents: 9578
diff changeset
    31
  and diff_induct = diff_induct [case_names 0 0_succ succ_succ, consumes 2]
9032bdbc2125 new-style theory;
wenzelm
parents: 9578
diff changeset
    32
9032bdbc2125 new-style theory;
wenzelm
parents: 9578
diff changeset
    33
(* belongs to theory Epsilon *)
9032bdbc2125 new-style theory;
wenzelm
parents: 9578
diff changeset
    34
lemmas eclose_induct = eclose_induct [induct set: eclose]
9032bdbc2125 new-style theory;
wenzelm
parents: 9578
diff changeset
    35
  and eclose_induct_down = eclose_induct_down [consumes 1]
9032bdbc2125 new-style theory;
wenzelm
parents: 9578
diff changeset
    36
9032bdbc2125 new-style theory;
wenzelm
parents: 9578
diff changeset
    37
(* belongs to theory Finite *)
9032bdbc2125 new-style theory;
wenzelm
parents: 9578
diff changeset
    38
lemmas Fin_induct = Fin_induct [case_names 0 cons, induct set: Fin]
9032bdbc2125 new-style theory;
wenzelm
parents: 9578
diff changeset
    39
9032bdbc2125 new-style theory;
wenzelm
parents: 9578
diff changeset
    40
(* belongs to theory CardinalArith *)
9032bdbc2125 new-style theory;
wenzelm
parents: 9578
diff changeset
    41
lemmas Finite_induct = Finite_induct [case_names 0 cons, induct set: Finite]
9032bdbc2125 new-style theory;
wenzelm
parents: 9578
diff changeset
    42
9032bdbc2125 new-style theory;
wenzelm
parents: 9578
diff changeset
    43
(* belongs to theory List *)
9032bdbc2125 new-style theory;
wenzelm
parents: 9578
diff changeset
    44
lemmas list_append_induct = list_append_induct [case_names Nil snoc, consumes 1]
9032bdbc2125 new-style theory;
wenzelm
parents: 9578
diff changeset
    45
9032bdbc2125 new-style theory;
wenzelm
parents: 9578
diff changeset
    46
(* belongs to theory IntDiv *)
9032bdbc2125 new-style theory;
wenzelm
parents: 9578
diff changeset
    47
lemmas posDivAlg_induct = posDivAlg_induct [consumes 2]
9032bdbc2125 new-style theory;
wenzelm
parents: 9578
diff changeset
    48
  and negDivAlg_induct = negDivAlg_induct [consumes 2]
9032bdbc2125 new-style theory;
wenzelm
parents: 9578
diff changeset
    49
12620
4e6626725e21 Some new theorems for ordinals
paulson
parents: 12552
diff changeset
    50
4e6626725e21 Some new theorems for ordinals
paulson
parents: 12552
diff changeset
    51
(* belongs to theory Epsilon *)
4e6626725e21 Some new theorems for ordinals
paulson
parents: 12552
diff changeset
    52
4e6626725e21 Some new theorems for ordinals
paulson
parents: 12552
diff changeset
    53
lemma def_transrec2:
4e6626725e21 Some new theorems for ordinals
paulson
parents: 12552
diff changeset
    54
     "(!!x. f(x)==transrec2(x,a,b))
4e6626725e21 Some new theorems for ordinals
paulson
parents: 12552
diff changeset
    55
      ==> f(0) = a & 
4e6626725e21 Some new theorems for ordinals
paulson
parents: 12552
diff changeset
    56
          f(succ(i)) = b(i, f(i)) & 
4e6626725e21 Some new theorems for ordinals
paulson
parents: 12552
diff changeset
    57
          (Limit(K) --> f(K) = (UN j<K. f(j)))"
4e6626725e21 Some new theorems for ordinals
paulson
parents: 12552
diff changeset
    58
by (simp add: transrec2_Limit)
4e6626725e21 Some new theorems for ordinals
paulson
parents: 12552
diff changeset
    59
12426
9032bdbc2125 new-style theory;
wenzelm
parents: 9578
diff changeset
    60
end