src/HOL/HOLCF/Universal.thy
author wenzelm
Sun, 11 Feb 2018 12:51:23 +0100
changeset 67593 5efb88c90051
parent 67399 eab6ce8368fa
child 68780 54fdc8bc73a3
permissions -rw-r--r--
clarified command-line defaults;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
42151
4da4fc77664b tuned headers;
wenzelm
parents: 41529
diff changeset
     1
(*  Title:      HOL/HOLCF/Universal.thy
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
     2
    Author:     Brian Huffman
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
     3
*)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
     4
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61605
diff changeset
     5
section \<open>A universal bifinite domain\<close>
35794
8cd7134275cc use headers consistently
huffman
parents: 35701
diff changeset
     6
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
     7
theory Universal
66453
cc19f7ca2ed6 session-qualified theory imports: isabelle imports -U -i -d '~~/src/Benchmarks' -a;
wenzelm
parents: 65552
diff changeset
     8
imports Bifinite Completion "HOL-Library.Nat_Bijection"
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
     9
begin
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    10
65552
f533820e7248 theories "GCD" and "Binomial" are already included in "Main": this avoids improper imports in applications;
wenzelm
parents: 64267
diff changeset
    11
no_notation binomial  (infixl "choose" 65)
f533820e7248 theories "GCD" and "Binomial" are already included in "Main": this avoids improper imports in applications;
wenzelm
parents: 64267
diff changeset
    12
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61605
diff changeset
    13
subsection \<open>Basis for universal domain\<close>
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
    14
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61605
diff changeset
    15
subsubsection \<open>Basis datatype\<close>
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    16
41295
5b5388d4ccc9 types -> type_synonym
huffman
parents: 41286
diff changeset
    17
type_synonym ubasis = nat
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    18
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    19
definition
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    20
  node :: "nat \<Rightarrow> ubasis \<Rightarrow> ubasis set \<Rightarrow> ubasis"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    21
where
35701
0f5bf989da42 adapt HOLCF to use Nat_Bijection library
huffman
parents: 34915
diff changeset
    22
  "node i a S = Suc (prod_encode (i, prod_encode (a, set_encode S)))"
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    23
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
    24
lemma node_not_0 [simp]: "node i a S \<noteq> 0"
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    25
unfolding node_def by simp
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    26
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
    27
lemma node_gt_0 [simp]: "0 < node i a S"
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    28
unfolding node_def by simp
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    29
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    30
lemma node_inject [simp]:
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
    31
  "\<lbrakk>finite S; finite T\<rbrakk>
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
    32
    \<Longrightarrow> node i a S = node j b T \<longleftrightarrow> i = j \<and> a = b \<and> S = T"
35701
0f5bf989da42 adapt HOLCF to use Nat_Bijection library
huffman
parents: 34915
diff changeset
    33
unfolding node_def by (simp add: prod_encode_eq set_encode_eq)
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    34
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
    35
lemma node_gt0: "i < node i a S"
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    36
unfolding node_def less_Suc_eq_le
35701
0f5bf989da42 adapt HOLCF to use Nat_Bijection library
huffman
parents: 34915
diff changeset
    37
by (rule le_prod_encode_1)
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    38
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
    39
lemma node_gt1: "a < node i a S"
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    40
unfolding node_def less_Suc_eq_le
35701
0f5bf989da42 adapt HOLCF to use Nat_Bijection library
huffman
parents: 34915
diff changeset
    41
by (rule order_trans [OF le_prod_encode_1 le_prod_encode_2])
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    42
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    43
lemma nat_less_power2: "n < 2^n"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    44
by (induct n) simp_all
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    45
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
    46
lemma node_gt2: "\<lbrakk>finite S; b \<in> S\<rbrakk> \<Longrightarrow> b < node i a S"
35701
0f5bf989da42 adapt HOLCF to use Nat_Bijection library
huffman
parents: 34915
diff changeset
    47
unfolding node_def less_Suc_eq_le set_encode_def
0f5bf989da42 adapt HOLCF to use Nat_Bijection library
huffman
parents: 34915
diff changeset
    48
apply (rule order_trans [OF _ le_prod_encode_2])
0f5bf989da42 adapt HOLCF to use Nat_Bijection library
huffman
parents: 34915
diff changeset
    49
apply (rule order_trans [OF _ le_prod_encode_2])
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66453
diff changeset
    50
apply (rule order_trans [where y="sum ((^) 2) {b}"])
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    51
apply (simp add: nat_less_power2 [THEN order_less_imp_le])
64267
b9a1486e79be setsum -> sum
nipkow
parents: 62390
diff changeset
    52
apply (erule sum_mono2, simp, simp)
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    53
done
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    54
35701
0f5bf989da42 adapt HOLCF to use Nat_Bijection library
huffman
parents: 34915
diff changeset
    55
lemma eq_prod_encode_pairI:
0f5bf989da42 adapt HOLCF to use Nat_Bijection library
huffman
parents: 34915
diff changeset
    56
  "\<lbrakk>fst (prod_decode x) = a; snd (prod_decode x) = b\<rbrakk> \<Longrightarrow> x = prod_encode (a, b)"
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    57
by (erule subst, erule subst, simp)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    58
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    59
lemma node_cases:
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    60
  assumes 1: "x = 0 \<Longrightarrow> P"
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
    61
  assumes 2: "\<And>i a S. \<lbrakk>finite S; x = node i a S\<rbrakk> \<Longrightarrow> P"
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    62
  shows "P"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    63
 apply (cases x)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    64
  apply (erule 1)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    65
 apply (rule 2)
35701
0f5bf989da42 adapt HOLCF to use Nat_Bijection library
huffman
parents: 34915
diff changeset
    66
  apply (rule finite_set_decode)
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    67
 apply (simp add: node_def)
35701
0f5bf989da42 adapt HOLCF to use Nat_Bijection library
huffman
parents: 34915
diff changeset
    68
 apply (rule eq_prod_encode_pairI [OF refl])
0f5bf989da42 adapt HOLCF to use Nat_Bijection library
huffman
parents: 34915
diff changeset
    69
 apply (rule eq_prod_encode_pairI [OF refl refl])
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    70
done
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    71
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    72
lemma node_induct:
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    73
  assumes 1: "P 0"
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
    74
  assumes 2: "\<And>i a S. \<lbrakk>P a; finite S; \<forall>b\<in>S. P b\<rbrakk> \<Longrightarrow> P (node i a S)"
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    75
  shows "P x"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    76
 apply (induct x rule: nat_less_induct)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    77
 apply (case_tac n rule: node_cases)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    78
  apply (simp add: 1)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    79
 apply (simp add: 2 node_gt1 node_gt2)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    80
done
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    81
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61605
diff changeset
    82
subsubsection \<open>Basis ordering\<close>
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    83
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    84
inductive
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    85
  ubasis_le :: "nat \<Rightarrow> nat \<Rightarrow> bool"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    86
where
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
    87
  ubasis_le_refl: "ubasis_le a a"
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    88
| ubasis_le_trans:
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
    89
    "\<lbrakk>ubasis_le a b; ubasis_le b c\<rbrakk> \<Longrightarrow> ubasis_le a c"
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    90
| ubasis_le_lower:
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
    91
    "finite S \<Longrightarrow> ubasis_le a (node i a S)"
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    92
| ubasis_le_upper:
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
    93
    "\<lbrakk>finite S; b \<in> S; ubasis_le a b\<rbrakk> \<Longrightarrow> ubasis_le (node i a S) b"
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    94
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    95
lemma ubasis_le_minimal: "ubasis_le 0 x"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    96
apply (induct x rule: node_induct)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    97
apply (rule ubasis_le_refl)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    98
apply (erule ubasis_le_trans)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
    99
apply (erule ubasis_le_lower)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   100
done
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   101
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   102
interpretation udom: preorder ubasis_le
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 59667
diff changeset
   103
apply standard
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   104
apply (rule ubasis_le_refl)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   105
apply (erule (1) ubasis_le_trans)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   106
done
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   107
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61605
diff changeset
   108
subsubsection \<open>Generic take function\<close>
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   109
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   110
function
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   111
  ubasis_until :: "(ubasis \<Rightarrow> bool) \<Rightarrow> ubasis \<Rightarrow> ubasis"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   112
