src/HOL/BNF/BNF_LFP.thy
author bulwahn
Sat, 20 Oct 2012 09:09:34 +0200
changeset 49945 fb696ff1f086
parent 49635 fc0777f04205
child 51739 3514b90d0a8b
permissions -rw-r--r--
adjusting proofs
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
49509
163914705f8d renamed top-level theory from "Codatatype" to "BNF"
blanchet
parents: 49326
diff changeset
     1
(*  Title:      HOL/BNF/BNF_LFP.thy
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     2
    Author:     Dmitriy Traytel, TU Muenchen
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     3
    Copyright   2012
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     4
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     5
Least fixed point operation on bounded natural functors.
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     6
*)
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     7
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     8
header {* Least Fixed Point Operation on Bounded Natural Functors *}
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
     9
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    10
theory BNF_LFP
49308
6190b701e4f4 reorganized dependencies so that the sugar does not depend on GFP -- this will be essential for bootstrapping
blanchet
parents: 49074
diff changeset
    11
imports BNF_FP
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    12
keywords
49308
6190b701e4f4 reorganized dependencies so that the sugar does not depend on GFP -- this will be essential for bootstrapping
blanchet
parents: 49074
diff changeset
    13
  "data" :: thy_decl
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    14
begin
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
    15
49312
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    16
lemma subset_emptyI: "(\<And>x. x \<in> A \<Longrightarrow> False) \<Longrightarrow> A \<subseteq> {}"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    17
by blast
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    18
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    19
lemma image_Collect_subsetI:
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    20
  "(\<And>x. P x \<Longrightarrow> f x \<in> B) \<Longrightarrow> f ` {x. P x} \<subseteq> B"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    21
by blast
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    22
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    23
lemma Collect_restrict: "{x. x \<in> X \<and> P x} \<subseteq> X"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    24
by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    25
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    26
lemma prop_restrict: "\<lbrakk>x \<in> Z; Z \<subseteq> {x. x \<in> X \<and> P x}\<rbrakk> \<Longrightarrow> P x"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    27
by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    28
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    29
lemma underS_I: "\<lbrakk>i \<noteq> j; (i, j) \<in> R\<rbrakk> \<Longrightarrow> i \<in> rel.underS R j"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    30
unfolding rel.underS_def by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    31
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    32
lemma underS_E: "i \<in> rel.underS R j \<Longrightarrow> i \<noteq> j \<and> (i, j) \<in> R"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    33
unfolding rel.underS_def by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    34
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    35
lemma underS_Field: "i \<in> rel.underS R j \<Longrightarrow> i \<in> Field R"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    36
unfolding rel.underS_def Field_def by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    37
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    38
lemma FieldI2: "(i, j) \<in> R \<Longrightarrow> j \<in> Field R"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    39
unfolding Field_def by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    40
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    41
lemma fst_convol': "fst (<f, g> x) = f x"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    42
using fst_convol unfolding convol_def by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    43
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    44
lemma snd_convol': "snd (<f, g> x) = g x"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    45
using snd_convol unfolding convol_def by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    46
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    47
lemma convol_expand_snd: "fst o f = g \<Longrightarrow>  <g, snd o f> = f"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    48
unfolding convol_def by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    49
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    50
definition inver where
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    51
  "inver g f A = (ALL a : A. g (f a) = a)"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    52
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    53
lemma bij_betw_iff_ex:
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    54
  "bij_betw f A B = (EX g. g ` B = A \<and> inver g f A \<and> inver f g B)" (is "?L = ?R")
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    55
proof (rule iffI)
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    56
  assume ?L
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    57
  hence f: "f ` A = B" and inj_f: "inj_on f A" unfolding bij_betw_def by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    58
  let ?phi = "% b a. a : A \<and> f a = b"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    59
  have "ALL b : B. EX a. ?phi b a" using f by blast
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    60
  then obtain g where g: "ALL b : B. g b : A \<and> f (g b) = b"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    61
    using bchoice[of B ?phi] by blast
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    62
  hence gg: "ALL b : f ` A. g b : A \<and> f (g b) = b" using f by blast
