src/ZF/Constructible/Wellorderings.thy
author wenzelm
Fri, 17 Nov 2006 02:20:03 +0100
changeset 21404 eb85850d3eb7
parent 21233 5a5c8ea5f66a
child 32960 69916a850301
permissions -rw-r--r--
more robust syntax for definition/abbreviation/notation;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13505
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13428
diff changeset
     1
(*  Title:      ZF/Constructible/Wellorderings.thy
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13428
diff changeset
     2
    ID:         $Id$
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13428
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13428
diff changeset
     4
*)
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13428
diff changeset
     5
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
     6
header {*Relativized Wellorderings*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
     7
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 13780
diff changeset
     8
theory Wellorderings imports Relative begin
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
     9
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    10
text{*We define functions analogous to @{term ordermap} @{term ordertype} 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    11
      but without using recursion.  Instead, there is a direct appeal
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    12
      to Replacement.  This will be the basis for a version relativized
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    13
      to some class @{text M}.  The main result is Theorem I 7.6 in Kunen,
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    14
      page 17.*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    15
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    16
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    17
subsection{*Wellorderings*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    18
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 16417
diff changeset
    19
definition
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    20
  irreflexive :: "[i=>o,i,i]=>o" where
13299
3a932abf97e8 More use of relativized quantifiers
paulson
parents: 13298
diff changeset
    21
    "irreflexive(M,A,r) == \<forall>x[M]. x\<in>A --> <x,x> \<notin> r"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    22
  
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    23
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    24
  transitive_rel :: "[i=>o,i,i]=>o" where
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    25
    "transitive_rel(M,A,r) == 
13299
3a932abf97e8 More use of relativized quantifiers
paulson
parents: 13298
diff changeset
    26
	\<forall>x[M]. x\<in>A --> (\<forall>y[M]. y\<in>A --> (\<forall>z[M]. z\<in>A --> 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    27
                          <x,y>\<in>r --> <y,z>\<in>r --> <x,z>\<in>r))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    28
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    29
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    30
  linear_rel :: "[i=>o,i,i]=>o" where
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    31
    "linear_rel(M,A,r) == 
13299
3a932abf97e8 More use of relativized quantifiers
paulson
parents: 13298
diff changeset
    32
	\<forall>x[M]. x\<in>A --> (\<forall>y[M]. y\<in>A --> <x,y>\<in>r | x=y | <y,x>\<in>r)"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    33
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    34
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    35
  wellfounded :: "[i=>o,i]=>o" where
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    36
    --{*EVERY non-empty set has an @{text r}-minimal element*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    37
    "wellfounded(M,r) == 
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
    38
	\<forall>x[M]. x\<noteq>0 --> (\<exists>y[M]. y\<in>x & ~(\<exists>z[M]. z\<in>x & <z,y> \<in> r))"
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    39
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    40
  wellfounded_on :: "[i=>o,i,i]=>o" where
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    41
    --{*every non-empty SUBSET OF @{text A} has an @{text r}-minimal element*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    42
    "wellfounded_on(M,A,r) == 
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
    43
	\<forall>x[M]. x\<noteq>0 --> x\<subseteq>A --> (\<exists>y[M]. y\<in>x & ~(\<exists>z[M]. z\<in>x & <z,y> \<in> r))"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    44
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    45
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    46
  wellordered :: "[i=>o,i,i]=>o" where
13513
paulson
parents: 13505
diff changeset
    47
    --{*linear and wellfounded on @{text A}*}
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    48
    "wellordered(M,A,r) == 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    49
	transitive_rel(M,A,r) & linear_rel(M,A,r) & wellfounded_on(M,A,r)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    50
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    51
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    52
subsubsection {*Trivial absoluteness proofs*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    53
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13513
diff changeset
    54
lemma (in M_basic) irreflexive_abs [simp]: 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    55
     "M(A) ==> irreflexive(M,A,r) <-> irrefl(A,r)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    56
by (simp add: irreflexive_def irrefl_def)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    57
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13513
diff changeset
    58
lemma (in M_basic) transitive_rel_abs [simp]: 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    59
     "M(A) ==> transitive_rel(M,A,r) <-> trans[A](r)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    60
by (simp add: transitive_rel_def trans_on_def)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    61
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13513
diff changeset
    62
lemma (in M_basic) linear_rel_abs [simp]: 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    63
     "M(A) ==> linear_rel(M,A,r) <-> linear(A,r)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    64
by (simp add: linear_rel_def linear_def)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    65
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13513
diff changeset
    66
lemma (in M_basic) wellordered_is_trans_on: 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    67
    "[| wellordered(M,A,r); M(A) |] ==> trans[A](r)"
13505
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13428
diff changeset
    68
by (auto simp add: wellordered_def)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    69
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13513
diff changeset
    70
lemma (in M_basic) wellordered_is_linear: 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    71
    "[| wellordered(M,A,r); M(A) |] ==> linear(A,r)"
13505
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13428
diff changeset
    72
by (auto simp add: wellordered_def)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    73
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13513
diff changeset
    74
lemma (in M_basic) wellordered_is_wellfounded_on: 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    75
    "[| wellordered(M,A,r); M(A) |] ==> wellfounded_on(M,A,r)"
13505
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13428
diff changeset
    76
by (auto simp add: wellordered_def)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    77
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13513
diff changeset
    78
lemma (in M_basic) wellfounded_imp_wellfounded_on: 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    79
    "[| wellfounded(M,r); M(A) |] ==> wellfounded_on(M,A,r)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    80
by (auto simp add: wellfounded_def wellfounded_on_def)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    81
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13513
diff changeset
    82
lemma (in M_basic) wellfounded_on_subset_A:
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13251
diff changeset
    83
     "[| wellfounded_on(M,A,r);  B<=A |] ==> wellfounded_on(M,B,r)"
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13251
diff changeset
    84
by (simp add: wellfounded_on_def, blast)
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13251
diff changeset
    85
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    86
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    87
subsubsection {*Well-founded relations*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    88
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13513
diff changeset
    89
lemma  (in M_basic) wellfounded_on_iff_wellfounded:
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    90
     "wellfounded_on(M,A,r) <-> wellfounded(M, r \<inter> A*A)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    91
apply (simp add: wellfounded_on_def wellfounded_def, safe)
13780
af7b79271364 more new-style theories
paulson
parents: 13634
diff changeset
    92
 apply force
13299
3a932abf97e8 More use of relativized quantifiers
paulson
parents: 13298
diff changeset
    93
apply (drule_tac x=x in rspec, assumption, blast) 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    94
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    95
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13513
diff changeset
    96
lemma (in M_basic) wellfounded_on_imp_wellfounded:
13247
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13245
diff changeset
    97
     "[|wellfounded_on(M,A,r); r \<subseteq> A*A|] ==> wellfounded(M,r)"
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13245
diff changeset
    98
by (simp add: wellfounded_on_iff_wellfounded subset_Int_iff)
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13245
diff changeset
    99
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13513
diff changeset
   100
lemma (in M_basic) wellfounded_on_field_imp_wellfounded:
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13251
diff changeset
   101
     "wellfounded_on(M, field(r), r) ==> wellfounded(M,r)"
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13251
diff changeset
   102
by (simp add: wellfounded_def wellfounded_on_iff_wellfounded, fast)
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13251
diff changeset
   103
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13513
diff changeset
   104
lemma (in M_basic) wellfounded_iff_wellfounded_on_field:
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13251
diff changeset
   105
     "M(r) ==> wellfounded(M,r) <-> wellfounded_on(M, field(r), r)"
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13251
diff changeset
   106
by (blast intro: wellfounded_imp_wellfounded_on
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13251
diff changeset
   107
                 wellfounded_on_field_imp_wellfounded)
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13251
diff changeset
   108
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   109
(*Consider the least z in domain(r) such that P(z) does not hold...*)
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13513
diff changeset
   110
lemma (in M_basic) wellfounded_induct: 
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   111
     "[| wellfounded(M,r); M(a); M(r); separation(M, \<lambda>x. ~P(x));  
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   112
         \<forall>x. M(x) & (\<forall>y. <y,x> \<in> r --> P(y)) --> P(x) |]
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   113
      ==> P(a)";
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   114
apply (simp (no_asm_use) add: wellfounded_def)
13299
3a932abf97e8 More use of relativized quantifiers
paulson
parents: 13298
diff changeset
   115
apply (drule_tac x="{z \<in> domain(r). ~P(z)}" in rspec)
3a932abf97e8 More use of relativized quantifiers
paulson
parents: 13298
diff changeset
   116
apply (blast dest: transM)+
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   117
done
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   118
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13513
diff changeset
   119
lemma (in M_basic) wellfounded_on_induct: 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   120
     "[| a\<in>A;  wellfounded_on(M,A,r);  M(A);  
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   121
       separation(M, \<lambda>x. x\<in>A --> ~P(x));  
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   122
       \<forall>x\<in>A. M(x) & (\<forall>y\<in>A. <y,x> \<in> r --> P(y)) --> P(x) |]
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   123
      ==> P(a)";
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   124
apply (simp (no_asm_use) add: wellfounded_on_def)
13299
3a932abf97e8 More use of relativized quantifiers
paulson
parents: 13298
diff changeset
   125
apply (drule_tac x="{z\<in>A. z\<in>A --> ~P(z)}" in rspec)
3a932abf97e8 More use of relativized quantifiers
paulson
parents: 13298
diff changeset
   126
apply (blast intro: transM)+
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   127
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   128
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   129
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   130
subsubsection {*Kunen's lemma IV 3.14, page 123*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   131
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13513
diff changeset
   132
lemma (in M_basic) linear_imp_relativized: 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   133
     "linear(A,r) ==> linear_rel(M,A,r)" 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   134
by (simp add: linear_def linear_rel_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   135
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13513
diff changeset
   136
lemma (in M_basic) trans_on_imp_relativized: 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   137
     "trans[A](r) ==> transitive_rel(M,A,r)" 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   138
by (unfold transitive_rel_def trans_on_def, blast) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   139
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13513
diff changeset
   140
lemma (in M_basic) wf_on_imp_relativized: 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   141
     "wf[A](r) ==> wellfounded_on(M,A,r)" 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   142
apply (simp add: wellfounded_on_def wf_def wf_on_def, clarify) 
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13306
diff changeset
   143
apply (drule_tac x=x in spec, blast) 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   144
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   145
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13513
diff changeset
   146
lemma (in M_basic) wf_imp_relativized: 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   147
     "wf(r) ==> wellfounded(M,r)" 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   148
apply (simp add: wellfounded_def wf_def, clarify) 
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13306
diff changeset
   149
apply (drule_tac x=x in spec, blast) 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   150
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   151
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13513
diff changeset
   152
lemma (in M_basic) well_ord_imp_relativized: 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   153
     "well_ord(A,r) ==> wellordered(M,A,r)" 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   154
by (simp add: wellordered_def well_ord_def tot_ord_def part_ord_def
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   155
       linear_imp_relativized trans_on_imp_relativized wf_on_imp_relativized)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   156
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   157
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   158
subsection{* Relativized versions of order-isomorphisms and order types *}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   159
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13513
diff changeset
   160
lemma (in M_basic) order_isomorphism_abs [simp]: 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   161
     "[| M(A); M(B); M(f) |] 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   162
      ==> order_isomorphism(M,A,r,B,s,f) <-> f \<in> ord_iso(A,r,B,s)"
13352
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13339
diff changeset
   163
by (simp add: apply_closed order_isomorphism_def ord_iso_def)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   164
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13513
diff changeset
   165
lemma (in M_basic) pred_set_abs [simp]: 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   166
     "[| M(r); M(B) |] ==> pred_set(M,A,x,r,B) <-> B = Order.pred(A,x,r)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   167
apply (simp add: pred_set_def Order.pred_def)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   168
apply (blast dest: transM) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   169
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   170
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13513
diff changeset
   171
lemma (in M_basic) pred_closed [intro,simp]: 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   172
     "[| M(A); M(r); M(x) |] ==> M(Order.pred(A,x,r))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   173
apply (simp add: Order.pred_def) 
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   174
apply (insert pred_separation [of r x], simp) 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   175
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   176
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13513
diff changeset
   177
lemma (in M_basic) membership_abs [simp]: 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   178
     "[| M(r); M(A) |] ==> membership(M,A,r) <-> r = Memrel(A)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   179
apply (simp add: membership_def Memrel_def, safe)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   180
  apply (rule equalityI) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   181
   apply clarify 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   182
   apply (frule transM, assumption)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   183
   apply blast
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   184
  apply clarify 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   185
  apply (subgoal_tac "M(<xb,ya>)", blast) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   186
  apply (blast dest: transM) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   187
 apply auto 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   188
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   189
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13513
diff changeset
   190
lemma (in M_basic) M_Memrel_iff:
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   191
     "M(A) ==> 
13298
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13295
diff changeset
   192
      Memrel(A) = {z \<in> A*A. \<exists>x[M]. \<exists>y[M]. z = \<langle>x,y\<rangle> & x \<in> y}"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   193
apply (simp add: Memrel_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   194
apply (blast dest: transM)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   195
done 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   196
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13513
diff changeset
   197
lemma (in M_basic) Memrel_closed [intro,simp]: 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   198
     "M(A) ==> M(Memrel(A))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   199
apply (simp add: M_Memrel_iff) 
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   200
apply (insert Memrel_separation, simp)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   201
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   202
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   203
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   204
subsection {* Main results of Kunen, Chapter 1 section 6 *}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   205
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   206
text{*Subset properties-- proved outside the locale*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   207
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   208
lemma linear_rel_subset: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   209
    "[| linear_rel(M,A,r);  B<=A |] ==> linear_rel(M,B,r)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   210
by (unfold linear_rel_def, blast)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   211
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   212
lemma transitive_rel_subset: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   213
    "[| transitive_rel(M,A,r);  B<=A |] ==> transitive_rel(M,B,r)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   214
by (unfold transitive_rel_def, blast)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   215
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   216
lemma wellfounded_on_subset: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   217
    "[| wellfounded_on(M,A,r);  B<=A |] ==> wellfounded_on(M,B,r)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   218
by (unfold wellfounded_on_def subset_def, blast)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   219
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   220
lemma wellordered_subset: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   221
    "[| wellordered(M,A,r);  B<=A |] ==> wellordered(M,B,r)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   222
apply (unfold wellordered_def)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   223
apply (blast intro: linear_rel_subset transitive_rel_subset 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   224
		    wellfounded_on_subset)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   225
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   226
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13513
diff changeset
   227
lemma (in M_basic) wellfounded_on_asym:
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   228
     "[| wellfounded_on(M,A,r);  <a,x>\<in>r;  a\<in>A; x\<in>A;  M(A) |] ==> <x,a>\<notin>r"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   229
apply (simp add: wellfounded_on_def) 
13299
3a932abf97e8 More use of relativized quantifiers
paulson
parents: 13298
diff changeset
   230
apply (drule_tac x="{x,a}" in rspec) 
3a932abf97e8 More use of relativized quantifiers
paulson
parents: 13298
diff changeset
   231
apply (blast dest: transM)+
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   232
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   233
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13513
diff changeset
   234
lemma (in M_basic) wellordered_asym:
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   235
     "[| wellordered(M,A,r);  <a,x>\<in>r;  a\<in>A; x\<in>A;  M(A) |] ==> <x,a>\<notin>r"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   236
by (simp add: wellordered_def, blast dest: wellfounded_on_asym)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   237
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   238
end