where
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   113
  "ubasis_until P 0 = 0"
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   114
| "finite S \<Longrightarrow> ubasis_until P (node i a S) =
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   115
    (if P (node i a S) then node i a S else ubasis_until P a)"
56073
29e308b56d23 enhanced simplifier solver for preconditions of rewrite rule, can now deal with conjunctions
nipkow
parents: 54863
diff changeset
   116
   apply clarify
29e308b56d23 enhanced simplifier solver for preconditions of rewrite rule, can now deal with conjunctions
nipkow
parents: 54863
diff changeset
   117
   apply (rule_tac x=b in node_cases)
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   118
    apply simp
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   119
   apply simp
56073
29e308b56d23 enhanced simplifier solver for preconditions of rewrite rule, can now deal with conjunctions
nipkow
parents: 54863
diff changeset
   120
   apply fast
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   121
  apply simp
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   122
 apply simp
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   123
done
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   124
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   125
termination ubasis_until
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   126
apply (relation "measure snd")
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   127
apply (rule wf_measure)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   128
apply (simp add: node_gt1)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   129
done
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   130
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   131
lemma ubasis_until: "P 0 \<Longrightarrow> P (ubasis_until P x)"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   132
by (induct x rule: node_induct) simp_all
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   133
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   134
lemma ubasis_until': "0 < ubasis_until P x \<Longrightarrow> P (ubasis_until P x)"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   135
by (induct x rule: node_induct) auto
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   136
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   137
lemma ubasis_until_same: "P x \<Longrightarrow> ubasis_until P x = x"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   138
by (induct x rule: node_induct) simp_all
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   139
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   140
lemma ubasis_until_idem:
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   141
  "P 0 \<Longrightarrow> ubasis_until P (ubasis_until P x) = ubasis_until P x"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   142
by (rule ubasis_until_same [OF ubasis_until])
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   143
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   144
lemma ubasis_until_0:
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   145
  "\<forall>x. x \<noteq> 0 \<longrightarrow> \<not> P x \<Longrightarrow> ubasis_until P x = 0"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   146
by (induct x rule: node_induct) simp_all
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   147
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   148
lemma ubasis_until_less: "ubasis_le (ubasis_until P x) x"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   149
apply (induct x rule: node_induct)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   150
apply (simp add: ubasis_le_refl)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   151
apply (simp add: ubasis_le_refl)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   152
apply (rule impI)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   153
apply (erule ubasis_le_trans)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   154
apply (erule ubasis_le_lower)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   155
done
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   156
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   157
lemma ubasis_until_chain:
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   158
  assumes PQ: "\<And>x. P x \<Longrightarrow> Q x"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   159
  shows "ubasis_le (ubasis_until P x) (ubasis_until Q x)"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   160
apply (induct x rule: node_induct)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   161
apply (simp add: ubasis_le_refl)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   162
apply (simp add: ubasis_le_refl)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   163
apply (simp add: PQ)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   164
apply clarify
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   165
apply (rule ubasis_le_trans)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   166
apply (rule ubasis_until_less)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   167
apply (erule ubasis_le_lower)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   168
done
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   169
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   170
lemma ubasis_until_mono:
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   171
  assumes "\<And>i a S b. \<lbrakk>finite S; P (node i a S); b \<in> S; ubasis_le a b\<rbrakk> \<Longrightarrow> P b"
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   172
  shows "ubasis_le a b \<Longrightarrow> ubasis_le (ubasis_until P a) (ubasis_until P b)"
30561
5e6088e1d6df clean up proofs
huffman
parents: 30505
diff changeset
   173
proof (induct set: ubasis_le)
5e6088e1d6df clean up proofs
huffman
parents: 30505
diff changeset
   174
  case (ubasis_le_refl a) show ?case by (rule ubasis_le.ubasis_le_refl)
5e6088e1d6df clean up proofs
huffman
parents: 30505
diff changeset
   175
next
5e6088e1d6df clean up proofs
huffman
parents: 30505
diff changeset
   176
  case (ubasis_le_trans a b c) thus ?case by - (rule ubasis_le.ubasis_le_trans)
5e6088e1d6df clean up proofs
huffman
parents: 30505
diff changeset
   177
next
5e6088e1d6df clean up proofs
huffman
parents: 30505
diff changeset
   178
  case (ubasis_le_lower S a i) thus ?case
5e6088e1d6df clean up proofs
huffman
parents: 30505
diff changeset
   179
    apply (clarsimp simp add: ubasis_le_refl)
5e6088e1d6df clean up proofs
huffman
parents: 30505
diff changeset
   180
    apply (rule ubasis_le_trans [OF ubasis_until_less])
5e6088e1d6df clean up proofs
huffman
parents: 30505
diff changeset
   181
    apply (erule ubasis_le.ubasis_le_lower)
5e6088e1d6df clean up proofs
huffman
parents: 30505
diff changeset
   182
    done
5e6088e1d6df clean up proofs
huffman
parents: 30505
diff changeset
   183
next
5e6088e1d6df clean up proofs
huffman
parents: 30505
diff changeset
   184
  case (ubasis_le_upper S b a i) thus ?case
5e6088e1d6df clean up proofs
huffman
parents: 30505
diff changeset
   185
    apply clarsimp
5e6088e1d6df clean up proofs
huffman
parents: 30505
diff changeset
   186
    apply (subst ubasis_until_same)
41529
ba60efa2fd08 eliminated global prems;
wenzelm
parents: 41430
diff changeset
   187
     apply (erule (3) assms)
30561
5e6088e1d6df clean up proofs
huffman
parents: 30505
diff changeset
   188
    apply (erule (2) ubasis_le.ubasis_le_upper)
5e6088e1d6df clean up proofs
huffman
parents: 30505
diff changeset
   189
    done
5e6088e1d6df clean up proofs
huffman
parents: 30505
diff changeset
   190
qed
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   191
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   192
lemma finite_range_ubasis_until:
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   193
  "finite {x. P x} \<Longrightarrow> finite (range (ubasis_until P))"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   194
apply (rule finite_subset [where B="insert 0 {x. P x}"])
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   195
apply (clarsimp simp add: ubasis_until')
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   196
apply simp
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   197
done
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   198
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   199
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61605
diff changeset
   200
subsection \<open>Defining the universal domain by ideal completion\<close>
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   201
49834
b27bbb021df1 discontinued obsolete typedef (open) syntax;
wenzelm
parents: 46868
diff changeset
   202
typedef udom = "{S. udom.ideal S}"
40888
28cd51cff70c reformulate lemma preorder.ex_ideal, and use it for typedefs
huffman
parents: 40774
diff changeset
   203
by (rule udom.ex_ideal)
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   204
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 30729
diff changeset
   205
instantiation udom :: below
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   206
begin
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   207
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   208
definition
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   209
  "x \<sqsubseteq> y \<longleftrightarrow> Rep_udom x \<subseteq> Rep_udom y"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   210
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   211
instance ..
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   212
end
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   213
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   214
instance udom :: po
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   215
using type_definition_udom below_udom_def
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   216
by (rule udom.typedef_ideal_po)
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   217
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   218
instance udom :: cpo
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   219
using type_definition_udom below_udom_def
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   220
by (rule udom.typedef_ideal_cpo)
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   221
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   222
definition
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   223
  udom_principal :: "nat \<Rightarrow> udom" where
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   224
  "udom_principal t = Abs_udom {u. ubasis_le u t}"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   225
39984
0300d5170622 add lemma typedef_ideal_completion
huffman
parents: 39974
diff changeset
   226
lemma ubasis_countable: "\<exists>f::ubasis \<Rightarrow> nat. inj f"
0300d5170622 add lemma typedef_ideal_completion
huffman
parents: 39974
diff changeset
   227
by (rule exI, rule inj_on_id)
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   228
30729
461ee3e49ad3 interpretation/interpret: prefixes are mandatory by default;
wenzelm
parents: 30561
diff changeset
   229
interpretation udom:
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   230
  ideal_completion ubasis_le udom_principal Rep_udom
39984
0300d5170622 add lemma typedef_ideal_completion
huffman
parents: 39974
diff changeset
   231
using type_definition_udom below_udom_def
0300d5170622 add lemma typedef_ideal_completion
huffman
parents: 39974
diff changeset
   232
using udom_principal_def ubasis_countable
0300d5170622 add lemma typedef_ideal_completion
huffman
parents: 39974
diff changeset
   233
by (rule udom.typedef_ideal_completion)
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   234
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61605
diff changeset
   235
text \<open>Universal domain is pointed\<close>
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   236
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   237
lemma udom_minimal: "udom_principal 0 \<sqsubseteq> x"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   238
apply (induct x rule: udom.principal_induct)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   239
apply (simp, simp add: ubasis_le_minimal)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   240
done
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   241
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   242
instance udom :: pcpo
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   243
by intro_classes (fast intro: udom_minimal)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   244
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   245
lemma inst_udom_pcpo: "\<bottom> = udom_principal 0"
41430
1aa23e9f2c87 change some lemma names containing 'UU' to 'bottom'
huffman
parents: 41413
diff changeset
   246
by (rule udom_minimal [THEN bottomI, symmetric])
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   247
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   248
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61605
diff changeset
   249
subsection \<open>Compact bases of domains\<close>
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   250
49834
b27bbb021df1 discontinued obsolete typedef (open) syntax;
wenzelm
parents: 46868
diff changeset
   251
typedef 'a compact_basis = "{x::'a::pcpo. compact x}"
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   252
by auto
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   253
41370
17b09240893c declare more simp rules, rewrite proofs in Isar-style
huffman
parents: 41295
diff changeset
   254