49326
a063a96c8662 got rid of metis calls
blanchet
parents: 49312
diff changeset
    63
  have gf: "inver g f A" unfolding inver_def
a063a96c8662 got rid of metis calls
blanchet
parents: 49312
diff changeset
    64
    by (metis (no_types) gg imageI[of _ A f] the_inv_into_f_f[OF inj_f])
49312
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    65
  moreover have "g ` B \<le> A \<and> inver f g B" using g unfolding inver_def by blast
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    66
  moreover have "A \<le> g ` B"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    67
  proof safe
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    68
    fix a assume a: "a : A"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    69
    hence "f a : B" using f by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    70
    moreover have "a = g (f a)" using a gf unfolding inver_def by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    71
    ultimately show "a : g ` B" by blast
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    72
  qed
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    73
  ultimately show ?R by blast
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    74
next
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    75
  assume ?R
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    76
  then obtain g where g: "g ` B = A \<and> inver g f A \<and> inver f g B" by blast
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    77
  show ?L unfolding bij_betw_def
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    78
  proof safe
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    79
    show "inj_on f A" unfolding inj_on_def
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    80
    proof safe
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    81
      fix a1 a2 assume a: "a1 : A"  "a2 : A" and "f a1 = f a2"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    82
      hence "g (f a1) = g (f a2)" by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    83
      thus "a1 = a2" using a g unfolding inver_def by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    84
    qed
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    85
  next
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    86
    fix a assume "a : A"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    87
    then obtain b where b: "b : B" and a: "a = g b" using g by blast
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    88
    hence "b = f (g b)" using g unfolding inver_def by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    89
    thus "f a : B" unfolding a using b by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    90
  next
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    91
    fix b assume "b : B"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    92
    hence "g b : A \<and> b = f (g b)" using g unfolding inver_def by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    93
    thus "b : f ` A" by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    94
  qed
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    95
qed
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    96
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    97
lemma bij_betw_ex_weakE:
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    98
  "\<lbrakk>bij_betw f A B\<rbrakk> \<Longrightarrow> \<exists>g. g ` B \<subseteq> A \<and> inver g f A \<and> inver f g B"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    99
by (auto simp only: bij_betw_iff_ex)
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   100
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   101
lemma inver_surj: "\<lbrakk>g ` B \<subseteq> A; f ` A \<subseteq> B; inver g f A\<rbrakk> \<Longrightarrow> g ` B = A"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   102
unfolding inver_def by auto (rule rev_image_eqI, auto)
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   103
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   104
lemma inver_mono: "\<lbrakk>A \<subseteq> B; inver f g B\<rbrakk> \<Longrightarrow> inver f g A"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   105
unfolding inver_def by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   106
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   107
lemma inver_pointfree: "inver f g A = (\<forall>a \<in> A. (f o g) a = a)"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   108
unfolding inver_def by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   109
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   110
lemma bij_betwE: "bij_betw f A B \<Longrightarrow> \<forall>a\<in>A. f a \<in> B"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   111
unfolding bij_betw_def by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   112
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   113
lemma bij_betw_imageE: "bij_betw f A B \<Longrightarrow> f ` A = B"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   114
unfolding bij_betw_def by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   115
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   116
lemma inverE: "\<lbrakk>inver f f' A; x \<in> A\<rbrakk> \<Longrightarrow> f (f' x) = x"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   117
unfolding inver_def by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   118
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   119
lemma bij_betw_inver1: "bij_betw f A B \<Longrightarrow> inver (inv_into A f) f A"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   120
unfolding bij_betw_def inver_def by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   121
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   122
lemma bij_betw_inver2: "bij_betw f A B \<Longrightarrow> inver f (inv_into A f) B"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   123
unfolding bij_betw_def inver_def by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   124
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   125
lemma bij_betwI: "\<lbrakk>bij_betw g B A; inver g f A; inver f g B\<rbrakk> \<Longrightarrow> bij_betw f A B"
49326
a063a96c8662 got rid of metis calls
blanchet
parents: 49312
diff changeset
   126
