src/ZF/Order.ML
author lcp
Thu, 04 Aug 1994 11:45:59 +0200
changeset 507 a00301e9e64b
parent 484 70b789956bd3
child 760 f0200e91b272
permissions -rw-r--r--
addition of show_brackets
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
     1
(*  Title: 	ZF/Order.ML
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
     2
    ID:         $Id$
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
     3
    Author: 	Lawrence C Paulson, Cambridge University Computer Laboratory
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
     4
    Copyright   1994  University of Cambridge
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
     5
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
     6
For Order.thy.  Orders in Zermelo-Fraenkel Set Theory 
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
     7
*)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
     8
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
     9
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    10
open Order;
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    11
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    12
val bij_apply_cs = ZF_cs addSEs [bij_converse_bij]
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    13
                         addIs  [bij_is_fun, apply_type];
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    14
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    15
val bij_inverse_ss = 
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    16
    ZF_ss addsimps [bij_is_fun RS apply_type,
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    17
		    bij_converse_bij RS bij_is_fun RS apply_type,
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    18
		    left_inverse_bij, right_inverse_bij];
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    19
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    20
(** Basic properties of the definitions **)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    21
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    22
(*needed????*)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    23
goalw Order.thy [part_ord_def, irrefl_def, trans_on_def, asym_def]
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    24
    "!!r. part_ord(A,r) ==> asym(r Int A*A)";
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    25
by (fast_tac ZF_cs 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    26
val part_ord_Imp_asym = result();
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    27
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    28
(** Subset properties for the various forms of relation **)
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    29
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    30
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    31
(*Note: a relation s such that s<=r need not be a partial ordering*)
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    32
goalw Order.thy [part_ord_def, irrefl_def, trans_on_def]
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    33
    "!!A B r. [| part_ord(A,r);  B<=A |] ==> part_ord(B,r)";
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    34
by (fast_tac ZF_cs 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    35
val part_ord_subset = result();
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    36
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    37
goalw Order.thy [linear_def]
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    38
    "!!A B r. [| linear(A,r);  B<=A |] ==> linear(B,r)";
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    39
by (fast_tac ZF_cs 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    40
val linear_subset = result();
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    41
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    42
goalw Order.thy [tot_ord_def]
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    43
    "!!A B r. [| tot_ord(A,r);  B<=A |] ==> tot_ord(B,r)";
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    44
by (fast_tac (ZF_cs addSEs [part_ord_subset, linear_subset]) 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    45
val tot_ord_subset = result();
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    46
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    47
goalw Order.thy [well_ord_def]
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    48
    "!!A B r. [| well_ord(A,r);  B<=A |] ==> well_ord(B,r)";
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    49
by (fast_tac (ZF_cs addSEs [tot_ord_subset, wf_on_subset_A]) 1);
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    50
val well_ord_subset = result();
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    51
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    52
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    53
(** Order-isomorphisms **)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    54
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    55
goalw Order.thy [ord_iso_def] 
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    56
    "!!f. f: ord_iso(A,r,B,s) ==> f: bij(A,B)";
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    57
by (etac CollectD1 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    58
val ord_iso_is_bij = result();
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    59
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    60
goalw Order.thy [ord_iso_def] 
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    61
    "!!f. [| f: ord_iso(A,r,B,s);  <x,y>: r;  x:A;  y:A |] ==> \
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    62
\         <f`x, f`y> : s";
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    63
by (fast_tac ZF_cs 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    64
val ord_iso_apply = result();
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    65
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    66
goalw Order.thy [ord_iso_def] 
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    67
    "!!f. [| f: ord_iso(A,r,B,s);  <x,y>: s;  x:B;  y:B |] ==> \
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    68
\         <converse(f) ` x, converse(f) ` y> : r";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
    69
by (etac CollectE 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
    70
by (etac (bspec RS bspec RS iffD2) 1);
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    71
by (REPEAT (eresolve_tac [asm_rl, 
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    72
			  bij_converse_bij RS bij_is_fun RS apply_type] 1));
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    73
by (asm_simp_tac (ZF_ss addsimps [right_inverse_bij]) 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    74
val ord_iso_converse = result();
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    75
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    76
val major::premx::premy::prems = goalw Order.thy [linear_def]
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    77
    "[| linear(A,r);  x:A;  y:A;  \
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    78
\       <x,y>:r ==> P;  x=y ==> P;  <y,x>:r ==> P |] ==> P";
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    79
by (cut_facts_tac [major,premx,premy] 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    80
by (REPEAT_FIRST (eresolve_tac [ballE,disjE]));
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    81
by (EVERY1 (map etac prems));
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    82
by (ALLGOALS contr_tac);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    83
val linearE = result();
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    84
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    85
(*Does the case analysis, deleting linear(A,r) and proving trivial subgoals*)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    86
val linear_case_tac =
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    87
    SELECT_GOAL (EVERY [etac linearE 1, assume_tac 1, assume_tac 1,
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    88
			REPEAT_SOME assume_tac]);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    89
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    90
(*Inductive argument for proving Kunen's Lemma 6.1*)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    91
goalw Order.thy [well_ord_def, tot_ord_def, part_ord_def, pred_def]
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    92
  "!!r. [| well_ord(A,r);  x: A;  f: ord_iso(A, r, pred(A,x,r), r);  y: A |] \
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    93
\       ==> f`y = y";
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    94
by (safe_tac ZF_cs);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    95
by (wf_on_ind_tac "y" [] 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    96
by (forward_tac [ord_iso_is_bij] 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    97
by (subgoal_tac "f`y1 : {y: A . <y, x> : r}" 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    98
by (fast_tac (ZF_cs addSEs [bij_is_fun RS apply_type]) 2);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    99
(*Now we know f`y1 : A  and  <f`y1, x> : r  *)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   100
by (etac CollectE 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   101
by (linear_case_tac 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   102
(*Case  <f`y1, y1> : r *)   (*use induction hyp*)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   103
by (dtac bspec 1 THEN mp_tac 2 THEN assume_tac 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   104
by (dres_inst_tac [("t", "op `(converse(f))")] subst_context 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   105
by (asm_full_simp_tac (ZF_ss addsimps [left_inverse_bij]) 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   106
(*The case  <y1, f`y1> : r  *)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   107
by (subgoal_tac "<y1,x> : r" 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   108
by (fast_tac (ZF_cs addSEs [trans_onD]) 2);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   109
by (dtac ord_iso_converse 1 THEN assume_tac 1 THEN fast_tac ZF_cs 2);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   110
by (fast_tac ZF_cs 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   111
by (asm_full_simp_tac (ZF_ss addsimps [left_inverse_bij]) 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   112
(*now use induction hyp*)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   113
by (dtac bspec 1 THEN mp_tac 2 THEN fast_tac bij_apply_cs 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   114
by (dres_inst_tac [("t", "op `(f)")] subst_context 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   115
by (asm_full_simp_tac (ZF_ss addsimps [right_inverse_bij]) 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   116
val not_well_ord_iso_pred_lemma = result();
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   117
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   118
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   119
(*Kunen's Lemma 6.1: there's no order-isomorphism to an initial segment
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   120
                     of a well-ordering*)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   121
goal Order.thy
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   122
    "!!r. [| well_ord(A,r);  x:A |] ==> \
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   123
\         ALL f. f ~: ord_iso(A, r, pred(A,x,r), r)";
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   124
by (safe_tac ZF_cs);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   125
by (metacut_tac not_well_ord_iso_pred_lemma 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   126
by (REPEAT_SOME assume_tac);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   127
(*Now we know f`x = x*)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   128
by (EVERY1 [dtac (ord_iso_is_bij RS bij_is_fun RS apply_type),
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   129
	     assume_tac]);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   130
(*Now we know f`x : pred(A,x,r) *)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   131
by (asm_full_simp_tac (ZF_ss addsimps [well_ord_def, pred_def]) 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   132
by (fast_tac (ZF_cs addSEs [wf_on_not_refl RS notE]) 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   133
val not_well_ord_iso_pred = result();
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   134
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   135
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   136
(*Inductive argument for proving Kunen's Lemma 6.2*)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   137
goalw Order.thy [well_ord_def, tot_ord_def, part_ord_def]
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   138
    "!!r. [| well_ord(A,r);  well_ord(B,s);  \
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   139
\            f: ord_iso(A,r, B,s);  g: ord_iso(A,r, B,s);  y: A |] \
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   140
\         ==> f`y = g`y";
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   141
by (safe_tac ZF_cs);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   142
by (wf_on_ind_tac "y" [] 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   143
by (subgoals_tac ["f: bij(A,B)", "g: bij(A,B)", "f`y1 : B", "g`y1 : B"] 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   144
by (REPEAT (fast_tac (bij_apply_cs addSEs [ord_iso_is_bij]) 2));
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   145
by (linear_case_tac 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   146
(*The case  <f`y1, g`y1> : s  *)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   147
by (forw_inst_tac [("f","g")] ord_iso_converse 1 THEN REPEAT (assume_tac 1));
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   148
by (asm_full_simp_tac (ZF_ss addsimps [left_inverse_bij]) 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   149
by (dtac bspec 1 THEN mp_tac 2 THEN fast_tac bij_apply_cs 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   150
by (dres_inst_tac [("t", "op `(converse(f))")] subst_context 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   151
by (dres_inst_tac [("t", "op `(g)")] subst_context 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   152
by (asm_full_simp_tac bij_inverse_ss 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   153
(*The case  <g`y1, f`y1> : s  *)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   154
by (forw_inst_tac [("f","f")] ord_iso_converse 1 THEN REPEAT (assume_tac 1));
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   155
by (asm_full_simp_tac (ZF_ss addsimps [left_inverse_bij]) 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   156
by (dtac bspec 1 THEN mp_tac 2 THEN fast_tac bij_apply_cs 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   157
by (dres_inst_tac [("t", "op `(converse(g))")] subst_context 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   158
by (dres_inst_tac [("t", "op `(f)")] subst_context 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   159
by (asm_full_simp_tac bij_inverse_ss 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   160
val well_ord_iso_unique_lemma = result();
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   161
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   162
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   163
(*Kunen's Lemma 6.2: Order-isomorphisms between well-orderings are unique*)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   164
goal Order.thy
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   165
    "!!r. [| well_ord(A,r);  well_ord(B,s);  \
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   166
\            f: ord_iso(A,r, B,s);  g: ord_iso(A,r, B,s) |] ==> f = g";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   167
by (rtac fun_extension 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   168
by (etac (ord_iso_is_bij RS bij_is_fun) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   169
by (etac (ord_iso_is_bij RS bij_is_fun) 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   170
by (rtac well_ord_iso_unique_lemma 1);
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   171
by (REPEAT_SOME assume_tac);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   172
val well_ord_iso_unique = result();
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   173
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   174
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   175
goalw Order.thy [irrefl_def, part_ord_def, tot_ord_def, 
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   176
		 trans_on_def, well_ord_def]
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   177
    "!!r. [| wf[A](r); linear(A,r) |] ==> well_ord(A,r)";
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   178
by (asm_simp_tac (ZF_ss addsimps [wf_on_not_refl]) 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   179
by (safe_tac ZF_cs);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   180
by (linear_case_tac 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   181
(*case x=xb*)
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   182
by (fast_tac (ZF_cs addSEs [wf_on_asym]) 1);
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   183
(*case x<xb*)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   184
by (fast_tac (ZF_cs addSEs [wf_on_chain3]) 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   185
val well_ordI = result();
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   186
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   187
goalw Order.thy [well_ord_def]
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   188
    "!!r. well_ord(A,r) ==> wf[A](r)";
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   189
by (safe_tac ZF_cs);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   190
val well_ord_is_wf = result();
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   191
467
92868dab2939 new cardinal arithmetic developments
lcp
parents: 437
diff changeset
   192
goalw Order.thy [well_ord_def, tot_ord_def, part_ord_def]
92868dab2939 new cardinal arithmetic developments
lcp
parents: 437
diff changeset
   193
    "!!r. well_ord(A,r) ==> trans[A](r)";
92868dab2939 new cardinal arithmetic developments
lcp
parents: 437
diff changeset
   194
by (safe_tac ZF_cs);
92868dab2939 new cardinal arithmetic developments
lcp
parents: 437
diff changeset
   195
val well_ord_is_trans_on = result();
92868dab2939 new cardinal arithmetic developments
lcp
parents: 437
diff changeset
   196
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   197
(*** Derived rules for pred(A,x,r) ***)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   198
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   199
val [major,minor] = goalw Order.thy [pred_def]
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   200
    "[| y: pred(A,x,r);  [| y:A; <y,x>:r |] ==> P |] ==> P";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   201
by (rtac (major RS CollectE) 1);
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   202
by (REPEAT (ares_tac [minor] 1));
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   203
val predE = result();
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   204
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   205
goalw Order.thy [pred_def] "pred(A,x,r) <= r -`` {x}";
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   206
by (fast_tac ZF_cs 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   207
val pred_subset_under = result();
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   208
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   209
goalw Order.thy [pred_def] "pred(A,x,r) <= A";
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   210
by (fast_tac ZF_cs 1);
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   211
val pred_subset = result();
467
92868dab2939 new cardinal arithmetic developments
lcp
parents: 437
diff changeset
   212
92868dab2939 new cardinal arithmetic developments
lcp
parents: 437
diff changeset
   213
goalw Order.thy [pred_def] "y : pred(A,x,r) <-> <y,x>:r & y:A";
92868dab2939 new cardinal arithmetic developments
lcp
parents: 437
diff changeset
   214
by (fast_tac ZF_cs 1);
92868dab2939 new cardinal arithmetic developments
lcp
parents: 437
diff changeset
   215
val pred_iff = result();
92868dab2939 new cardinal arithmetic developments
lcp
parents: 437
diff changeset
   216
92868dab2939 new cardinal arithmetic developments
lcp
parents: 437
diff changeset
   217
goalw Order.thy [pred_def]
92868dab2939 new cardinal arithmetic developments
lcp
parents: 437
diff changeset
   218
    "pred(pred(A,x,r), y, r) = pred(A,x,r) Int pred(A,y,r)";
92868dab2939 new cardinal arithmetic developments
lcp
parents: 437
diff changeset
   219
by (fast_tac eq_cs 1);
92868dab2939 new cardinal arithmetic developments
lcp
parents: 437
diff changeset
   220
val pred_pred_eq = result();