lemma Rep_compact_basis' [simp]: "compact (Rep_compact_basis a)"
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   255
by (rule Rep_compact_basis [unfolded mem_Collect_eq])
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   256
41370
17b09240893c declare more simp rules, rewrite proofs in Isar-style
huffman
parents: 41295
diff changeset
   257
lemma Abs_compact_basis_inverse' [simp]:
17b09240893c declare more simp rules, rewrite proofs in Isar-style
huffman
parents: 41295
diff changeset
   258
   "compact x \<Longrightarrow> Rep_compact_basis (Abs_compact_basis x) = x"
17b09240893c declare more simp rules, rewrite proofs in Isar-style
huffman
parents: 41295
diff changeset
   259
by (rule Abs_compact_basis_inverse [unfolded mem_Collect_eq])
17b09240893c declare more simp rules, rewrite proofs in Isar-style
huffman
parents: 41295
diff changeset
   260
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   261
instantiation compact_basis :: (pcpo) below
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   262
begin
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   263
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   264
definition
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   265
  compact_le_def:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 66453
diff changeset
   266
    "(\<sqsubseteq>) \<equiv> (\<lambda>x y. Rep_compact_basis x \<sqsubseteq> Rep_compact_basis y)"
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   267
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   268
instance ..
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   269
end
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   270
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   271
instance compact_basis :: (pcpo) po
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   272
using type_definition_compact_basis compact_le_def
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   273
by (rule typedef_po)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   274
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   275
definition
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   276
  approximants :: "'a \<Rightarrow> 'a compact_basis set" where
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   277
  "approximants = (\<lambda>x. {a. Rep_compact_basis a \<sqsubseteq> x})"
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   278
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   279
definition
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   280
  compact_bot :: "'a::pcpo compact_basis" where
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   281
  "compact_bot = Abs_compact_basis \<bottom>"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   282
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   283
lemma Rep_compact_bot [simp]: "Rep_compact_basis compact_bot = \<bottom>"
41370
17b09240893c declare more simp rules, rewrite proofs in Isar-style
huffman
parents: 41295
diff changeset
   284
unfolding compact_bot_def by simp
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   285
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   286
lemma compact_bot_minimal [simp]: "compact_bot \<sqsubseteq> a"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   287
unfolding compact_le_def Rep_compact_bot by simp
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   288
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   289
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61605
diff changeset
   290
subsection \<open>Universality of \emph{udom}\<close>
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   291
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61605
diff changeset
   292
text \<open>We use a locale to parameterize the construction over a chain
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61605
diff changeset
   293
of approx functions on the type to be embedded.\<close>
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   294
46868
6c250adbe101 eliminated old-fashioned 'constrains' element;
wenzelm
parents: 42151
diff changeset
   295
locale bifinite_approx_chain =
6c250adbe101 eliminated old-fashioned 'constrains' element;
wenzelm
parents: 42151
diff changeset
   296
  approx_chain approx for approx :: "nat \<Rightarrow> 'a::bifinite \<rightarrow> 'a"
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   297
begin
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   298
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61605
diff changeset
   299
subsubsection \<open>Choosing a maximal element from a finite set\<close>
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   300
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   301
lemma finite_has_maximal:
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   302
  fixes A :: "'a compact_basis set"
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   303
  shows "\<lbrakk>finite A; A \<noteq> {}\<rbrakk> \<Longrightarrow> \<exists>x\<in>A. \<forall>y\<in>A. x \<sqsubseteq> y \<longrightarrow> x = y"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   304
proof (induct rule: finite_ne_induct)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   305
  case (singleton x)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   306
    show ?case by simp
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   307
next
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   308
  case (insert a A)
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61605
diff changeset
   309
  from \<open>\<exists>x\<in>A. \<forall>y\<in>A. x \<sqsubseteq> y \<longrightarrow> x = y\<close>
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   310
  obtain x where x: "x \<in> A"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   311
           and x_eq: "\<And>y. \<lbrakk>y \<in> A; x \<sqsubseteq> y\<rbrakk> \<Longrightarrow> x = y" by fast
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   312
  show ?case
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   313
  proof (intro bexI ballI impI)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   314
    fix y
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   315
    assume "y \<in> insert a A" and "(if x \<sqsubseteq> a then a else x) \<sqsubseteq> y"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   316
    thus "(if x \<sqsubseteq> a then a else x) = y"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   317
      apply auto
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 30729
diff changeset
   318
      apply (frule (1) below_trans)
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   319
      apply (frule (1) x_eq)
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 30729
diff changeset
   320
      apply (rule below_antisym, assumption)
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   321
      apply simp
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   322
      apply (erule (1) x_eq)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   323
      done
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   324
  next
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   325
    show "(if x \<sqsubseteq> a then a else x) \<in> insert a A"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   326
      by (simp add: x)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   327
  qed
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   328
qed
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   329
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   330
definition
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   331
  choose :: "'a compact_basis set \<Rightarrow> 'a compact_basis"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   332
where
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   333
  "choose A = (SOME x. x \<in> {x\<in>A. \<forall>y\<in>A. x \<sqsubseteq> y \<longrightarrow> x = y})"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   334
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   335
lemma choose_lemma:
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   336
  "\<lbrakk>finite A; A \<noteq> {}\<rbrakk> \<Longrightarrow> choose A \<in> {x\<in>A. \<forall>y\<in>A. x \<sqsubseteq> y \<longrightarrow> x = y}"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   337
unfolding choose_def
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   338
apply (rule someI_ex)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   339
apply (frule (1) finite_has_maximal, fast)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   340
done
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   341
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   342
lemma maximal_choose:
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   343
  "\<lbrakk>finite A; y \<in> A; choose A \<sqsubseteq> y\<rbrakk> \<Longrightarrow> choose A = y"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   344
apply (cases "A = {}", simp)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   345
apply (frule (1) choose_lemma, simp)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   346
done
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   347
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   348
lemma choose_in: "\<lbrakk>finite A; A \<noteq> {}\<rbrakk> \<Longrightarrow> choose A \<in> A"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   349
by (frule (1) choose_lemma, simp)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   350
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   351
function
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   352
  choose_pos :: "'a compact_basis set \<Rightarrow> 'a compact_basis \<Rightarrow> nat"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   353
where
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   354
  "choose_pos A x =
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   355
    (if finite A \<and> x \<in> A \<and> x \<noteq> choose A
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   356
      then Suc (choose_pos (A - {choose A}) x) else 0)"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   357