by (drule bij_betw_imageE, unfold bij_betw_iff_ex) blast
49312
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   127
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   128
lemma bij_betwI':
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   129
  "\<lbrakk>\<And>x y. \<lbrakk>x \<in> X; y \<in> X\<rbrakk> \<Longrightarrow> (f x = f y) = (x = y);
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   130
    \<And>x. x \<in> X \<Longrightarrow> f x \<in> Y;
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   131
    \<And>y. y \<in> Y \<Longrightarrow> \<exists>x \<in> X. y = f x\<rbrakk> \<Longrightarrow> bij_betw f X Y"
49326
a063a96c8662 got rid of metis calls
blanchet
parents: 49312
diff changeset
   132
unfolding bij_betw_def inj_on_def
a063a96c8662 got rid of metis calls
blanchet
parents: 49312
diff changeset
   133
apply (rule conjI)
a063a96c8662 got rid of metis calls
blanchet
parents: 49312
diff changeset
   134
 apply blast
a063a96c8662 got rid of metis calls
blanchet
parents: 49312
diff changeset
   135
by (erule thin_rl) blast
49312
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   136
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   137
lemma surj_fun_eq:
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   138
  assumes surj_on: "f ` X = UNIV" and eq_on: "\<forall>x \<in> X. (g1 o f) x = (g2 o f) x"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   139
  shows "g1 = g2"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   140
proof (rule ext)
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   141
  fix y
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   142
  from surj_on obtain x where "x \<in> X" and "y = f x" by blast
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   143
  thus "g1 y = g2 y" using eq_on by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   144
qed
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   145
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   146
lemma Card_order_wo_rel: "Card_order r \<Longrightarrow> wo_rel r"
49514
45e3e564e306 tuned whitespace
blanchet
parents: 49510
diff changeset
   147
unfolding wo_rel_def card_order_on_def by blast
49312
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   148
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   149
lemma Cinfinite_limit: "\<lbrakk>x \<in> Field r; Cinfinite r\<rbrakk> \<Longrightarrow>
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   150
  \<exists>y \<in> Field r. x \<noteq> y \<and> (x, y) \<in> r"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   151
unfolding cinfinite_def by (auto simp add: infinite_Card_order_limit)
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   152
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   153
lemma Card_order_trans:
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   154
  "\<lbrakk>Card_order r; x \<noteq> y; (x, y) \<in> r; y \<noteq> z; (y, z) \<in> r\<rbrakk> \<Longrightarrow> x \<noteq> z \<and> (x, z) \<in> r"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   155
unfolding card_order_on_def well_order_on_def linear_order_on_def
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   156
  partial_order_on_def preorder_on_def trans_def antisym_def by blast
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   157
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   158
lemma Cinfinite_limit2:
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   159
 assumes x1: "x1 \<in> Field r" and x2: "x2 \<in> Field r" and r: "Cinfinite r"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   160
 shows "\<exists>y \<in> Field r. (x1 \<noteq> y \<and> (x1, y) \<in> r) \<and> (x2 \<noteq> y \<and> (x2, y) \<in> r)"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   161
proof -
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   162
  from r have trans: "trans r" and total: "Total r" and antisym: "antisym r"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   163
    unfolding card_order_on_def well_order_on_def linear_order_on_def
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   164
      partial_order_on_def preorder_on_def by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   165
  obtain y1 where y1: "y1 \<in> Field r" "x1 \<noteq> y1" "(x1, y1) \<in> r"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   166
    using Cinfinite_limit[OF x1 r] by blast
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   167
  obtain y2 where y2: "y2 \<in> Field r" "x2 \<noteq> y2" "(x2, y2) \<in> r"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   168
    using Cinfinite_limit[OF x2 r] by blast
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   169
  show ?thesis
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   170
  proof (cases "y1 = y2")
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   171
    case True with y1 y2 show ?thesis by blast
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   172
  next
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   173
    case False
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   174
    with y1(1) y2(1) total have "(y1, y2) \<in> r \<or> (y2, y1) \<in> r"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   175
      unfolding total_on_def by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   176
    thus ?thesis
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   177
    proof
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   178
      assume *: "(y1, y2) \<in> r"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   179
      with trans y1(3) have "(x1, y2) \<in> r" unfolding trans_def by blast
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   180
      with False y1 y2 * antisym show ?thesis by (cases "x1 = y2") (auto simp: antisym_def)
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   181
    next
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   182
      assume *: "(y2, y1) \<in> r"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   183
      with trans y2(3) have "(x2, y1) \<in> r" unfolding trans_def by blast
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   184
      with False y1 y2 * antisym show ?thesis by (cases "x2 = y1") (auto simp: antisym_def)
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   185
    qed
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   186
  qed
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   187
qed
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   188
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   189
lemma Cinfinite_limit_finite: "\<lbrakk>finite X; X \<subseteq> Field r; Cinfinite r\<rbrakk>
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   190
 \<Longrightarrow> \<exists>y \<in> Field r. \<forall>x \<in> X. (x \<noteq> y \<and> (x, y) \<in> r)"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   191
