src/HOL/BNF/BNF_GFP.thy
author traytel
Tue, 13 Nov 2012 12:06:43 +0100
changeset 50058 bb1fadeba35e
parent 49635 fc0777f04205
child 51446 a6ebb12cc003
permissions -rw-r--r--
import Sublist rather than PrefixOrder to avoid unnecessary class instantiation
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
49509
163914705f8d renamed top-level theory from "Codatatype" to "BNF"
blanchet
parents: 49328
diff changeset
     1
(*  Title:      HOL/BNF/BNF_GFP.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
Greatest 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 {* Greatest 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_GFP
50058
bb1fadeba35e import Sublist rather than PrefixOrder to avoid unnecessary class instantiation
traytel
parents: 49635
diff changeset
    11
imports BNF_FP Equiv_Relations_More "~~/src/HOL/Library/Sublist"
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
  "codata" :: 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 sum_case_comp_Inl:
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    17
"sum_case f g \<circ> Inl = f"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    18
unfolding comp_def by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    19
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    20
lemma sum_case_expand_Inr: "f o Inl = g \<Longrightarrow> f x = sum_case g (f o Inr) x"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    21
by (auto split: sum.splits)
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 converse_Times: "(A \<times> B) ^-1 = B \<times> A"
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 equiv_triv1:
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    27
assumes "equiv A R" and "(a, b) \<in> R" and "(a, c) \<in> R"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    28
shows "(b, c) \<in> R"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    29
using assms unfolding equiv_def sym_def trans_def by blast
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    30
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    31
lemma equiv_triv2:
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    32
assumes "equiv A R" and "(a, b) \<in> R" and "(b, c) \<in> R"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    33
shows "(a, c) \<in> R"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    34
using assms unfolding equiv_def trans_def by blast
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    35
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    36
lemma equiv_proj:
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    37
  assumes e: "equiv A R" and "z \<in> R"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    38
  shows "(proj R o fst) z = (proj R o snd) z"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    39
proof -
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    40
  from assms(2) have z: "(fst z, snd z) \<in> R" by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    41
  have P: "\<And>x. (fst z, x) \<in> R \<Longrightarrow> (snd z, x) \<in> R" by (erule equiv_triv1[OF e z])
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    42
  have "\<And>x. (snd z, x) \<in> R \<Longrightarrow> (fst z, x) \<in> R" by (erule equiv_triv2[OF e z])
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    43
  with P show ?thesis unfolding proj_def[abs_def] by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    44
qed
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    45
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    46
(* Operators: *)
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    47
definition diag where "diag A \<equiv> {(a,a) | a. a \<in> A}"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    48
definition image2 where "image2 A f g = {(f a, g a) | a. a \<in> A}"
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
lemma diagI: "x \<in> A \<Longrightarrow> (x, x) \<in> diag A"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    51
unfolding diag_def by simp
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 diagE: "(a, b) \<in> diag A \<Longrightarrow> a = b"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    54
unfolding diag_def by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    55
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    56
lemma diagE': "x \<in> diag A \<Longrightarrow> fst x = snd x"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    57
unfolding diag_def by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    58
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    59
lemma diag_fst: "x \<in> diag A \<Longrightarrow> fst x \<in> A"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    60
unfolding diag_def by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    61
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    62
lemma diag_UNIV: "diag UNIV = Id"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    63
unfolding diag_def by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    64
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    65
lemma diag_converse: "diag A = (diag A) ^-1"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    66
unfolding diag_def by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    67
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    68
lemma diag_Comp: "diag A = diag A O diag A"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    69
unfolding diag_def by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    70
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    71
lemma diag_Gr: "diag A = Gr A id"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    72
unfolding diag_def Gr_def by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    73
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    74
lemma diag_UNIV_I: "x = y \<Longrightarrow> (x, y) \<in> diag UNIV"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    75
unfolding diag_def by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    76
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    77
lemma image2_eqI: "\<lbrakk>b = f x; c = g x; x \<in> A\<rbrakk> \<Longrightarrow> (b, c) \<in> image2 A f g"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    78
unfolding image2_def by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    79
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    80
lemma Id_subset: "Id \<subseteq> {(a, b). P a b \<or> a = b}"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    81
by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    82
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    83
lemma IdD: "(a, b) \<in> Id \<Longrightarrow> a = b"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    84
by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    85
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    86
lemma image2_Gr: "image2 A f g = (Gr A f)^-1 O (Gr A g)"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    87
unfolding image2_def Gr_def by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    88
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    89
lemma GrI: "\<lbrakk>x \<in> A; f x = fx\<rbrakk> \<Longrightarrow> (x, fx) \<in> Gr A f"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    90
unfolding Gr_def by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    91
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    92
lemma GrE: "(x, fx) \<in> Gr A f \<Longrightarrow> (x \<in> A \<Longrightarrow> f x = fx \<Longrightarrow> P) \<Longrightarrow> P"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    93
unfolding Gr_def by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    94
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    95
lemma GrD1: "(x, fx) \<in> Gr A f \<Longrightarrow> x \<in> A"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    96
unfolding Gr_def by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    97
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    98
lemma GrD2: "(x, fx) \<in> Gr A f \<Longrightarrow> f x = fx"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
    99
unfolding Gr_def by simp
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 Gr_incl: "Gr A f \<subseteq> A <*> B \<longleftrightarrow> f ` A \<subseteq> B"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   102
unfolding Gr_def by 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
definition relImage where
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   105
"relImage R f \<equiv> {(f a1, f a2) | a1 a2. (a1,a2) \<in> R}"
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
definition relInvImage where
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   108
"relInvImage A R f \<equiv> {(a1, a2) | a1 a2. a1 \<in> A \<and> a2 \<in> A \<and> (f a1, f a2) \<in> R}"
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 relImage_Gr:
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   111
"\<lbrakk>R \<subseteq> A \<times> A\<rbrakk> \<Longrightarrow> relImage R f = (Gr A f)^-1 O R O Gr A f"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   112
unfolding relImage_def Gr_def relcomp_def by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   113
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   114
lemma relInvImage_Gr: "\<lbrakk>R \<subseteq> B \<times> B\<rbrakk> \<Longrightarrow> relInvImage A R f = Gr A f O R O (Gr A f)^-1"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   115
unfolding Gr_def relcomp_def image_def relInvImage_def by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   116
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   117
lemma relImage_mono:
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   118
"R1 \<subseteq> R2 \<Longrightarrow> relImage R1 f \<subseteq> relImage R2 f"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   119
unfolding relImage_def by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   120
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   121
lemma relInvImage_mono:
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   122
"R1 \<subseteq> R2 \<Longrightarrow> relInvImage A R1 f \<subseteq> relInvImage A R2 f"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   123
unfolding relInvImage_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 relInvImage_diag:
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   126
"(\<And>a1 a2. f a1 = f a2 \<longleftrightarrow> a1 = a2) \<Longrightarrow> relInvImage A (diag B) f \<subseteq> Id"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   127
unfolding relInvImage_def diag_def by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   128
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   129
lemma relInvImage_UNIV_relImage:
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   130
"R \<subseteq> relInvImage UNIV (relImage R f) f"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   131
unfolding relInvImage_def relImage_def by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   132
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   133
lemma equiv_Image: "equiv A R \<Longrightarrow> (\<And>a b. (a, b) \<in> R \<Longrightarrow> a \<in> A \<and> b \<in> A \<and> R `` {a} = R `` {b})"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   134
unfolding equiv_def refl_on_def Image_def by (auto intro: transD symD)
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   135
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   136
lemma relImage_proj:
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   137
assumes "equiv A R"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   138
shows "relImage R (proj R) \<subseteq> diag (A//R)"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   139
unfolding relImage_def diag_def apply safe
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   140
using proj_iff[OF assms]
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   141
by (metis assms equiv_Image proj_def proj_preserves)
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   142
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   143
lemma relImage_relInvImage:
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   144
assumes "R \<subseteq> f ` A <*> f ` A"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   145
shows "relImage (relInvImage A R f) f = R"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   146
using assms unfolding relImage_def relInvImage_def by fastforce
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   147
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   148
lemma subst_Pair: "P x y \<Longrightarrow> a = (x, y) \<Longrightarrow> P (fst a) (snd a)"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   149
by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   150
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   151
lemma fst_diag_id: "(fst \<circ> (%x. (x, x))) z = id z"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   152
by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   153
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   154
lemma snd_diag_id: "(snd \<circ> (%x. (x, x))) z = id z"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   155
by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   156
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   157
lemma Collect_restrict': "{(x, y) | x y. phi x y \<and> P x y} \<subseteq> {(x, y) | x y. phi x y}"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   158
by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   159
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   160
lemma image_convolD: "\<lbrakk>(a, b) \<in> <f, g> ` X\<rbrakk> \<Longrightarrow> \<exists>x. x \<in> X \<and> a = f x \<and> b = g x"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   161
unfolding convol_def by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   162
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   163
(*Extended Sublist*)
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   164
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   165
definition prefCl where
50058
bb1fadeba35e import Sublist rather than PrefixOrder to avoid unnecessary class instantiation
traytel
parents: 49635
diff changeset
   166
  "prefCl Kl = (\<forall> kl1 kl2. prefixeq kl1 kl2 \<and> kl2 \<in> Kl \<longrightarrow> kl1 \<in> Kl)"
49312
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   167
definition PrefCl where
50058
bb1fadeba35e import Sublist rather than PrefixOrder to avoid unnecessary class instantiation
traytel
parents: 49635
diff changeset
   168
  "PrefCl A n = (\<forall>kl kl'. kl \<in> A n \<and> prefixeq kl' kl \<longrightarrow> (\<exists>m\<le>n. kl' \<in> A m))"
49312
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   169
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   170
lemma prefCl_UN:
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   171
  "\<lbrakk>\<And>n. PrefCl A n\<rbrakk> \<Longrightarrow> prefCl (\<Union>n. A n)"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   172
unfolding prefCl_def PrefCl_def by fastforce
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   173
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   174
definition Succ where "Succ Kl kl = {k . kl @ [k] \<in> Kl}"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   175
definition Shift where "Shift Kl k = {kl. k # kl \<in> Kl}"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   176
definition shift where "shift lab k = (\<lambda>kl. lab (k # kl))"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   177
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   178
lemma empty_Shift: "\<lbrakk>[] \<in> Kl; k \<in> Succ Kl []\<rbrakk> \<Longrightarrow> [] \<in> Shift Kl k"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   179
unfolding Shift_def Succ_def by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   180
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   181
lemma Shift_clists: "Kl \<subseteq> Field (clists r) \<Longrightarrow> Shift Kl k \<subseteq> Field (clists r)"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   182
unfolding Shift_def clists_def Field_card_of by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   183
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   184
lemma Shift_prefCl: "prefCl Kl \<Longrightarrow> prefCl (Shift Kl k)"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   185
unfolding prefCl_def Shift_def
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   186
proof safe
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   187
  fix kl1 kl2
50058
bb1fadeba35e import Sublist rather than PrefixOrder to avoid unnecessary class instantiation
traytel
parents: 49635
diff changeset
   188
  assume "\<forall>kl1 kl2. prefixeq kl1 kl2 \<and> kl2 \<in> Kl \<longrightarrow> kl1 \<in> Kl"
bb1fadeba35e import Sublist rather than PrefixOrder to avoid unnecessary class instantiation
traytel
parents: 49635
diff changeset
   189
    "prefixeq kl1 kl2" "k # kl2 \<in> Kl"
bb1fadeba35e import Sublist rather than PrefixOrder to avoid unnecessary class instantiation
traytel
parents: 49635
diff changeset
   190
  thus "k # kl1 \<in> Kl" using Cons_prefixeq_Cons[of k kl1 k kl2] by blast
49312
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   191
qed
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   192
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   193
lemma not_in_Shift: "kl \<notin> Shift Kl x \<Longrightarrow> x # kl \<notin> Kl"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   194
unfolding Shift_def by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   195
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   196
lemma prefCl_Succ: "\<lbrakk>prefCl Kl; k # kl \<in> Kl\<rbrakk> \<Longrightarrow> k \<in> Succ Kl []"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   197
unfolding Succ_def proof
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   198
  assume "prefCl Kl" "k # kl \<in> Kl"
50058
bb1fadeba35e import Sublist rather than PrefixOrder to avoid unnecessary class instantiation
traytel
parents: 49635
diff changeset
   199
  moreover have "prefixeq (k # []) (k # kl)" by auto
49312
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   200
  ultimately have "k # [] \<in> Kl" unfolding prefCl_def by blast
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   201
  thus "[] @ [k] \<in> Kl" by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   202
qed
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   203
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   204
lemma SuccD: "k \<in> Succ Kl kl \<Longrightarrow> kl @ [k] \<in> Kl"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   205
unfolding Succ_def by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   206
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   207
lemmas SuccE = SuccD[elim_format]
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 SuccI: "kl @ [k] \<in> Kl \<Longrightarrow> k \<in> Succ Kl kl"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   210
unfolding Succ_def by simp
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
lemma ShiftD: "kl \<in> Shift Kl k \<Longrightarrow> k # kl \<in> Kl"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   213
unfolding Shift_def by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   214
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   215
lemma Succ_Shift: "Succ (Shift Kl k) kl = Succ Kl (k # kl)"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   216
unfolding Succ_def Shift_def by auto
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 ShiftI: "k # kl \<in> Kl \<Longrightarrow> kl \<in> Shift Kl k"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   219
unfolding Shift_def by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   220
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   221
lemma Func_cexp: "|Func A B| =o |B| ^c |A|"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   222
unfolding cexp_def Field_card_of by (simp only: card_of_refl)
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   223
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   224
lemma clists_bound: "A \<in> Field (cpow (clists r)) - {{}} \<Longrightarrow> |A| \<le>o clists r"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   225
unfolding cpow_def clists_def Field_card_of by (auto simp: card_of_mono1)
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   226
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   227
lemma cpow_clists_czero: "\<lbrakk>A \<in> Field (cpow (clists r)) - {{}}; |A| =o czero\<rbrakk> \<Longrightarrow> False"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   228
unfolding cpow_def clists_def
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   229
by (auto simp add: card_of_ordIso_czero_iff_empty[symmetric])
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   230
   (erule notE, erule ordIso_transitive, rule czero_ordIso)
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   231
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   232
lemma incl_UNION_I:
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   233
assumes "i \<in> I" and "A \<subseteq> F i"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   234
shows "A \<subseteq> UNION I F"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   235
using assms by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   236
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   237
lemma Nil_clists: "{[]} \<subseteq> Field (clists r)"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   238
unfolding clists_def Field_card_of by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   239
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   240
lemma Cons_clists:
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   241
  "\<lbrakk>x \<in> Field r; xs \<in> Field (clists r)\<rbrakk> \<Longrightarrow> x # xs \<in> Field (clists r)"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   242
unfolding clists_def Field_card_of by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   243
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   244
lemma length_Cons: "length (x # xs) = Suc (length xs)"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   245
by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   246
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   247
lemma length_append_singleton: "length (xs @ [x]) = Suc (length xs)"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   248
by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   249
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   250
(*injection into the field of a cardinal*)
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   251
definition "toCard_pred A r f \<equiv> inj_on f A \<and> f ` A \<subseteq> Field r \<and> Card_order r"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   252
definition "toCard A r \<equiv> SOME f. toCard_pred A r f"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   253
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   254
lemma ex_toCard_pred:
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   255
"\<lbrakk>|A| \<le>o r; Card_order r\<rbrakk> \<Longrightarrow> \<exists> f. toCard_pred A r f"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   256
unfolding toCard_pred_def
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   257
using card_of_ordLeq[of A "Field r"]
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   258
      ordLeq_ordIso_trans[OF _ card_of_unique[of "Field r" r], of "|A|"]
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   259
by blast
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   260
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   261
lemma toCard_pred_toCard:
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   262
  "\<lbrakk>|A| \<le>o r; Card_order r\<rbrakk> \<Longrightarrow> toCard_pred A r (toCard A r)"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   263
unfolding toCard_def using someI_ex[OF ex_toCard_pred] .
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   264
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   265
lemma toCard_inj: "\<lbrakk>|A| \<le>o r; Card_order r; x \<in> A; y \<in> A\<rbrakk> \<Longrightarrow>
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   266
  toCard A r x = toCard A r y \<longleftrightarrow> x = y"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   267
using toCard_pred_toCard unfolding inj_on_def toCard_pred_def by blast
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   268
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   269
lemma toCard: "\<lbrakk>|A| \<le>o r; Card_order r; b \<in> A\<rbrakk> \<Longrightarrow> toCard A r b \<in> Field r"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   270
using toCard_pred_toCard unfolding toCard_pred_def by blast
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   271
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   272
definition "fromCard A r k \<equiv> SOME b. b \<in> A \<and> toCard A r b = k"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   273
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   274
lemma fromCard_toCard:
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   275
"\<lbrakk>|A| \<le>o r; Card_order r; b \<in> A\<rbrakk> \<Longrightarrow> fromCard A r (toCard A r b) = b"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   276
unfolding fromCard_def by (rule some_equality) (auto simp add: toCard_inj)
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   277
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   278
(* pick according to the weak pullback *)
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   279
definition pickWP_pred where
49325
blanchet
parents: 49314
diff changeset
   280
"pickWP_pred A p1 p2 b1 b2 a \<equiv> a \<in> A \<and> p1 a = b1 \<and> p2 a = b2"
49312
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   281
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   282
definition pickWP where
49325
blanchet
parents: 49314
diff changeset
   283
"pickWP A p1 p2 b1 b2 \<equiv> SOME a. pickWP_pred A p1 p2 b1 b2 a"
49312
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   284
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   285
lemma pickWP_pred:
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   286
assumes "wpull A B1 B2 f1 f2 p1 p2" and
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   287
"b1 \<in> B1" and "b2 \<in> B2" and "f1 b1 = f2 b2"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   288
shows "\<exists> a. pickWP_pred A p1 p2 b1 b2 a"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   289
using assms unfolding wpull_def pickWP_pred_def by blast
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   290
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   291
lemma pickWP_pred_pickWP:
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   292
assumes "wpull A B1 B2 f1 f2 p1 p2" and
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   293
"b1 \<in> B1" and "b2 \<in> B2" and "f1 b1 = f2 b2"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   294
shows "pickWP_pred A p1 p2 b1 b2 (pickWP A p1 p2 b1 b2)"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   295
unfolding pickWP_def using assms by(rule someI_ex[OF pickWP_pred])
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   296
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   297
lemma pickWP:
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   298
assumes "wpull A B1 B2 f1 f2 p1 p2" and
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   299
"b1 \<in> B1" and "b2 \<in> B2" and "f1 b1 = f2 b2"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   300
shows "pickWP A p1 p2 b1 b2 \<in> A"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   301
      "p1 (pickWP A p1 p2 b1 b2) = b1"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   302
      "p2 (pickWP A p1 p2 b1 b2) = b2"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   303
using assms pickWP_pred_pickWP unfolding pickWP_pred_def by fastforce+
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   304
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   305
lemma Inl_Field_csum: "a \<in> Field r \<Longrightarrow> Inl a \<in> Field (r +c s)"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   306
unfolding Field_card_of csum_def by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   307
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   308
lemma Inr_Field_csum: "a \<in> Field s \<Longrightarrow> Inr a \<in> Field (r +c s)"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   309
unfolding Field_card_of csum_def by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   310
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   311
lemma nat_rec_0: "f = nat_rec f1 (%n rec. f2 n rec) \<Longrightarrow> f 0 = f1"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   312
by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   313
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   314
lemma nat_rec_Suc: "f = nat_rec f1 (%n rec. f2 n rec) \<Longrightarrow> f (Suc n) = f2 n (f n)"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   315
by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   316
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   317
lemma list_rec_Nil: "f = list_rec f1 (%x xs rec. f2 x xs rec) \<Longrightarrow> f [] = f1"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   318
by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   319
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   320
lemma list_rec_Cons: "f = list_rec f1 (%x xs rec. f2 x xs rec) \<Longrightarrow> f (x # xs) = f2 x xs (f xs)"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   321
by auto
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   322
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   323
lemma not_arg_cong_Inr: "x \<noteq> y \<Longrightarrow> Inr x \<noteq> Inr y"
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   324
by simp
c874ff5658dc moved theorems closer to where they are used
blanchet
parents: 49309
diff changeset
   325
49309
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents: 49308
diff changeset
   326
ML_file "Tools/bnf_gfp_util.ML"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents: 49308
diff changeset
   327
ML_file "Tools/bnf_gfp_tactics.ML"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents: 49308
diff changeset
   328
ML_file "Tools/bnf_gfp.ML"
f20b24214ac2 split basic BNFs into really basic ones and others, and added Andreas Lochbihler's "option" BNF
blanchet
parents: 49308
diff changeset
   329
48975
7f79f94a432c added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff changeset
   330
end