by auto
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   358
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   359
termination choose_pos
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   360
apply (relation "measure (card \<circ> fst)", simp)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   361
apply clarsimp
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   362
apply (rule card_Diff1_less)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   363
apply assumption
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   364
apply (erule choose_in)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   365
apply clarsimp
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   366
done
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   367
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   368
declare choose_pos.simps [simp del]
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   369
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   370
lemma choose_pos_choose: "finite A \<Longrightarrow> choose_pos A (choose A) = 0"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   371
by (simp add: choose_pos.simps)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   372
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   373
lemma inj_on_choose_pos [OF refl]:
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   374
  "\<lbrakk>card A = n; finite A\<rbrakk> \<Longrightarrow> inj_on (choose_pos A) A"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   375
 apply (induct n arbitrary: A)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   376
  apply simp
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   377
 apply (case_tac "A = {}", simp)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   378
 apply (frule (1) choose_in)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   379
 apply (rule inj_onI)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   380
 apply (drule_tac x="A - {choose A}" in meta_spec, simp)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   381
 apply (simp add: choose_pos.simps)
62390
842917225d56 more canonical names
nipkow
parents: 62175
diff changeset
   382
 apply (simp split: if_split_asm)
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   383
 apply (erule (1) inj_onD, simp, simp)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   384
done
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   385
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   386
lemma choose_pos_bounded [OF refl]:
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   387
  "\<lbrakk>card A = n; finite A; x \<in> A\<rbrakk> \<Longrightarrow> choose_pos A x < n"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   388
apply (induct n arbitrary: A)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   389
apply simp
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   390
 apply (case_tac "A = {}", simp)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   391
 apply (frule (1) choose_in)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   392
apply (subst choose_pos.simps)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   393
apply simp
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   394
done
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   395
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   396
lemma choose_pos_lessD:
41182
717404c7d59a add notsqsubseteq syntax
huffman
parents: 40888
diff changeset
   397
  "\<lbrakk>choose_pos A x < choose_pos A y; finite A; x \<in> A; y \<in> A\<rbrakk> \<Longrightarrow> x \<notsqsubseteq> y"
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   398
 apply (induct A x arbitrary: y rule: choose_pos.induct)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   399
 apply simp
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   400
 apply (case_tac "x = choose A")
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   401
  apply simp
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   402
  apply (rule notI)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   403
  apply (frule (2) maximal_choose)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   404
  apply simp
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   405
 apply (case_tac "y = choose A")
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   406
  apply (simp add: choose_pos_choose)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   407
 apply (drule_tac x=y in meta_spec)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   408
 apply simp
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   409
 apply (erule meta_mp)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   410
 apply (simp add: choose_pos.simps)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   411
done
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   412
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61605
diff changeset
   413
subsubsection \<open>Compact basis take function\<close>
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   414
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   415
primrec
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   416
  cb_take :: "nat \<Rightarrow> 'a compact_basis \<Rightarrow> 'a compact_basis" where
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   417
  "cb_take 0 = (\<lambda>x. compact_bot)"
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   418
| "cb_take (Suc n) = (\<lambda>a. Abs_compact_basis (approx n\<cdot>(Rep_compact_basis a)))"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   419
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   420
declare cb_take.simps [simp del]
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   421
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   422
lemma cb_take_zero [simp]: "cb_take 0 a = compact_bot"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   423
by (simp only: cb_take.simps)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   424
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   425
lemma Rep_cb_take:
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   426
  "Rep_compact_basis (cb_take (Suc n) a) = approx n\<cdot>(Rep_compact_basis a)"
41370
17b09240893c declare more simp rules, rewrite proofs in Isar-style
huffman
parents: 41295
diff changeset
   427
by (simp add: cb_take.simps(2))
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   428
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   429
lemmas approx_Rep_compact_basis = Rep_cb_take [symmetric]
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   430
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   431
lemma cb_take_covers: "\<exists>n. cb_take n x = x"
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   432
apply (subgoal_tac "\<exists>n. cb_take (Suc n) x = x", fast)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   433
apply (simp add: Rep_compact_basis_inject [symmetric])
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   434
apply (simp add: Rep_cb_take)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   435
apply (rule compact_eq_approx)
41370
17b09240893c declare more simp rules, rewrite proofs in Isar-style
huffman
parents: 41295
diff changeset
   436
