src/ZF/Main_ZF.thy
author wenzelm
Sat, 10 Oct 2015 21:43:07 +0200
changeset 61391 2332d9be352b
parent 61378 3e04c9ca001a
child 61397 6204c86280ff
permissions -rw-r--r--
tuned syntax -- more symbols;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59647
diff changeset
     1
section\<open>Theory Main: Everything Except AC\<close>
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
     2
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
     3
theory Main_ZF imports List_ZF IntDiv_ZF CardinalArith begin
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
     4
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
     5
(*The theory of "iterates" logically belongs to Nat, but can't go there because
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
     6
  primrec isn't available into after Datatype.*)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
     7
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59647
diff changeset
     8
subsection\<open>Iteration of the function @{term F}\<close>
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
     9
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    10
consts  iterates :: "[i=>i,i,i] => i"   ("(_^_ '(_'))" [60,1000,1000] 60)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    11
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    12
primrec
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    13
    "F^0 (x) = x"
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    14
    "F^(succ(n)) (x) = F(F^n (x))"
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    15
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    16
definition
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    17
  iterates_omega :: "[i=>i,i] => i"  where
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    18
    "iterates_omega(F,x) == \<Union>n\<in>nat. F^n (x)"
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    19
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    20
notation (xsymbols)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    21
  iterates_omega  ("(_^\<omega> '(_'))" [60,1000] 60)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    22
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    23
lemma iterates_triv:
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46820
diff changeset
    24
     "[| n\<in>nat;  F(x) = x |] ==> F^n (x) = x"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    25
by (induct n rule: nat_induct, simp_all)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    26
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    27
lemma iterates_type [TC]:
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46820
diff changeset
    28
     "[| n \<in> nat;  a \<in> A; !!x. x \<in> A ==> F(x) \<in> A |]
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46820
diff changeset
    29
      ==> F^n (a) \<in> A"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    30
by (induct n rule: nat_induct, simp_all)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    31
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    32
lemma iterates_omega_triv:
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46820
diff changeset
    33
    "F(x) = x ==> F^\<omega> (x) = x"
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46820
diff changeset
    34
by (simp add: iterates_omega_def iterates_triv)
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    35
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    36
lemma Ord_iterates [simp]:
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46820
diff changeset
    37
     "[| n\<in>nat;  !!i. Ord(i) ==> Ord(F(i));  Ord(x) |]
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46820
diff changeset
    38
      ==> Ord(F^n (x))"
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    39
by (induct n rule: nat_induct, simp_all)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    40
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    41
lemma iterates_commute: "n \<in> nat ==> F(F^n (x)) = F^n (F(x))"
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    42
by (induct_tac n, simp_all)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    43
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    44
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59647
diff changeset
    45
subsection\<open>Transfinite Recursion\<close>
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    46
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59647
diff changeset
    47
text\<open>Transfinite recursion for definitions based on the
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59647
diff changeset
    48
    three cases of ordinals\<close>
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    49
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    50
definition
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    51
  transrec3 :: "[i, i, [i,i]=>i, [i,i]=>i] =>i" where
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46820
diff changeset
    52
    "transrec3(k, a, b, c) ==
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    53
       transrec(k, \<lambda>x r.
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    54
         if x=0 then a
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    55
         else if Limit(x) then c(x, \<lambda>y\<in>x. r`y)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    56
         else b(Arith.pred(x), r ` Arith.pred(x)))"
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    57
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    58
lemma transrec3_0 [simp]: "transrec3(0,a,b,c) = a"
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    59
by (rule transrec3_def [THEN def_transrec, THEN trans], simp)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    60
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    61
lemma transrec3_succ [simp]:
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    62
     "transrec3(succ(i),a,b,c) = b(i, transrec3(i,a,b,c))"
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    63
by (rule transrec3_def [THEN def_transrec, THEN trans], simp)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    64
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    65
lemma transrec3_Limit:
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46820
diff changeset
    66
     "Limit(i) ==>
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    67
      transrec3(i,a,b,c) = c(i, \<lambda>j\<in>i. transrec3(j,a,b,c))"
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    68
by (rule transrec3_def [THEN def_transrec, THEN trans], force)
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    69
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    70
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59647
diff changeset
    71
declaration \<open>fn _ =>
59647
c6f413b660cf clarified Drule.gen_all: observe context more carefully;
wenzelm
parents: 58871
diff changeset
    72
  Simplifier.map_ss (Simplifier.set_mksimps (fn ctxt =>
60822
4f58f3662e7d more explicit context;
wenzelm
parents: 60770
diff changeset
    73
    map mk_eq o Ord_atomize o Variable.gen_all ctxt))
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59647
diff changeset
    74
\<close>
26056
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    75
6a0801279f4c Made theory names in ZF disjoint from HOL theory names to allow loading both developments
krauss
parents:
diff changeset
    76
end