src/ZF/Order.ML
author paulson
Wed, 09 Apr 1997 12:37:44 +0200
changeset 2925 b0ae2e13db93
parent 2637 e9b203f854ae
child 3207 fe79ad367d77
permissions -rw-r--r--
Using Blast_tac
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
     1
(*  Title:      ZF/Order.ML
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
     2
    ID:         $Id$
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
435
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
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
     6
Orders in Zermelo-Fraenkel Set Theory 
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
     7
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
     8
Results from the book "Set Theory: an Introduction to Independence Proofs"
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
     9
        by Ken Kunen.  Chapter 1, section 6.
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    10
*)
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
open Order;
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    13
2637
e9b203f854ae reflecting my recent changes of the simplifier and classical reasoner
oheimb
parents: 2493
diff changeset
    14
val op addss = op unsafe_addss;
e9b203f854ae reflecting my recent changes of the simplifier and classical reasoner
oheimb
parents: 2493
diff changeset
    15
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    16
(** Basic properties of the definitions **)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    17
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    18
(*needed?*)
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    19
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
    20
    "!!r. part_ord(A,r) ==> asym(r Int A*A)";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
    21
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 484
diff changeset
    22
qed "part_ord_Imp_asym";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
    23
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    24
val major::premx::premy::prems = goalw Order.thy [linear_def]
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    25
    "[| linear(A,r);  x:A;  y:A;  \
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    26
\       <x,y>:r ==> P;  x=y ==> P;  <y,x>:r ==> P |] ==> P";
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    27
by (cut_facts_tac [major,premx,premy] 1);
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    28
by (REPEAT_FIRST (eresolve_tac [ballE,disjE]));
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    29
by (EVERY1 (map etac prems));
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    30
by (ALLGOALS contr_tac);
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    31
qed "linearE";
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    32
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    33
(*Does the case analysis, deleting linear(A,r) and proving trivial subgoals*)
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    34
val linear_case_tac =
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    35
    SELECT_GOAL (EVERY [etac linearE 1, assume_tac 1, assume_tac 1,
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
    36
                        REPEAT_SOME (assume_tac ORELSE' contr_tac)]);
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    37
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    38
(** General properties of well_ord **)
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    39
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    40
goalw Order.thy [irrefl_def, part_ord_def, tot_ord_def, 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
    41
                 trans_on_def, well_ord_def]
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    42
    "!!r. [| wf[A](r); linear(A,r) |] ==> well_ord(A,r)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
    43
by (asm_simp_tac (!simpset addsimps [wf_on_not_refl]) 1);
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
    44
by (fast_tac (!claset addEs [linearE, wf_on_asym, wf_on_chain3]) 1);
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    45
qed "well_ordI";
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    46
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    47
goalw Order.thy [well_ord_def]
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    48
    "!!r. well_ord(A,r) ==> wf[A](r)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
    49
by (safe_tac (!claset));
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    50
qed "well_ord_is_wf";
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    51
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    52
goalw Order.thy [well_ord_def, tot_ord_def, part_ord_def]
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    53
    "!!r. well_ord(A,r) ==> trans[A](r)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
    54
by (safe_tac (!claset));
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    55
qed "well_ord_is_trans_on";
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    56
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    57
goalw Order.thy [well_ord_def, tot_ord_def]
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    58
  "!!r. well_ord(A,r) ==> linear(A,r)";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
    59
by (Blast_tac 1);
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    60
qed "well_ord_is_linear";
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    61
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    62
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    63
(** Derived rules for pred(A,x,r) **)
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    64
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    65
goalw Order.thy [pred_def] "y : pred(A,x,r) <-> <y,x>:r & y:A";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
    66
by (Blast_tac 1);
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    67
qed "pred_iff";
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    68
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    69
bind_thm ("predI", conjI RS (pred_iff RS iffD2));
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    70
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    71
val [major,minor] = goalw Order.thy [pred_def]
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    72
    "[| y: pred(A,x,r);  [| y:A; <y,x>:r |] ==> P |] ==> P";
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    73
by (rtac (major RS CollectE) 1);
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    74
by (REPEAT (ares_tac [minor] 1));
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    75
qed "predE";
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    76
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    77
goalw Order.thy [pred_def] "pred(A,x,r) <= r -`` {x}";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
    78
by (Blast_tac 1);
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    79
qed "pred_subset_under";
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    80
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    81
goalw Order.thy [pred_def] "pred(A,x,r) <= A";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
    82
by (Blast_tac 1);
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    83
qed "pred_subset";
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    84
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    85
goalw Order.thy [pred_def]
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    86
    "pred(pred(A,x,r), y, r) = pred(A,x,r) Int pred(A,y,r)";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
    87
by (Blast_tac 1);
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    88
qed "pred_pred_eq";
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    89
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    90
goalw Order.thy [trans_on_def, pred_def]
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    91
    "!!r. [| trans[A](r);  <y,x>:r;  x:A;  y:A \
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    92
\         |] ==> pred(pred(A,x,r), y, r) = pred(A,y,r)";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
    93
by (Blast_tac 1);
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    94
qed "trans_pred_pred_eq";
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    95
836
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
    96
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    97
(** The ordering's properties hold over all subsets of its domain 
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
    98
    [including initial segments of the form pred(A,x,r) **)
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
    99
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   100
(*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
   101
goalw Order.thy [part_ord_def, irrefl_def, trans_on_def]
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   102
    "!!A B r. [| part_ord(A,r);  B<=A |] ==> part_ord(B,r)";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   103
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 484
diff changeset
   104
qed "part_ord_subset";
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   105
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   106
goalw Order.thy [linear_def]
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   107
    "!!A B r. [| linear(A,r);  B<=A |] ==> linear(B,r)";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   108
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 484
diff changeset
   109
qed "linear_subset";
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   110
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   111
goalw Order.thy [tot_ord_def]
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   112
    "!!A B r. [| tot_ord(A,r);  B<=A |] ==> tot_ord(B,r)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   113
by (fast_tac (!claset addSEs [part_ord_subset, linear_subset]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 484
diff changeset
   114
qed "tot_ord_subset";
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   115
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   116
goalw Order.thy [well_ord_def]
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   117
    "!!A B r. [| well_ord(A,r);  B<=A |] ==> well_ord(B,r)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   118
by (fast_tac (!claset addSEs [tot_ord_subset, wf_on_subset_A]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 484
diff changeset
   119
qed "well_ord_subset";
484
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   120
70b789956bd3 Axiom of choice, cardinality results, etc.
lcp
parents: 467
diff changeset
   121
836
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   122
(** Relations restricted to a smaller domain, by Krzysztof Grabczewski **)
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   123
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   124
goalw Order.thy [irrefl_def] "irrefl(A,r Int A*A) <-> irrefl(A,r)";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   125
by (Blast_tac 1);
836
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   126
qed "irrefl_Int_iff";
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   127
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   128
goalw Order.thy [trans_on_def] "trans[A](r Int A*A) <-> trans[A](r)";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   129
by (Blast_tac 1);
836
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   130
qed "trans_on_Int_iff";
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   131
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   132
goalw Order.thy [part_ord_def] "part_ord(A,r Int A*A) <-> part_ord(A,r)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   133
by (simp_tac (!simpset addsimps [irrefl_Int_iff, trans_on_Int_iff]) 1);
836
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   134
qed "part_ord_Int_iff";
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   135
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   136
goalw Order.thy [linear_def] "linear(A,r Int A*A) <-> linear(A,r)";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   137
by (Blast_tac 1);
836
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   138
qed "linear_Int_iff";
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   139
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   140
goalw Order.thy [tot_ord_def] "tot_ord(A,r Int A*A) <-> tot_ord(A,r)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   141
by (simp_tac (!simpset addsimps [part_ord_Int_iff, linear_Int_iff]) 1);
836
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   142
qed "tot_ord_Int_iff";
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   143
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   144
goalw Order.thy [wf_on_def, wf_def] "wf[A](r Int A*A) <-> wf[A](r)";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   145
by (Blast_tac 1);
836
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   146
qed "wf_on_Int_iff";
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   147
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   148
goalw Order.thy [well_ord_def] "well_ord(A,r Int A*A) <-> well_ord(A,r)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   149
by (simp_tac (!simpset addsimps [tot_ord_Int_iff, wf_on_Int_iff]) 1);
836
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   150
qed "well_ord_Int_iff";
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   151
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   152
848
b1dc15d86081 Proved ord_isoI, ord_iso_refl. Simplified proof of
lcp
parents: 836
diff changeset
   153
(** Relations over the Empty Set **)
b1dc15d86081 Proved ord_isoI, ord_iso_refl. Simplified proof of
lcp
parents: 836
diff changeset
   154
b1dc15d86081 Proved ord_isoI, ord_iso_refl. Simplified proof of
lcp
parents: 836
diff changeset
   155
goalw Order.thy [irrefl_def] "irrefl(0,r)";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   156
by (Blast_tac 1);
848
b1dc15d86081 Proved ord_isoI, ord_iso_refl. Simplified proof of
lcp
parents: 836
diff changeset
   157
qed "irrefl_0";
b1dc15d86081 Proved ord_isoI, ord_iso_refl. Simplified proof of
lcp
parents: 836
diff changeset
   158
b1dc15d86081 Proved ord_isoI, ord_iso_refl. Simplified proof of
lcp
parents: 836
diff changeset
   159
goalw Order.thy [trans_on_def] "trans[0](r)";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   160
by (Blast_tac 1);
848
b1dc15d86081 Proved ord_isoI, ord_iso_refl. Simplified proof of
lcp
parents: 836
diff changeset
   161
qed "trans_on_0";
b1dc15d86081 Proved ord_isoI, ord_iso_refl. Simplified proof of
lcp
parents: 836
diff changeset
   162
b1dc15d86081 Proved ord_isoI, ord_iso_refl. Simplified proof of
lcp
parents: 836
diff changeset
   163
goalw Order.thy [part_ord_def] "part_ord(0,r)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   164
by (simp_tac (!simpset addsimps [irrefl_0, trans_on_0]) 1);
848
b1dc15d86081 Proved ord_isoI, ord_iso_refl. Simplified proof of
lcp
parents: 836
diff changeset
   165
qed "part_ord_0";
b1dc15d86081 Proved ord_isoI, ord_iso_refl. Simplified proof of
lcp
parents: 836
diff changeset
   166
b1dc15d86081 Proved ord_isoI, ord_iso_refl. Simplified proof of
lcp
parents: 836
diff changeset
   167
goalw Order.thy [linear_def] "linear(0,r)";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   168
by (Blast_tac 1);
848
b1dc15d86081 Proved ord_isoI, ord_iso_refl. Simplified proof of
lcp
parents: 836
diff changeset
   169
qed "linear_0";
b1dc15d86081 Proved ord_isoI, ord_iso_refl. Simplified proof of
lcp
parents: 836
diff changeset
   170
b1dc15d86081 Proved ord_isoI, ord_iso_refl. Simplified proof of
lcp
parents: 836
diff changeset
   171
goalw Order.thy [tot_ord_def] "tot_ord(0,r)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   172
by (simp_tac (!simpset addsimps [part_ord_0, linear_0]) 1);
848
b1dc15d86081 Proved ord_isoI, ord_iso_refl. Simplified proof of
lcp
parents: 836
diff changeset
   173
qed "tot_ord_0";
b1dc15d86081 Proved ord_isoI, ord_iso_refl. Simplified proof of
lcp
parents: 836
diff changeset
   174
b1dc15d86081 Proved ord_isoI, ord_iso_refl. Simplified proof of
lcp
parents: 836
diff changeset
   175
goalw Order.thy [wf_on_def, wf_def] "wf[0](r)";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   176
by (Blast_tac 1);
848
b1dc15d86081 Proved ord_isoI, ord_iso_refl. Simplified proof of
lcp
parents: 836
diff changeset
   177
qed "wf_on_0";
b1dc15d86081 Proved ord_isoI, ord_iso_refl. Simplified proof of
lcp
parents: 836
diff changeset
   178
b1dc15d86081 Proved ord_isoI, ord_iso_refl. Simplified proof of
lcp
parents: 836
diff changeset
   179
goalw Order.thy [well_ord_def] "well_ord(0,r)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   180
by (simp_tac (!simpset addsimps [tot_ord_0, wf_on_0]) 1);
848
b1dc15d86081 Proved ord_isoI, ord_iso_refl. Simplified proof of
lcp
parents: 836
diff changeset
   181
qed "well_ord_0";
836
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   182
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   183
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   184
(** Order-preserving (monotone) maps **)
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   185
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   186
goalw Order.thy [mono_map_def] 
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   187
    "!!f. f: mono_map(A,r,B,s) ==> f: A->B";
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   188
by (etac CollectD1 1);
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   189
qed "mono_map_is_fun";
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   190
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   191
goalw Order.thy [mono_map_def, inj_def] 
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   192
    "!!f. [| linear(A,r);  wf[B](s);  f: mono_map(A,r,B,s) |] ==> f: inj(A,B)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   193
by (step_tac (!claset) 1);
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   194
by (linear_case_tac 1);
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   195
by (REPEAT 
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   196
    (EVERY [eresolve_tac [wf_on_not_refl RS notE] 1,
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
   197
            etac ssubst 2,
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   198
            Fast_tac 2,
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
   199
            REPEAT (ares_tac [apply_type] 1)]));
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   200
qed "mono_map_is_inj";
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   201
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   202
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   203
(** Order-isomorphisms -- or similarities, as Suppes calls them **)
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   204
848
b1dc15d86081 Proved ord_isoI, ord_iso_refl. Simplified proof of
lcp
parents: 836
diff changeset
   205
val prems = goalw Order.thy [ord_iso_def]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
   206
    "[| f: bij(A, B);   \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
   207
\       !!x y. [| x:A; y:A |] ==> <x, y> : r <-> <f`x, f`y> : s \
848
b1dc15d86081 Proved ord_isoI, ord_iso_refl. Simplified proof of
lcp
parents: 836
diff changeset
   208
\    |] ==> f: ord_iso(A,r,B,s)";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   209
by (blast_tac (!claset addSIs prems) 1);
848
b1dc15d86081 Proved ord_isoI, ord_iso_refl. Simplified proof of
lcp
parents: 836
diff changeset
   210
qed "ord_isoI";
b1dc15d86081 Proved ord_isoI, ord_iso_refl. Simplified proof of
lcp
parents: 836
diff changeset
   211
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   212
goalw Order.thy [ord_iso_def, mono_map_def]
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   213
    "!!f. f: ord_iso(A,r,B,s) ==> f: mono_map(A,r,B,s)";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   214
by (blast_tac (!claset addSDs [bij_is_fun]) 1);
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   215
qed "ord_iso_is_mono_map";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   216
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   217
goalw Order.thy [ord_iso_def] 
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   218
    "!!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
   219
by (etac CollectD1 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 484
diff changeset
   220
qed "ord_iso_is_bij";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   221
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   222
(*Needed?  But ord_iso_converse is!*)
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   223
goalw Order.thy [ord_iso_def] 
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   224
    "!!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
   225
\         <f`x, f`y> : s";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   226
by (Blast_tac 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 484
diff changeset
   227
qed "ord_iso_apply";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   228
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   229
goalw Order.thy [ord_iso_def] 
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   230
    "!!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
   231
\         <converse(f) ` x, converse(f) ` y> : r";
437
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   232
by (etac CollectE 1);
435875e4b21d modifications for cardinal arithmetic
lcp
parents: 435
diff changeset
   233
by (etac (bspec RS bspec RS iffD2) 1);
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   234
by (REPEAT (eresolve_tac [asm_rl, 
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
   235
                          bij_converse_bij RS bij_is_fun RS apply_type] 1));
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   236
by (asm_simp_tac (!simpset addsimps [right_inverse_bij]) 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 484
diff changeset
   237
qed "ord_iso_converse";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   238
812
bf4b7c37db2c Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents: 794
diff changeset
   239
bf4b7c37db2c Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents: 794
diff changeset
   240
(*Rewriting with bijections and converse (function inverse)*)
bf4b7c37db2c Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents: 794
diff changeset
   241
val bij_inverse_ss = 
2637
e9b203f854ae reflecting my recent changes of the simplifier and classical reasoner
oheimb
parents: 2493
diff changeset
   242
    !simpset setSolver (type_auto_tac [ord_iso_is_bij, bij_is_fun, apply_type, 
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   243
                                       bij_converse_bij, comp_fun, comp_bij])
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   244
          addsimps [right_inverse_bij, left_inverse_bij];
812
bf4b7c37db2c Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents: 794
diff changeset
   245
bf4b7c37db2c Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents: 794
diff changeset
   246
(** Symmetry and Transitivity Rules **)
bf4b7c37db2c Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents: 794
diff changeset
   247
848
b1dc15d86081 Proved ord_isoI, ord_iso_refl. Simplified proof of
lcp
parents: 836
diff changeset
   248
(*Reflexivity of similarity*)
b1dc15d86081 Proved ord_isoI, ord_iso_refl. Simplified proof of
lcp
parents: 836
diff changeset
   249
goal Order.thy "id(A): ord_iso(A,r,A,r)";
b1dc15d86081 Proved ord_isoI, ord_iso_refl. Simplified proof of
lcp
parents: 836
diff changeset
   250
by (resolve_tac [id_bij RS ord_isoI] 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   251
by (Asm_simp_tac 1);
848
b1dc15d86081 Proved ord_isoI, ord_iso_refl. Simplified proof of
lcp
parents: 836
diff changeset
   252
qed "ord_iso_refl";
b1dc15d86081 Proved ord_isoI, ord_iso_refl. Simplified proof of
lcp
parents: 836
diff changeset
   253
b1dc15d86081 Proved ord_isoI, ord_iso_refl. Simplified proof of
lcp
parents: 836
diff changeset
   254
(*Symmetry of similarity*)
769
66cdfde4ec5d not_well_ord_iso_pred: removed needless quantifier
lcp
parents: 760
diff changeset
   255
goalw Order.thy [ord_iso_def] 
66cdfde4ec5d not_well_ord_iso_pred: removed needless quantifier
lcp
parents: 760
diff changeset
   256
    "!!f. f: ord_iso(A,r,B,s) ==> converse(f): ord_iso(B,s,A,r)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   257
by (fast_tac (!claset addss bij_inverse_ss) 1);
782
200a16083201 added bind_thm for theorems defined by "standard ..."
clasohm
parents: 769
diff changeset
   258
qed "ord_iso_sym";
769
66cdfde4ec5d not_well_ord_iso_pred: removed needless quantifier
lcp
parents: 760
diff changeset
   259
848
b1dc15d86081 Proved ord_isoI, ord_iso_refl. Simplified proof of
lcp
parents: 836
diff changeset
   260
(*Transitivity of similarity*)
812
bf4b7c37db2c Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents: 794
diff changeset
   261
goalw Order.thy [mono_map_def] 
bf4b7c37db2c Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents: 794
diff changeset
   262
    "!!f. [| g: mono_map(A,r,B,s);  f: mono_map(B,s,C,t) |] ==> \
bf4b7c37db2c Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents: 794
diff changeset
   263
\         (f O g): mono_map(A,r,C,t)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   264
by (fast_tac (!claset addss bij_inverse_ss) 1);
812
bf4b7c37db2c Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents: 794
diff changeset
   265
qed "mono_map_trans";
bf4b7c37db2c Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents: 794
diff changeset
   266
bf4b7c37db2c Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents: 794
diff changeset
   267
(*Transitivity of similarity: the order-isomorphism relation*)
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   268
goalw Order.thy [ord_iso_def] 
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   269
    "!!f. [| g: ord_iso(A,r,B,s);  f: ord_iso(B,s,C,t) |] ==> \
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   270
\         (f O g): ord_iso(A,r,C,t)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   271
by (fast_tac (!claset addss bij_inverse_ss) 1);
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   272
qed "ord_iso_trans";
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   273
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   274
(** Two monotone maps can make an order-isomorphism **)
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   275
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   276
goalw Order.thy [ord_iso_def, mono_map_def]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
   277
    "!!f g. [| f: mono_map(A,r,B,s);  g: mono_map(B,s,A,r);     \
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   278
\              f O g = id(B);  g O f = id(A) |] ==> f: ord_iso(A,r,B,s)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   279
by (safe_tac (!claset));
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   280
by (REPEAT_FIRST (ares_tac [fg_imp_bijective]));
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   281
by (Blast_tac 1);
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   282
by (subgoal_tac "<g`(f`x), g`(f`y)> : r" 1);
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   283
by (blast_tac (!claset addIs [apply_funtype]) 2);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   284
by (asm_full_simp_tac (!simpset addsimps [comp_eq_id_iff RS iffD1]) 1);
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   285
qed "mono_ord_isoI";
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   286
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   287
goal Order.thy
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
   288
    "!!B. [| well_ord(A,r);  well_ord(B,s);                             \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
   289
\            f: mono_map(A,r,B,s);  converse(f): mono_map(B,s,A,r)      \
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   290
\         |] ==> f: ord_iso(A,r,B,s)";
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   291
by (REPEAT (ares_tac [mono_ord_isoI] 1));
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   292
by (forward_tac [mono_map_is_fun RS fun_is_rel] 1);
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   293
by (etac (converse_converse RS subst) 1 THEN rtac left_comp_inverse 1);
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   294
by (DEPTH_SOLVE (ares_tac [mono_map_is_inj, left_comp_inverse] 1
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   295
          ORELSE eresolve_tac [well_ord_is_linear, well_ord_is_wf] 1));
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   296
qed "well_ord_mono_ord_isoI";
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   297
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   298
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   299
(** Order-isomorphisms preserve the ordering's properties **)
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   300
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   301
goalw Order.thy [part_ord_def, irrefl_def, trans_on_def, ord_iso_def]
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   302
    "!!A B r. [| part_ord(B,s);  f: ord_iso(A,r,B,s) |] ==> part_ord(A,r)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   303
by (Asm_simp_tac 1);
2493
bdeb5024353a Removal of sum_cs and eq_cs
paulson
parents: 2469
diff changeset
   304
by (fast_tac (!claset addIs [bij_is_fun RS apply_type]) 1);
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   305
qed "part_ord_ord_iso";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   306
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   307
goalw Order.thy [linear_def, ord_iso_def]
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   308
    "!!A B r. [| linear(B,s);  f: ord_iso(A,r,B,s) |] ==> linear(A,r)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   309
by (Asm_simp_tac 1);
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   310
by (safe_tac (!claset));
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   311
by (dres_inst_tac [("x1", "f`x"), ("x", "f`xa")] (bspec RS bspec) 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   312
by (safe_tac (!claset addSEs [bij_is_fun RS apply_type]));
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   313
by (dres_inst_tac [("t", "op `(converse(f))")] subst_context 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   314
by (asm_full_simp_tac (!simpset addsimps [left_inverse_bij]) 1);
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   315
qed "linear_ord_iso";
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   316
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   317
goalw Order.thy [wf_on_def, wf_def, ord_iso_def]
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   318
    "!!A B r. [| wf[B](s);  f: ord_iso(A,r,B,s) |] ==> wf[A](r)";
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   319
(*reversed &-congruence rule handles context of membership in A*)
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   320
by (asm_full_simp_tac (!simpset addcongs [conj_cong2]) 1);
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   321
by (safe_tac (!claset));
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   322
by (dres_inst_tac [("x", "{f`z. z:Z Int A}")] spec 1);
2493
bdeb5024353a Removal of sum_cs and eq_cs
paulson
parents: 2469
diff changeset
   323
by (safe_tac (!claset addSIs [equalityI]));
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   324
by (ALLGOALS (blast_tac
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   325
	      (!claset addSDs [equalityD1] addIs [bij_is_fun RS apply_type])));
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   326
qed "wf_on_ord_iso";
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   327
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   328
goalw Order.thy [well_ord_def, tot_ord_def]
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   329
    "!!A B r. [| well_ord(B,s);  f: ord_iso(A,r,B,s) |] ==> well_ord(A,r)";
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   330
by (fast_tac
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   331
    (!claset addSEs [part_ord_ord_iso, linear_ord_iso, wf_on_ord_iso]) 1);
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   332
qed "well_ord_ord_iso";
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   333
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   334
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   335
(*** Main results of Kunen, Chapter 1 section 6 ***)
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   336
812
bf4b7c37db2c Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents: 794
diff changeset
   337
(*Inductive argument for Kunen's Lemma 6.1, etc. 
bf4b7c37db2c Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents: 794
diff changeset
   338
  Simple proof from Halmos, page 72*)
bf4b7c37db2c Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents: 794
diff changeset
   339
goalw Order.thy [well_ord_def, ord_iso_def]
bf4b7c37db2c Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents: 794
diff changeset
   340
  "!!r. [| well_ord(A,r);  f: ord_iso(A,r, A',r);  A'<= A;  y: A |] \
bf4b7c37db2c Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents: 794
diff changeset
   341
\       ==> ~ <f`y, y>: r";
bf4b7c37db2c Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents: 794
diff changeset
   342
by (REPEAT (eresolve_tac [conjE, CollectE] 1));
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   343
by (wf_on_ind_tac "y" [] 1);
812
bf4b7c37db2c Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents: 794
diff changeset
   344
by (dres_inst_tac [("a","y1")] (bij_is_fun RS apply_type) 1);
bf4b7c37db2c Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents: 794
diff changeset
   345
by (assume_tac 1);
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   346
by (Blast_tac 1);
812
bf4b7c37db2c Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents: 794
diff changeset
   347
qed "well_ord_iso_subset_lemma";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   348
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   349
(*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
   350
                     of a well-ordering*)
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   351
goal Order.thy
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   352
    "!!r. [| well_ord(A,r);  f : ord_iso(A, r, pred(A,x,r), r);  x:A |] ==> P";
812
bf4b7c37db2c Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents: 794
diff changeset
   353
by (metacut_tac well_ord_iso_subset_lemma 1);
bf4b7c37db2c Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents: 794
diff changeset
   354
by (REPEAT_FIRST (ares_tac [pred_subset]));
bf4b7c37db2c Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents: 794
diff changeset
   355
(*Now we know  f`x < x *)
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   356
by (EVERY1 [dtac (ord_iso_is_bij RS bij_is_fun RS apply_type),
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
   357
             assume_tac]);
812
bf4b7c37db2c Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents: 794
diff changeset
   358
(*Now we also know f`x : pred(A,x,r);  contradiction! *)
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   359
by (asm_full_simp_tac (!simpset addsimps [well_ord_def, pred_def]) 1);
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   360
qed "well_ord_iso_predE";
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   361
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   362
(*Simple consequence of Lemma 6.1*)
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   363
goal Order.thy
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
   364
 "!!r. [| well_ord(A,r);  f : ord_iso(pred(A,a,r), r, pred(A,c,r), r);  \
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   365
\         a:A;  c:A |] ==> a=c";
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   366
by (forward_tac [well_ord_is_trans_on] 1);
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   367
by (forward_tac [well_ord_is_linear] 1);
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   368
by (linear_case_tac 1);
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   369
by (dtac ord_iso_sym 1);
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   370
by (REPEAT   (*because there are two symmetric cases*)
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   371
    (EVERY [eresolve_tac [pred_subset RSN (2, well_ord_subset) RS
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
   372
                          well_ord_iso_predE] 1,
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   373
            blast_tac (!claset addSIs [predI]) 2,
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   374
            asm_simp_tac (!simpset addsimps [trans_pred_pred_eq]) 1]));
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   375
qed "well_ord_iso_pred_eq";
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   376
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   377
(*Does not assume r is a wellordering!*)
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   378
goalw Order.thy [ord_iso_def, pred_def]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
   379
 "!!r. [| f : ord_iso(A,r,B,s);   a:A |] ==>    \
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   380
\      f `` pred(A,a,r) = pred(B, f`a, s)";
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   381
by (etac CollectE 1);
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   382
by (asm_simp_tac 
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   383
    (!simpset addsimps [[bij_is_fun, Collect_subset] MRS image_fun]) 1);
2493
bdeb5024353a Removal of sum_cs and eq_cs
paulson
parents: 2469
diff changeset
   384
by (rtac equalityI 1);
bdeb5024353a Removal of sum_cs and eq_cs
paulson
parents: 2469
diff changeset
   385
by (safe_tac (!claset addSEs [bij_is_fun RS apply_type]));
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
   386
by (rtac RepFun_eqI 1);
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   387
by (blast_tac (!claset addSIs [right_inverse_bij RS sym]) 1);
812
bf4b7c37db2c Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents: 794
diff changeset
   388
by (asm_simp_tac bij_inverse_ss 1);
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   389
qed "ord_iso_image_pred";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   390
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   391
(*But in use, A and B may themselves be initial segments.  Then use
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   392
  trans_pred_pred_eq to simplify the pred(pred...) terms.  See just below.*)
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   393
goal Order.thy
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
   394
 "!!r. [| f : ord_iso(A,r,B,s);   a:A |] ==>    \
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   395
\      restrict(f, pred(A,a,r)) : ord_iso(pred(A,a,r), r, pred(B, f`a, s), s)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   396
by (asm_simp_tac (!simpset addsimps [ord_iso_image_pred RS sym]) 1);
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
   397
by (rewtac ord_iso_def);
848
b1dc15d86081 Proved ord_isoI, ord_iso_refl. Simplified proof of
lcp
parents: 836
diff changeset
   398
by (etac CollectE 1);
b1dc15d86081 Proved ord_isoI, ord_iso_refl. Simplified proof of
lcp
parents: 836
diff changeset
   399
by (rtac CollectI 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   400
by (asm_full_simp_tac (!simpset addsimps [pred_def]) 2);
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   401
by (eresolve_tac [[bij_is_inj, pred_subset] MRS restrict_bij] 1);
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   402
qed "ord_iso_restrict_pred";
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   403
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   404
(*Tricky; a lot of forward proof!*)
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   405
goal Order.thy
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
   406
 "!!r. [| well_ord(A,r);  well_ord(B,s);  <a,c>: r;     \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
   407
\         f : ord_iso(pred(A,a,r), r, pred(B,b,s), s);  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
   408
\         g : ord_iso(pred(A,c,r), r, pred(B,d,s), s);  \
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   409
\         a:A;  c:A;  b:B;  d:B |] ==> <b,d>: s";
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   410
by (forward_tac [ord_iso_is_bij RS bij_is_fun RS apply_type] 1  THEN
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   411
    REPEAT1 (eresolve_tac [asm_rl, predI, predE] 1));
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   412
by (subgoal_tac "b = g`a" 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   413
by (Asm_simp_tac 1);
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
   414
by (rtac well_ord_iso_pred_eq 1);
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   415
by (REPEAT_SOME assume_tac);
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   416
by (forward_tac [ord_iso_restrict_pred] 1  THEN
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   417
    REPEAT1 (eresolve_tac [asm_rl, predI] 1));
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   418
by (asm_full_simp_tac
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   419
    (!simpset addsimps [well_ord_is_trans_on, trans_pred_pred_eq]) 1);
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   420
by (eresolve_tac [ord_iso_sym RS ord_iso_trans] 1);
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   421
by (assume_tac 1);
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   422
qed "well_ord_iso_preserving";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   423
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   424
val  bij_apply_cs = !claset addSIs [bij_converse_bij]
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   425
                            addIs  [ord_iso_is_bij, bij_is_fun, apply_funtype];
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   426
812
bf4b7c37db2c Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents: 794
diff changeset
   427
(*See Halmos, page 72*)
bf4b7c37db2c Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents: 794
diff changeset
   428
goal Order.thy
bf4b7c37db2c Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents: 794
diff changeset
   429
    "!!r. [| well_ord(A,r);  \
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   430
\            f: ord_iso(A,r, B,s);  g: ord_iso(A,r, B,s);  y: A |] \
812
bf4b7c37db2c Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents: 794
diff changeset
   431
\         ==> ~ <g`y, f`y> : s";
bf4b7c37db2c Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents: 794
diff changeset
   432
by (forward_tac [well_ord_iso_subset_lemma] 1);
bf4b7c37db2c Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents: 794
diff changeset
   433
by (res_inst_tac [("f","converse(f)"), ("g","g")] ord_iso_trans 1);
bf4b7c37db2c Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents: 794
diff changeset
   434
by (REPEAT_FIRST (ares_tac [subset_refl, ord_iso_sym]));
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   435
by (safe_tac (!claset));
812
bf4b7c37db2c Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents: 794
diff changeset
   436
by (forward_tac [ord_iso_converse] 1);
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   437
by (EVERY (map (blast_tac bij_apply_cs) [1,1,1]));
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   438
by (asm_full_simp_tac bij_inverse_ss 1);
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 484
diff changeset
   439
qed "well_ord_iso_unique_lemma";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   440
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   441
(*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
   442
goal Order.thy
812
bf4b7c37db2c Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents: 794
diff changeset
   443
    "!!r. [| well_ord(A,r);  \
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   444
\            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
   445
by (rtac fun_extension 1);
812
bf4b7c37db2c Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents: 794
diff changeset
   446
by (REPEAT (etac (ord_iso_is_bij RS bij_is_fun) 1));
bf4b7c37db2c Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents: 794
diff changeset
   447
by (subgoals_tac ["f`x : B", "g`x : B", "linear(B,s)"] 1);
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   448
by (REPEAT (blast_tac bij_apply_cs 3));
812
bf4b7c37db2c Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents: 794
diff changeset
   449
by (dtac well_ord_ord_iso 2 THEN etac ord_iso_sym 2);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   450
by (asm_full_simp_tac (!simpset addsimps [tot_ord_def, well_ord_def]) 2);
812
bf4b7c37db2c Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents: 794
diff changeset
   451
by (linear_case_tac 1);
bf4b7c37db2c Simplified proof of ord_iso_image_pred using bij_inverse_ss.
lcp
parents: 794
diff changeset
   452
by (DEPTH_SOLVE (eresolve_tac [asm_rl, well_ord_iso_unique_lemma RS notE] 1));
760
f0200e91b272 added qed and qed_goal[w]
clasohm
parents: 484
diff changeset
   453
qed "well_ord_iso_unique";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   454
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   455
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   456
(** Towards Kunen's Theorem 6.3: linearity of the similarity relation **)
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   457
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   458
goalw Order.thy [ord_iso_map_def] "ord_iso_map(A,r,B,s) <= A*B";
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   459
by (Blast_tac 1);
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   460
qed "ord_iso_map_subset";
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   461
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   462
goalw Order.thy [ord_iso_map_def] "domain(ord_iso_map(A,r,B,s)) <= A";
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   463
by (Blast_tac 1);
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   464
qed "domain_ord_iso_map";
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   465
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   466
goalw Order.thy [ord_iso_map_def] "range(ord_iso_map(A,r,B,s)) <= B";
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   467
by (Blast_tac 1);
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   468
qed "range_ord_iso_map";
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   469
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   470
goalw Order.thy [ord_iso_map_def]
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   471
    "converse(ord_iso_map(A,r,B,s)) = ord_iso_map(B,s,A,r)";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   472
by (blast_tac (!claset addIs [ord_iso_sym]) 1);
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   473
qed "converse_ord_iso_map";
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   474
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   475
goalw Order.thy [ord_iso_map_def, function_def]
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   476
    "!!B. well_ord(B,s) ==> function(ord_iso_map(A,r,B,s))";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   477
by (blast_tac (!claset addIs [well_ord_iso_pred_eq, 
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   478
			      ord_iso_sym, ord_iso_trans]) 1);
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   479
qed "function_ord_iso_map";
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   480
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   481
goal Order.thy
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
   482
    "!!B. well_ord(B,s) ==> ord_iso_map(A,r,B,s)        \
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   483
\          : domain(ord_iso_map(A,r,B,s)) -> range(ord_iso_map(A,r,B,s))";
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   484
by (asm_simp_tac 
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   485
    (!simpset addsimps [Pi_iff, function_ord_iso_map,
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
   486
                     ord_iso_map_subset RS domain_times_range]) 1);
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   487
qed "ord_iso_map_fun";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   488
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   489
goalw Order.thy [mono_map_def]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
   490
    "!!B. [| well_ord(A,r);  well_ord(B,s) |] ==> ord_iso_map(A,r,B,s)  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
   491
\          : mono_map(domain(ord_iso_map(A,r,B,s)), r,  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
   492
\                     range(ord_iso_map(A,r,B,s)), s)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   493
by (asm_simp_tac (!simpset addsimps [ord_iso_map_fun]) 1);
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   494
by (safe_tac (!claset));
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   495
by (subgoals_tac ["x:A", "xa:A", "y:B", "ya:B"] 1);
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   496
by (REPEAT 
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   497
    (blast_tac (!claset addSEs [ord_iso_map_subset RS subsetD RS SigmaE]) 2));
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   498
by (asm_simp_tac 
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   499
    (!simpset addsimps [ord_iso_map_fun RSN (2,apply_equality)]) 1);
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
   500
by (rewtac ord_iso_map_def);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   501
by (safe_tac (!claset addSEs [UN_E]));
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
   502
by (rtac well_ord_iso_preserving 1 THEN REPEAT_FIRST assume_tac);
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   503
qed "ord_iso_map_mono_map";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   504
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   505
goalw Order.thy [mono_map_def]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
   506
    "!!B. [| well_ord(A,r);  well_ord(B,s) |] ==> ord_iso_map(A,r,B,s)  \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
   507
\          : ord_iso(domain(ord_iso_map(A,r,B,s)), r,   \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
   508
\                     range(ord_iso_map(A,r,B,s)), s)";
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
   509
by (rtac well_ord_mono_ord_isoI 1);
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   510
by (resolve_tac [converse_ord_iso_map RS subst] 4);
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   511
by (asm_simp_tac 
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   512
    (!simpset addsimps [ord_iso_map_subset RS converse_converse]) 4);
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   513
by (REPEAT (ares_tac [ord_iso_map_mono_map] 3));
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   514
by (ALLGOALS (etac well_ord_subset));
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   515
by (ALLGOALS (resolve_tac [domain_ord_iso_map, range_ord_iso_map]));
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   516
qed "ord_iso_map_ord_iso";
467
92868dab2939 new cardinal arithmetic developments
lcp
parents: 437
diff changeset
   517
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   518
(*One way of saying that domain(ord_iso_map(A,r,B,s)) is downwards-closed*)
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   519
goalw Order.thy [ord_iso_map_def]
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
   520
  "!!B. [| well_ord(A,r);  well_ord(B,s);               \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
   521
\          a: A;  a ~: domain(ord_iso_map(A,r,B,s))     \
1015
75110179587d Changed proof of domain_ord_iso_map_subset for new hyp_subst_tac
lcp
parents: 990
diff changeset
   522
\       |] ==>  domain(ord_iso_map(A,r,B,s)) <= pred(A, a, r)";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   523
by (safe_tac (!claset addSIs [predI]));
1015
75110179587d Changed proof of domain_ord_iso_map_subset for new hyp_subst_tac
lcp
parents: 990
diff changeset
   524
(*Case analysis on  xaa vs a in r *)
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   525
by (forw_inst_tac [("A","A")] well_ord_is_linear 1);
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   526
by (linear_case_tac 1);
1015
75110179587d Changed proof of domain_ord_iso_map_subset for new hyp_subst_tac
lcp
parents: 990
diff changeset
   527
(*Trivial case: a=xa*)
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   528
by (Blast_tac 2);
1015
75110179587d Changed proof of domain_ord_iso_map_subset for new hyp_subst_tac
lcp
parents: 990
diff changeset
   529
(*Harder case: <a, xa>: r*)
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   530
by (forward_tac [ord_iso_is_bij RS bij_is_fun RS apply_type] 1  THEN
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   531
    REPEAT1 (eresolve_tac [asm_rl, predI, predE] 1));
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   532
by (forward_tac [ord_iso_restrict_pred] 1  THEN
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   533
    REPEAT1 (eresolve_tac [asm_rl, predI] 1));
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   534
by (asm_full_simp_tac
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   535
    (!simpset addsimps [well_ord_is_trans_on, trans_pred_pred_eq]) 1);
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   536
by (Blast_tac 1);
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   537
qed "domain_ord_iso_map_subset";
435
ca5356bd315a Addition of cardinals and order types, various tidying
lcp
parents:
diff changeset
   538
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   539
(*For the 4-way case analysis in the main result*)
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   540
goal Order.thy
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   541
  "!!B. [| well_ord(A,r);  well_ord(B,s) |] ==> \
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
   542
\       domain(ord_iso_map(A,r,B,s)) = A |      \
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   543
\       (EX x:A. domain(ord_iso_map(A,r,B,s)) = pred(A,x,r))";
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   544
by (forward_tac [well_ord_is_wf] 1);
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   545
by (rewrite_goals_tac [wf_on_def, wf_def]);
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   546
by (dres_inst_tac [("x", "A-domain(ord_iso_map(A,r,B,s))")] spec 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   547
by (step_tac (!claset) 1);
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   548
(*The first case: the domain equals A*)
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   549
by (rtac (domain_ord_iso_map RS equalityI) 1);
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   550
by (etac (Diff_eq_0_iff RS iffD1) 1);
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   551
(*The other case: the domain equals an initial segment*)
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   552
by (swap_res_tac [bexI] 1);
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   553
by (assume_tac 2);
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   554
by (rtac equalityI 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   555
(*not (!claset) below; that would use rules like domainE!*)
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   556
by (blast_tac (!claset addSEs [predE]) 2);
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   557
by (REPEAT (ares_tac [domain_ord_iso_map_subset] 1));
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   558
qed "domain_ord_iso_map_cases";
467
92868dab2939 new cardinal arithmetic developments
lcp
parents: 437
diff changeset
   559
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   560
(*As above, by duality*)
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   561
goal Order.thy
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
   562
  "!!B. [| well_ord(A,r);  well_ord(B,s) |] ==> \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
   563
\       range(ord_iso_map(A,r,B,s)) = B |       \
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   564
\       (EX y:B. range(ord_iso_map(A,r,B,s))= pred(B,y,s))";
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   565
by (resolve_tac [converse_ord_iso_map RS subst] 1);
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   566
by (asm_simp_tac
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   567
    (!simpset addsimps [range_converse, domain_ord_iso_map_cases]) 1);
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   568
qed "range_ord_iso_map_cases";
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   569
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   570
(*Kunen's Theorem 6.3: Fundamental Theorem for Well-Ordered Sets*)
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   571
goal Order.thy
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
   572
  "!!B. [| well_ord(A,r);  well_ord(B,s) |] ==>         \
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
   573
\       ord_iso_map(A,r,B,s) : ord_iso(A, r, B, s) |    \
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   574
\       (EX x:A. ord_iso_map(A,r,B,s) : ord_iso(pred(A,x,r), r, B, s)) | \
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   575
\       (EX y:B. ord_iso_map(A,r,B,s) : ord_iso(A, r, pred(B,y,s), s))";
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   576
by (forw_inst_tac [("B","B")] domain_ord_iso_map_cases 1);
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   577
by (forw_inst_tac [("B","B")] range_ord_iso_map_cases 2);
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   578
by (REPEAT_FIRST (eresolve_tac [asm_rl, disjE, bexE]));
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
   579
by (ALLGOALS (dtac ord_iso_map_ord_iso THEN' assume_tac THEN' 
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   580
              asm_full_simp_tac (!simpset addsimps [bexI])));
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   581
by (resolve_tac [wf_on_not_refl RS notE] 1);
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
   582
by (etac well_ord_is_wf 1);
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   583
by (assume_tac 1);
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   584
by (subgoal_tac "<x,y>: ord_iso_map(A,r,B,s)" 1);
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
   585
by (dtac rangeI 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   586
by (asm_full_simp_tac (!simpset addsimps [pred_def]) 1);
1461
6bcb44e4d6e5 expanded tabs
clasohm
parents: 1015
diff changeset
   587
by (rewtac ord_iso_map_def);
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   588
by (Blast_tac 1);
789
30bdc59198ff well_ord_iso_predE replaces not_well_ord_iso_pred
lcp
parents: 782
diff changeset
   589
qed "well_ord_trichotomy";
467
92868dab2939 new cardinal arithmetic developments
lcp
parents: 437
diff changeset
   590
836
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   591
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   592
(*** Properties of converse(r), by Krzysztof Grabczewski ***)
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   593
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   594
goalw Order.thy [irrefl_def] 
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   595
            "!!A. irrefl(A,r) ==> irrefl(A,converse(r))";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   596
by (blast_tac (!claset addSIs [converseI]) 1);
836
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   597
qed "irrefl_converse";
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   598
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   599
goalw Order.thy [trans_on_def] 
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   600
    "!!A. trans[A](r) ==> trans[A](converse(r))";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   601
by (blast_tac (!claset addSIs [converseI]) 1);
836
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   602
qed "trans_on_converse";
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   603
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   604
goalw Order.thy [part_ord_def] 
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   605
    "!!A. part_ord(A,r) ==> part_ord(A,converse(r))";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   606
by (blast_tac (!claset addSIs [irrefl_converse, trans_on_converse]) 1);
836
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   607
qed "part_ord_converse";
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   608
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   609
goalw Order.thy [linear_def] 
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   610
    "!!A. linear(A,r) ==> linear(A,converse(r))";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   611
by (blast_tac (!claset addSIs [converseI]) 1);
836
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   612
qed "linear_converse";
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   613
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   614
goalw Order.thy [tot_ord_def] 
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   615
    "!!A. tot_ord(A,r) ==> tot_ord(A,converse(r))";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   616
by (blast_tac (!claset addSIs [part_ord_converse, linear_converse]) 1);
836
627f4842b020 Added Krzysztof's theorems irrefl_converse, trans_on_converse,
lcp
parents: 812
diff changeset
   617
qed "tot_ord_converse";
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   618
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   619
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   620
(** By Krzysztof Grabczewski.
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   621
    Lemmas involving the first element of a well ordered set **)
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   622
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   623
goalw thy [first_def] "!!b. first(b,B,r) ==> b:B";
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   624
by (Blast_tac 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   625
qed "first_is_elem";
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   626
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   627
goalw thy [well_ord_def, wf_on_def, wf_def,     first_def] 
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   628
        "!!r. [| well_ord(A,r); B<=A; B~=0 |] ==> (EX! b. first(b,B,r))";
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   629
by (REPEAT (eresolve_tac [conjE,allE,disjE] 1));
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   630
by (contr_tac 1);
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   631
by (etac bexE 1);
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   632
by (res_inst_tac [("a","x")] ex1I 1);
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   633
by (Blast_tac 2);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   634
by (rewrite_goals_tac [tot_ord_def, linear_def]);
2925
b0ae2e13db93 Using Blast_tac
paulson
parents: 2637
diff changeset
   635
by (Blast.depth_tac (!claset) 7 1);
2469
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   636
qed "well_ord_imp_ex1_first";
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   637
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   638
goal thy "!!r. [| well_ord(A,r); B<=A; B~=0 |] ==> (THE b. first(b,B,r)) : B";
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   639
by (dtac well_ord_imp_ex1_first 1 THEN REPEAT (assume_tac 1));
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   640
by (rtac first_is_elem 1);
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   641
by (etac theI 1);
b50b8c0eec01 Implicit simpsets and clasets for FOL and ZF
paulson
parents: 1791
diff changeset
   642
qed "the_first_in";