apply (rule Rep_compact_basis')
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   437
done
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   438
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   439
lemma cb_take_less: "cb_take n x \<sqsubseteq> x"
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   440
unfolding compact_le_def
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   441
by (cases n, simp, simp add: Rep_cb_take approx_below)
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   442
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   443
lemma cb_take_idem: "cb_take n (cb_take n x) = cb_take n x"
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   444
unfolding Rep_compact_basis_inject [symmetric]
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   445
by (cases n, simp, simp add: Rep_cb_take approx_idem)
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   446
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   447
lemma cb_take_mono: "x \<sqsubseteq> y \<Longrightarrow> cb_take n x \<sqsubseteq> cb_take n y"
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   448
unfolding compact_le_def
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   449
by (cases n, simp, simp add: Rep_cb_take monofun_cfun_arg)
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   450
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   451
lemma cb_take_chain_le: "m \<le> n \<Longrightarrow> cb_take m x \<sqsubseteq> cb_take n x"
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   452
unfolding compact_le_def
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   453
apply (cases m, simp, cases n, simp)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   454
apply (simp add: Rep_cb_take, rule chain_mono, simp, simp)
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   455
done
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   456
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   457
lemma finite_range_cb_take: "finite (range (cb_take n))"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   458
apply (cases n)
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   459
apply (subgoal_tac "range (cb_take 0) = {compact_bot}", simp, force)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   460
apply (rule finite_imageD [where f="Rep_compact_basis"])
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   461
apply (rule finite_subset [where B="range (\<lambda>x. approx (n - 1)\<cdot>x)"])
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   462
apply (clarsimp simp add: Rep_cb_take)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   463
apply (rule finite_range_approx)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   464
apply (rule inj_onI, simp add: Rep_compact_basis_inject)
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   465
done
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   466
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61605
diff changeset
   467
subsubsection \<open>Rank of basis elements\<close>
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   468
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   469
definition
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   470
  rank :: "'a compact_basis \<Rightarrow> nat"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   471
where
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   472
  "rank x = (LEAST n. cb_take n x = x)"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   473
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   474
lemma compact_approx_rank: "cb_take (rank x) x = x"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   475
unfolding rank_def
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   476
apply (rule LeastI_ex)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   477
apply (rule cb_take_covers)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   478
done
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   479
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   480
lemma rank_leD: "rank x \<le> n \<Longrightarrow> cb_take n x = x"
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 30729
diff changeset
   481
apply (rule below_antisym [OF cb_take_less])
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   482
apply (subst compact_approx_rank [symmetric])
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   483
apply (erule cb_take_chain_le)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   484
done
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   485
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   486
lemma rank_leI: "cb_take n x = x \<Longrightarrow> rank x \<le> n"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   487
unfolding rank_def by (rule Least_le)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   488
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   489
lemma rank_le_iff: "rank x \<le> n \<longleftrightarrow> cb_take n x = x"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   490
by (rule iffI [OF rank_leD rank_leI])
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   491
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   492
lemma rank_compact_bot [simp]: "rank compact_bot = 0"
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   493
using rank_leI [of 0 compact_bot] by simp
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   494
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   495
lemma rank_eq_0_iff [simp]: "rank x = 0 \<longleftrightarrow> x = compact_bot"
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   496
using rank_le_iff [of x 0] by auto
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   497
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   498
definition
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   499
  rank_le :: "'a compact_basis \<Rightarrow> 'a compact_basis set"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   500
where
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   501
  "rank_le x = {y. rank y \<le> rank x}"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   502
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   503
definition
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   504
  rank_lt :: "'a compact_basis \<Rightarrow> 'a compact_basis set"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   505
where
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   506
  "rank_lt x = {y. rank y < rank x}"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   507
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   508
definition
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   509
  rank_eq :: "'a compact_basis \<Rightarrow> 'a compact_basis set"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   510
where
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   511
  "rank_eq x = {y. rank y = rank x}"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   512
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   513
lemma rank_eq_cong: "rank x = rank y \<Longrightarrow> rank_eq x = rank_eq y"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   514
unfolding rank_eq_def by simp
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   515
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   516
lemma rank_lt_cong: "rank x = rank y \<Longrightarrow> rank_lt x = rank_lt y"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   517
unfolding rank_lt_def by simp
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   518
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   519
lemma rank_eq_subset: "rank_eq x \<subseteq> rank_le x"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   520
unfolding rank_eq_def rank_le_def by auto
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   521
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   522
lemma rank_lt_subset: "rank_lt x \<subseteq> rank_le x"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   523
unfolding rank_lt_def rank_le_def by auto
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   524
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   525
lemma finite_rank_le: "finite (rank_le x)"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   526
unfolding rank_le_def
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   527
apply (rule finite_subset [where B="range (cb_take (rank x))"])
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   528
apply clarify
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   529
apply (rule range_eqI)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   530
apply (erule rank_leD [symmetric])
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   531
apply (rule finite_range_cb_take)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   532
done
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   533
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   534
lemma finite_rank_eq: "finite (rank_eq x)"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   535
by (rule finite_subset [OF rank_eq_subset finite_rank_le])
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   536
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   537
lemma finite_rank_lt: "finite (rank_lt x)"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   538
by (rule finite_subset [OF rank_lt_subset finite_rank_le])
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   539
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   540
lemma rank_lt_Int_rank_eq: "rank_lt x \<inter> rank_eq x = {}"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   541
unfolding rank_lt_def rank_eq_def rank_le_def by auto
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   542
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   543
lemma rank_lt_Un_rank_eq: "rank_lt x \<union> rank_eq x = rank_le x"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   544
unfolding rank_lt_def rank_eq_def rank_le_def by auto
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   545
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61605
diff changeset
   546
subsubsection \<open>Sequencing basis elements\<close>
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   547
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   548
definition
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   549
  place :: "'a compact_basis \<Rightarrow> nat"
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   550
where
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   551
  "place x = card (rank_lt x) + choose_pos (rank_eq x) x"
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   552
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   553
lemma place_bounded: "place x < card (rank_le x)"
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   554
unfolding place_def
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   555
 apply (rule ord_less_eq_trans)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   556
  apply (rule add_strict_left_mono)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   557
  apply (rule choose_pos_bounded)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   558
   apply (rule finite_rank_eq)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   559
  apply (simp add: rank_eq_def)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   560
 apply (subst card_Un_disjoint [symmetric])
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   561
    apply (rule finite_rank_lt)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   562
   apply (rule finite_rank_eq)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   563
  apply (rule rank_lt_Int_rank_eq)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   564
 apply (simp add: rank_lt_Un_rank_eq)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   565
done
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   566
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   567
lemma place_ge: "card (rank_lt x) \<le> place x"
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   568
unfolding place_def by simp
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   569
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   570
lemma place_rank_mono:
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   571
  fixes x y :: "'a compact_basis"
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   572
  shows "rank x < rank y \<Longrightarrow> place x < place y"
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   573
apply (rule less_le_trans [OF place_bounded])
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   574
apply (rule order_trans [OF _ place_ge])
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   575
apply (rule card_mono)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   576
apply (rule finite_rank_lt)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   577
apply (simp add: rank_le_def rank_lt_def subset_eq)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   578
done
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   579
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   580
lemma place_eqD: "place x = place y \<Longrightarrow> x = y"
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   581
 apply (rule linorder_cases [where x="rank x" and y="rank y"])
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   582
   apply (drule place_rank_mono, simp)
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   583
  apply (simp add: place_def)
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   584
  apply (rule inj_on_choose_pos [where A="rank_eq x", THEN inj_onD])
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   585
     apply (rule finite_rank_eq)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   586
    apply (simp cong: rank_lt_cong rank_eq_cong)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   587
   apply (simp add: rank_eq_def)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   588
  apply (simp add: rank_eq_def)
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   589
 apply (drule place_rank_mono, simp)
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   590
done
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   591
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   592
lemma inj_place: "inj place"
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   593
by (rule inj_onI, erule place_eqD)
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   594
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61605
diff changeset
   595
subsubsection \<open>Embedding and projection on basis elements\<close>
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   596
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   597
definition
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   598
  sub :: "'a compact_basis \<Rightarrow> 'a compact_basis"
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   599
where
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   600
  "sub x = (case rank x of 0 \<Rightarrow> compact_bot | Suc k \<Rightarrow> cb_take k x)"
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   601
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   602
lemma rank_sub_less: "x \<noteq> compact_bot \<Longrightarrow> rank (sub x) < rank x"
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   603
unfolding sub_def
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   604
apply (cases "rank x", simp)
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   605
apply (simp add: less_Suc_eq_le)
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   606
apply (rule rank_leI)
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   607
apply (rule cb_take_idem)
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   608
done
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   609
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   610
lemma place_sub_less: "x \<noteq> compact_bot \<Longrightarrow> place (sub x) < place x"
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   611
apply (rule place_rank_mono)
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   612
apply (erule rank_sub_less)
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   613
done
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   614
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   615
lemma sub_below: "sub x \<sqsubseteq> x"
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   616
unfolding sub_def by (cases "rank x", simp_all add: cb_take_less)
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   617
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   618
lemma rank_less_imp_below_sub: "\<lbrakk>x \<sqsubseteq> y; rank x < rank y\<rbrakk> \<Longrightarrow> x \<sqsubseteq> sub y"
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   619
unfolding sub_def
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   620
apply (cases "rank y", simp)
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   621
apply (simp add: less_Suc_eq_le)
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   622
apply (subgoal_tac "cb_take nat x \<sqsubseteq> cb_take nat y")
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   623
apply (simp add: rank_leD)
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   624
apply (erule cb_take_mono)
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   625
done
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   626
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   627
function
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   628
  basis_emb :: "'a compact_basis \<Rightarrow> ubasis"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   629
where
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   630
  "basis_emb x = (if x = compact_bot then 0 else
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   631
    node (place x) (basis_emb (sub x))
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   632
      (basis_emb ` {y. place y < place x \<and> x \<sqsubseteq> y}))"
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   633
by auto
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   634
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   635
termination basis_emb
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   636
apply (relation "measure place", simp)
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   637
apply (simp add: place_sub_less)
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   638
apply simp
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   639
done
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   640
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   641
declare basis_emb.simps [simp del]
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   642
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   643
lemma basis_emb_compact_bot [simp]: "basis_emb compact_bot = 0"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   644
by (simp add: basis_emb.simps)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   645
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   646
lemma fin1: "finite {y. place y < place x \<and> x \<sqsubseteq> y}"
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   647
apply (subst Collect_conj_eq)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   648
apply (rule finite_Int)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   649
apply (rule disjI1)
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   650
apply (subgoal_tac "finite (place -` {n. n < place x})", simp)
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   651
apply (rule finite_vimageI [OF _ inj_place])
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   652
apply (simp add: lessThan_def [symmetric])
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   653
done
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   654
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   655
lemma fin2: "finite (basis_emb ` {y. place y < place x \<and> x \<sqsubseteq> y})"
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   656
by (rule finite_imageI [OF fin1])
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   657
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   658
lemma rank_place_mono:
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   659
  "\<lbrakk>place x < place y; x \<sqsubseteq> y\<rbrakk> \<Longrightarrow> rank x < rank y"
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   660
apply (rule linorder_cases, assumption)
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   661
apply (simp add: place_def cong: rank_lt_cong rank_eq_cong)
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   662
apply (drule choose_pos_lessD)
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   663
apply (rule finite_rank_eq)
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   664
apply (simp add: rank_eq_def)
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   665
apply (simp add: rank_eq_def)
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   666
apply simp
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   667
apply (drule place_rank_mono, simp)
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   668
done
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   669
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   670
lemma basis_emb_mono:
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   671
  "x \<sqsubseteq> y \<Longrightarrow> ubasis_le (basis_emb x) (basis_emb y)"
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 33071
diff changeset
   672
proof (induct "max (place x) (place y)" arbitrary: x y rule: less_induct)
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 33071
diff changeset
   673
  case less
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   674
  show ?case proof (rule linorder_cases)
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   675
    assume "place x < place y"
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   676
    then have "rank x < rank y"
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61605
diff changeset
   677
      using \<open>x \<sqsubseteq> y\<close> by (rule rank_place_mono)
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61605
diff changeset
   678
    with \<open>place x < place y\<close> show ?case
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   679
      apply (case_tac "y = compact_bot", simp)
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   680
      apply (simp add: basis_emb.simps [of y])
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   681
      apply (rule ubasis_le_trans [OF _ ubasis_le_lower [OF fin2]])
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 33071
diff changeset
   682
      apply (rule less)
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   683
       apply (simp add: less_max_iff_disj)
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   684
       apply (erule place_sub_less)
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61605
diff changeset
   685
      apply (erule rank_less_imp_below_sub [OF \<open>x \<sqsubseteq> y\<close>])
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   686
      done
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   687
  next
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   688
    assume "place x = place y"
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   689
    hence "x = y" by (rule place_eqD)
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   690
    thus ?case by (simp add: ubasis_le_refl)
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   691
  next
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   692
    assume "place x > place y"
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61605
diff changeset
   693
    with \<open>x \<sqsubseteq> y\<close> show ?case
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   694
      apply (case_tac "x = compact_bot", simp add: ubasis_le_minimal)
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   695
      apply (simp add: basis_emb.simps [of x])
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   696
      apply (rule ubasis_le_upper [OF fin2], simp)
34915
7894c7dab132 Adapted to changes in induct method.
berghofe
parents: 33071
diff changeset
   697
      apply (rule less)
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   698
       apply (simp add: less_max_iff_disj)
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   699
       apply (erule place_sub_less)
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 30729
diff changeset
   700
      apply (erule rev_below_trans)
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   701
      apply (rule sub_below)
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   702
      done
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   703
  qed
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   704
qed
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   705
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   706
lemma inj_basis_emb: "inj basis_emb"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   707
 apply (rule inj_onI)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   708
 apply (case_tac "x = compact_bot")
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   709
  apply (case_tac [!] "y = compact_bot")
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   710
    apply simp
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   711
   apply (simp add: basis_emb.simps)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   712
  apply (simp add: basis_emb.simps)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   713
 apply (simp add: basis_emb.simps)
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   714
 apply (simp add: fin2 inj_eq [OF inj_place])
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   715
done
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   716
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   717
definition
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   718
  basis_prj :: "ubasis \<Rightarrow> 'a compact_basis"
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   719
where
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   720
  "basis_prj x = inv basis_emb
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   721
    (ubasis_until (\<lambda>x. x \<in> range (basis_emb :: 'a compact_basis \<Rightarrow> ubasis)) x)"
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   722
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   723
lemma basis_prj_basis_emb: "\<And>x. basis_prj (basis_emb x) = x"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   724
unfolding basis_prj_def
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   725
 apply (subst ubasis_until_same)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   726
  apply (rule rangeI)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   727
 apply (rule inv_f_f)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   728
 apply (rule inj_basis_emb)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   729
done
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   730
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   731
lemma basis_prj_node:
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   732
  "\<lbrakk>finite S; node i a S \<notin> range (basis_emb :: 'a compact_basis \<Rightarrow> nat)\<rbrakk>
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   733
    \<Longrightarrow> basis_prj (node i a S) = (basis_prj a :: 'a compact_basis)"
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   734
unfolding basis_prj_def by simp
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   735
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   736
lemma basis_prj_0: "basis_prj 0 = compact_bot"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   737
apply (subst basis_emb_compact_bot [symmetric])
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   738
apply (rule basis_prj_basis_emb)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   739
done
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   740
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   741
lemma node_eq_basis_emb_iff:
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   742
  "finite S \<Longrightarrow> node i a S = basis_emb x \<longleftrightarrow>
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   743
    x \<noteq> compact_bot \<and> i = place x \<and> a = basis_emb (sub x) \<and>
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   744
        S = basis_emb ` {y. place y < place x \<and> x \<sqsubseteq> y}"
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   745
apply (cases "x = compact_bot", simp)
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   746
apply (simp add: basis_emb.simps [of x])
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   747
apply (simp add: fin2)
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   748
done
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   749
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   750
lemma basis_prj_mono: "ubasis_le a b \<Longrightarrow> basis_prj a \<sqsubseteq> basis_prj b"
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   751
proof (induct a b rule: ubasis_le.induct)
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 30729
diff changeset
   752
  case (ubasis_le_refl a) show ?case by (rule below_refl)
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   753
next
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 30729
diff changeset
   754
  case (ubasis_le_trans a b c) thus ?case by - (rule below_trans)
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   755
next
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   756
  case (ubasis_le_lower S a i) thus ?case
30561
5e6088e1d6df clean up proofs
huffman
parents: 30505
diff changeset
   757
    apply (cases "node i a S \<in> range (basis_emb :: 'a compact_basis \<Rightarrow> nat)")
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   758
     apply (erule rangeE, rename_tac x)
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   759
     apply (simp add: basis_prj_basis_emb)
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   760
     apply (simp add: node_eq_basis_emb_iff)
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   761
     apply (simp add: basis_prj_basis_emb)
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   762
     apply (rule sub_below)
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   763
    apply (simp add: basis_prj_node)
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   764
    done
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   765
next
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   766
  case (ubasis_le_upper S b a i) thus ?case
30561
5e6088e1d6df clean up proofs
huffman
parents: 30505
diff changeset
   767
    apply (cases "node i a S \<in> range (basis_emb :: 'a compact_basis \<Rightarrow> nat)")
30505
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   768
     apply (erule rangeE, rename_tac x)
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   769
     apply (simp add: basis_prj_basis_emb)
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   770
     apply (clarsimp simp add: node_eq_basis_emb_iff)
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   771
     apply (simp add: basis_prj_basis_emb)
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   772
    apply (simp add: basis_prj_node)
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   773
    done
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   774
qed
110e59507eec introduce new helper functions; clean up proofs
huffman
parents: 29252
diff changeset
   775
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   776
lemma basis_emb_prj_less: "ubasis_le (basis_emb (basis_prj x)) x"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   777
unfolding basis_prj_def
33071
362f59fe5092 renamed f_inv_onto_f to f_inv_into_f (cf. 764547b68538);
wenzelm
parents: 32997
diff changeset
   778
 apply (subst f_inv_into_f [where f=basis_emb])
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   779
  apply (rule ubasis_until)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   780
  apply (rule range_eqI [where x=compact_bot])
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   781
  apply simp
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   782
 apply (rule ubasis_until_less)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   783
done
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   784
41286
3d7685a4a5ff reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents: 41182
diff changeset
   785
lemma ideal_completion:
3d7685a4a5ff reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents: 41182
diff changeset
   786
  "ideal_completion below Rep_compact_basis (approximants :: 'a \<Rightarrow> _)"
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   787
proof
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   788
  fix w :: "'a"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   789
  show "below.ideal (approximants w)"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   790
  proof (rule below.idealI)
41370
17b09240893c declare more simp rules, rewrite proofs in Isar-style
huffman
parents: 41295
diff changeset
   791
    have "Abs_compact_basis (approx 0\<cdot>w) \<in> approximants w"
17b09240893c declare more simp rules, rewrite proofs in Isar-style
huffman
parents: 41295
diff changeset
   792
      by (simp add: approximants_def approx_below)
17b09240893c declare more simp rules, rewrite proofs in Isar-style
huffman
parents: 41295
diff changeset
   793
    thus "\<exists>x. x \<in> approximants w" ..
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   794
  next
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   795
    fix x y :: "'a compact_basis"
41370
17b09240893c declare more simp rules, rewrite proofs in Isar-style
huffman
parents: 41295
diff changeset
   796
    assume x: "x \<in> approximants w" and y: "y \<in> approximants w"
17b09240893c declare more simp rules, rewrite proofs in Isar-style
huffman
parents: 41295
diff changeset
   797
    obtain i where i: "approx i\<cdot>(Rep_compact_basis x) = Rep_compact_basis x"
17b09240893c declare more simp rules, rewrite proofs in Isar-style
huffman
parents: 41295
diff changeset
   798
      using compact_eq_approx Rep_compact_basis' by fast
17b09240893c declare more simp rules, rewrite proofs in Isar-style
huffman
parents: 41295
diff changeset
   799
    obtain j where j: "approx j\<cdot>(Rep_compact_basis y) = Rep_compact_basis y"
17b09240893c declare more simp rules, rewrite proofs in Isar-style
huffman
parents: 41295
diff changeset
   800
      using compact_eq_approx Rep_compact_basis' by fast
17b09240893c declare more simp rules, rewrite proofs in Isar-style
huffman
parents: 41295
diff changeset
   801
    let ?z = "Abs_compact_basis (approx (max i j)\<cdot>w)"
17b09240893c declare more simp rules, rewrite proofs in Isar-style
huffman
parents: 41295
diff changeset
   802
    have "?z \<in> approximants w"
17b09240893c declare more simp rules, rewrite proofs in Isar-style
huffman
parents: 41295
diff changeset
   803
      by (simp add: approximants_def approx_below)
17b09240893c declare more simp rules, rewrite proofs in Isar-style
huffman
parents: 41295
diff changeset
   804
    moreover from x y have "x \<sqsubseteq> ?z \<and> y \<sqsubseteq> ?z"
17b09240893c declare more simp rules, rewrite proofs in Isar-style
huffman
parents: 41295
diff changeset
   805
      by (simp add: approximants_def compact_le_def)
54863
82acc20ded73 prefer more canonical names for lemmas on min/max
haftmann
parents: 49834
diff changeset
   806
         (metis i j monofun_cfun chain_mono chain_approx max.cobounded1 max.cobounded2)
41370
17b09240893c declare more simp rules, rewrite proofs in Isar-style
huffman
parents: 41295
diff changeset
   807
    ultimately show "\<exists>z \<in> approximants w. x \<sqsubseteq> z \<and> y \<sqsubseteq> z" ..
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   808
  next
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   809
    fix x y :: "'a compact_basis"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   810
    assume "x \<sqsubseteq> y" "y \<in> approximants w" thus "x \<in> approximants w"
41370
17b09240893c declare more simp rules, rewrite proofs in Isar-style
huffman
parents: 41295
diff changeset
   811
      unfolding approximants_def compact_le_def
17b09240893c declare more simp rules, rewrite proofs in Isar-style
huffman
parents: 41295
diff changeset
   812
      by (auto elim: below_trans)
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   813
  qed
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   814
next
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   815
  fix Y :: "nat \<Rightarrow> 'a"
41370
17b09240893c declare more simp rules, rewrite proofs in Isar-style
huffman
parents: 41295
diff changeset
   816
  assume "chain Y"
17b09240893c declare more simp rules, rewrite proofs in Isar-style
huffman
parents: 41295
diff changeset
   817
  thus "approximants (\<Squnion>i. Y i) = (\<Union>i. approximants (Y i))"
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   818
    unfolding approximants_def
41370
17b09240893c declare more simp rules, rewrite proofs in Isar-style
huffman
parents: 41295
diff changeset
   819
    by (auto simp add: compact_below_lub_iff)
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   820
next
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   821
  fix a :: "'a compact_basis"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   822
  show "approximants (Rep_compact_basis a) = {b. b \<sqsubseteq> a}"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   823
    unfolding approximants_def compact_le_def ..
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   824
next
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   825
  fix x y :: "'a"
41370
17b09240893c declare more simp rules, rewrite proofs in Isar-style
huffman
parents: 41295
diff changeset
   826
  assume "approximants x \<subseteq> approximants y"
17b09240893c declare more simp rules, rewrite proofs in Isar-style
huffman
parents: 41295
diff changeset
   827
  hence "\<forall>z. compact z \<longrightarrow> z \<sqsubseteq> x \<longrightarrow> z \<sqsubseteq> y"
17b09240893c declare more simp rules, rewrite proofs in Isar-style
huffman
parents: 41295
diff changeset
   828
    by (simp add: approximants_def subset_eq)
17b09240893c declare more simp rules, rewrite proofs in Isar-style
huffman
parents: 41295
diff changeset
   829
       (metis Abs_compact_basis_inverse')
17b09240893c declare more simp rules, rewrite proofs in Isar-style
huffman
parents: 41295
diff changeset
   830
  hence "(\<Squnion>i. approx i\<cdot>x) \<sqsubseteq> y"
17b09240893c declare more simp rules, rewrite proofs in Isar-style
huffman
parents: 41295
diff changeset
   831
    by (simp add: lub_below approx_below)
17b09240893c declare more simp rules, rewrite proofs in Isar-style
huffman
parents: 41295
diff changeset
   832
  thus "x \<sqsubseteq> y"
17b09240893c declare more simp rules, rewrite proofs in Isar-style
huffman
parents: 41295
diff changeset
   833
    by (simp add: lub_distribs)
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   834
next
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   835
  show "\<exists>f::'a compact_basis \<Rightarrow> nat. inj f"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   836
    by (rule exI, rule inj_place)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   837
qed
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   838
41286
3d7685a4a5ff reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents: 41182
diff changeset
   839
end
3d7685a4a5ff reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents: 41182
diff changeset
   840
61605
1bf7b186542e qualifier is mandatory by default;
wenzelm
parents: 61169
diff changeset
   841
interpretation compact_basis:
41286
3d7685a4a5ff reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents: 41182
diff changeset
   842
  ideal_completion below Rep_compact_basis
3d7685a4a5ff reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents: 41182
diff changeset
   843
    "approximants :: 'a::bifinite \<Rightarrow> 'a compact_basis set"
3d7685a4a5ff reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents: 41182
diff changeset
   844
proof -
3d7685a4a5ff reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents: 41182
diff changeset
   845
  obtain a :: "nat \<Rightarrow> 'a \<rightarrow> 'a" where "approx_chain a"
3d7685a4a5ff reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents: 41182
diff changeset
   846
    using bifinite ..
3d7685a4a5ff reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents: 41182
diff changeset
   847
  hence "bifinite_approx_chain a"
3d7685a4a5ff reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents: 41182
diff changeset
   848
    unfolding bifinite_approx_chain_def .
3d7685a4a5ff reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents: 41182
diff changeset
   849
  thus "ideal_completion below Rep_compact_basis (approximants :: 'a \<Rightarrow> _)"
3d7685a4a5ff reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents: 41182
diff changeset
   850
    by (rule bifinite_approx_chain.ideal_completion)
3d7685a4a5ff reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents: 41182
diff changeset
   851
qed
3d7685a4a5ff reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents: 41182
diff changeset
   852
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61605
diff changeset
   853
subsubsection \<open>EP-pair from any bifinite domain into \emph{udom}\<close>
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   854
41286
3d7685a4a5ff reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents: 41182
diff changeset
   855
context bifinite_approx_chain begin
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   856
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   857
definition
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   858
  udom_emb :: "'a \<rightarrow> udom"
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   859
where
41394
51c866d1b53b rename function ideal_completion.basis_fun to ideal_completion.extension
huffman
parents: 41370
diff changeset
   860
  "udom_emb = compact_basis.extension (\<lambda>x. udom_principal (basis_emb x))"
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   861
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   862
definition
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   863
  udom_prj :: "udom \<rightarrow> 'a"
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   864
where
41394
51c866d1b53b rename function ideal_completion.basis_fun to ideal_completion.extension
huffman
parents: 41370
diff changeset
   865
  "udom_prj = udom.extension (\<lambda>x. Rep_compact_basis (basis_prj x))"
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   866
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   867
lemma udom_emb_principal:
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   868
  "udom_emb\<cdot>(Rep_compact_basis x) = udom_principal (basis_emb x)"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   869
unfolding udom_emb_def
41394
51c866d1b53b rename function ideal_completion.basis_fun to ideal_completion.extension
huffman
parents: 41370
diff changeset
   870
apply (rule compact_basis.extension_principal)
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   871
apply (rule udom.principal_mono)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   872
apply (erule basis_emb_mono)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   873
done
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   874
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   875
lemma udom_prj_principal:
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   876
  "udom_prj\<cdot>(udom_principal x) = Rep_compact_basis (basis_prj x)"
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   877
unfolding udom_prj_def
41394
51c866d1b53b rename function ideal_completion.basis_fun to ideal_completion.extension
huffman
parents: 41370
diff changeset
   878
apply (rule udom.extension_principal)
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   879
apply (rule compact_basis.principal_mono)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   880
apply (erule basis_prj_mono)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   881
done
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   882
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   883
lemma ep_pair_udom: "ep_pair udom_emb udom_prj"
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 59667
diff changeset
   884
 apply standard
27411
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   885
  apply (rule compact_basis.principal_induct, simp)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   886
  apply (simp add: udom_emb_principal udom_prj_principal)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   887
  apply (simp add: basis_prj_basis_emb)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   888
 apply (rule udom.principal_induct, simp)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   889
 apply (simp add: udom_emb_principal udom_prj_principal)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   890
 apply (rule basis_emb_prj_less)
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   891
done
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   892
60fad3219d32 universal bifinite domain
huffman
parents:
diff changeset
   893
end
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   894
41286
3d7685a4a5ff reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents: 41182
diff changeset
   895
abbreviation "udom_emb \<equiv> bifinite_approx_chain.udom_emb"
3d7685a4a5ff reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents: 41182
diff changeset
   896
abbreviation "udom_prj \<equiv> bifinite_approx_chain.udom_prj"
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   897
41286
3d7685a4a5ff reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents: 41182
diff changeset
   898
lemmas ep_pair_udom =
3d7685a4a5ff reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents: 41182
diff changeset
   899
  bifinite_approx_chain.ep_pair_udom [unfolded bifinite_approx_chain_def]
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   900
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61605
diff changeset
   901
subsection \<open>Chain of approx functions for type \emph{udom}\<close>
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   902
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   903
definition
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   904
  udom_approx :: "nat \<Rightarrow> udom \<rightarrow> udom"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   905
where
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   906
  "udom_approx i =
41394
51c866d1b53b rename function ideal_completion.basis_fun to ideal_completion.extension
huffman
parents: 41370
diff changeset
   907
    udom.extension (\<lambda>x. udom_principal (ubasis_until (\<lambda>y. y \<le> i) x))"
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   908
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   909
lemma udom_approx_mono:
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   910
  "ubasis_le a b \<Longrightarrow>
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   911
    udom_principal (ubasis_until (\<lambda>y. y \<le> i) a) \<sqsubseteq>
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   912
    udom_principal (ubasis_until (\<lambda>y. y \<le> i) b)"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   913
apply (rule udom.principal_mono)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   914
apply (rule ubasis_until_mono)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   915
apply (frule (2) order_less_le_trans [OF node_gt2])
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   916
apply (erule order_less_imp_le)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   917
apply assumption
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   918
done
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   919
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   920
lemma adm_mem_finite: "\<lbrakk>cont f; finite S\<rbrakk> \<Longrightarrow> adm (\<lambda>x. f x \<in> S)"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   921
by (erule adm_subst, induct set: finite, simp_all)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   922
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   923
lemma udom_approx_principal:
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   924
  "udom_approx i\<cdot>(udom_principal x) =
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   925
    udom_principal (ubasis_until (\<lambda>y. y \<le> i) x)"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   926
unfolding udom_approx_def
41394
51c866d1b53b rename function ideal_completion.basis_fun to ideal_completion.extension
huffman
parents: 41370
diff changeset
   927
apply (rule udom.extension_principal)
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   928
apply (erule udom_approx_mono)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   929
done
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   930
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   931
lemma finite_deflation_udom_approx: "finite_deflation (udom_approx i)"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   932
proof
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   933
  fix x show "udom_approx i\<cdot>(udom_approx i\<cdot>x) = udom_approx i\<cdot>x"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   934
    by (induct x rule: udom.principal_induct, simp)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   935
       (simp add: udom_approx_principal ubasis_until_idem)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   936
next
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   937
  fix x show "udom_approx i\<cdot>x \<sqsubseteq> x"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   938
    by (induct x rule: udom.principal_induct, simp)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   939
       (simp add: udom_approx_principal ubasis_until_less)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   940
next
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   941
  have *: "finite (range (\<lambda>x. udom_principal (ubasis_until (\<lambda>y. y \<le> i) x)))"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   942
    apply (subst range_composition [where f=udom_principal])
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   943
    apply (simp add: finite_range_ubasis_until)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   944
    done
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   945
  show "finite {x. udom_approx i\<cdot>x = x}"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   946
    apply (rule finite_range_imp_finite_fixes)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   947
    apply (rule rev_finite_subset [OF *])
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   948
    apply (clarsimp, rename_tac x)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   949
    apply (induct_tac x rule: udom.principal_induct)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   950
    apply (simp add: adm_mem_finite *)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   951
    apply (simp add: udom_approx_principal)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   952
    done
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   953
qed
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   954
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   955
interpretation udom_approx: finite_deflation "udom_approx i"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   956
by (rule finite_deflation_udom_approx)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   957
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   958
lemma chain_udom_approx [simp]: "chain (\<lambda>i. udom_approx i)"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   959
unfolding udom_approx_def
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   960
apply (rule chainI)
41394
51c866d1b53b rename function ideal_completion.basis_fun to ideal_completion.extension
huffman
parents: 41370
diff changeset
   961
apply (rule udom.extension_mono)
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   962
apply (erule udom_approx_mono)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   963
apply (erule udom_approx_mono)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   964
apply (rule udom.principal_mono)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   965
apply (rule ubasis_until_chain, simp)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   966
done
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   967
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   968
lemma lub_udom_approx [simp]: "(\<Squnion>i. udom_approx i) = ID"
40002
c5b5f7a3a3b1 new theorem names: fun_below_iff, fun_belowI, cfun_eq_iff, cfun_eqI, cfun_below_iff, cfun_belowI
huffman
parents: 39984
diff changeset
   969
apply (rule cfun_eqI, simp add: contlub_cfun_fun)
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   970
apply (rule below_antisym)
40500
ee9c8d36318e add lemmas lub_below, below_lub; simplify some proofs; remove some unused lemmas
huffman
parents: 40002
diff changeset
   971
apply (rule lub_below)
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   972
apply (simp)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   973
apply (rule udom_approx.below)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   974
apply (rule_tac x=x in udom.principal_induct)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   975
apply (simp add: lub_distribs)
40500
ee9c8d36318e add lemmas lub_below, below_lub; simplify some proofs; remove some unused lemmas
huffman
parents: 40002
diff changeset
   976
apply (rule_tac i=a in below_lub)
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   977
apply simp
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   978
apply (simp add: udom_approx_principal)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   979
apply (simp add: ubasis_until_same ubasis_le_refl)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   980
done
65552
f533820e7248 theories "GCD" and "Binomial" are already included in "Main": this avoids improper imports in applications;
wenzelm
parents: 64267
diff changeset
   981
41286
3d7685a4a5ff reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents: 41182
diff changeset
   982
lemma udom_approx [simp]: "approx_chain udom_approx"
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   983
proof
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   984
  show "chain (\<lambda>i. udom_approx i)"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   985
    by (rule chain_udom_approx)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   986
  show "(\<Squnion>i. udom_approx i) = ID"
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   987
    by (rule lub_udom_approx)
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   988
qed
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   989
41286
3d7685a4a5ff reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents: 41182
diff changeset
   990
instance udom :: bifinite
61169
4de9ff3ea29a tuned proofs -- less legacy;
wenzelm
parents: 59667
diff changeset
   991
  by standard (fast intro: udom_approx)
41286
3d7685a4a5ff reintroduce 'bifinite' class, now with existentially-quantified approx function (cf. b525988432e9)
huffman
parents: 41182
diff changeset
   992
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   993
hide_const (open) node
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   994
65552
f533820e7248 theories "GCD" and "Binomial" are already included in "Main": this avoids improper imports in applications;
wenzelm
parents: 64267
diff changeset
   995
notation binomial  (infixl "choose" 65)
f533820e7248 theories "GCD" and "Binomial" are already included in "Main": this avoids improper imports in applications;
wenzelm
parents: 64267
diff changeset
   996
39974
b525988432e9 major reorganization/simplification of HOLCF type classes:
huffman
parents: 36452
diff changeset
   997
end