proof (induct X rule: finite_induct)
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   192
  case empty thus ?case unfolding cinfinite_def using ex_in_conv[of "Field r"] finite.emptyI by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   193
next
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   194
  case (insert x X)
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   195
  then obtain y where y: "y \<in> Field r" "\<forall>x \<in> X. (x \<noteq> y \<and> (x, y) \<in> r)" by blast
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   196
  then obtain z where z: "z \<in> Field r" "x \<noteq> z \<and> (x, z) \<in> r" "y \<noteq> z \<and> (y, z) \<in> r"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   197
    using Cinfinite_limit2[OF _ y(1) insert(5), of x] insert(4) by blast
49326
a063a96c8662 got rid of metis calls
blanchet
parents: 49312
diff changeset
   198
  show ?case
a063a96c8662 got rid of metis calls
blanchet
parents: 49312
diff changeset
   199
    apply (intro bexI ballI)
a063a96c8662 got rid of metis calls
blanchet
parents: 49312
diff changeset
   200
    apply (erule insertE)
a063a96c8662 got rid of metis calls
blanchet
parents: 49312
diff changeset
   201
    apply hypsubst
a063a96c8662 got rid of metis calls
blanchet
parents: 49312
diff changeset
   202
    apply (rule z(2))
a063a96c8662 got rid of metis calls
blanchet
parents: 49312
diff changeset
   203
    using Card_order_trans[OF insert(5)[THEN conjunct2]] y(2) z(3)
a063a96c8662 got rid of metis calls
blanchet
parents: 49312
diff changeset
   204
    apply blast
a063a96c8662 got rid of metis calls
blanchet
parents: 49312
diff changeset
   205
    apply (rule z(1))
a063a96c8662 got rid of metis calls
blanchet
parents: 49312
diff changeset
   206
    done
49312
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   207
qed
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   208
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   209
lemma insert_subsetI: "\<lbrakk>x \<in> A; X \<subseteq> A\<rbrakk> \<Longrightarrow> insert x X \<subseteq> A"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   210
by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   211
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   212
(*helps resolution*)
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   213
lemma well_order_induct_imp:
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   214
  "wo_rel r \<Longrightarrow> (\<And>x. \<forall>y. y \<noteq> x \<and> (y, x) \<in> r \<longrightarrow> y \<in> Field r \<longrightarrow> P y \<Longrightarrow> x \<in> Field r \<longrightarrow> P x) \<Longrightarrow>
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   215
     x \<in> Field r \<longrightarrow> P x"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   216
by (erule wo_rel.well_order_induct)
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   217
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   218
lemma meta_spec2:
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   219
  assumes "(\<And>x y. PROP P x y)"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   220
  shows "PROP P x y"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   221
by (rule `(\<And>x y. PROP P x y)`)
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   222
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents: 49308
diff changeset
   223
ML_file "Tools/bnf_lfp_util.ML"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents: 49308
diff changeset
   224
ML_file "Tools/bnf_lfp_tactics.ML"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents: 49308
diff changeset
   225
ML_file "Tools/bnf_lfp.ML"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents: 49308
diff changeset
   226
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   227
end