src/HOLCF/Library/List_Cpo.thy
author huffman
Sat, 04 Sep 2010 07:26:34 -0700
changeset 39143 d80990d8b909
child 39199 720112792ba0
permissions -rw-r--r--
add List_Cpo.thy to HOLCF/Library
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
39143
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
     1
(*  Title:      HOLCF/Library/List_Cpo.thy
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
     2
    Author:     Brian Huffman
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
     3
*)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
     4
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
     5
header {* Lists as a complete partial order *}
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
     6
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
     7
theory List_Cpo
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
     8
imports HOLCF
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
     9
begin
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    10
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    11
subsection {* Lists are a partial order *}
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    12
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    13
instantiation list :: (po) po
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    14
begin
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    15
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    16
definition
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    17
  "xs \<sqsubseteq> ys \<longleftrightarrow> list_all2 (op \<sqsubseteq>) xs ys"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    18
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    19
instance proof
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    20
  fix xs :: "'a list"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    21
  from below_refl show "xs \<sqsubseteq> xs"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    22
    unfolding below_list_def
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    23
    by (rule list_all2_refl)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    24
next
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    25
  fix xs ys zs :: "'a list"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    26
  assume "xs \<sqsubseteq> ys" and "ys \<sqsubseteq> zs"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    27
  with below_trans show "xs \<sqsubseteq> zs"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    28
    unfolding below_list_def
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    29
    by (rule list_all2_trans)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    30
next
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    31
  fix xs ys zs :: "'a list"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    32
  assume "xs \<sqsubseteq> ys" and "ys \<sqsubseteq> xs"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    33
  with below_antisym show "xs = ys"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    34
    unfolding below_list_def
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    35
    by (rule list_all2_antisym)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    36
qed
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    37
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    38
end
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    39
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    40
lemma below_list_simps [simp]:
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    41
  "[] \<sqsubseteq> []"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    42
  "x # xs \<sqsubseteq> y # ys \<longleftrightarrow> x \<sqsubseteq> y \<and> xs \<sqsubseteq> ys"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    43
  "\<not> [] \<sqsubseteq> y # ys"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    44
  "\<not> x # xs \<sqsubseteq> []"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    45
by (simp_all add: below_list_def)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    46
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    47
lemma Nil_below_iff [simp]: "[] \<sqsubseteq> xs \<longleftrightarrow> xs = []"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    48
by (cases xs, simp_all)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    49
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    50
lemma below_Nil_iff [simp]: "xs \<sqsubseteq> [] \<longleftrightarrow> xs = []"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    51
by (cases xs, simp_all)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    52
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    53
text "Thanks to Joachim Breitner"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    54
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    55
lemma list_Cons_below:
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    56
  assumes "a # as \<sqsubseteq> xs"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    57
  obtains b and bs where "a \<sqsubseteq> b" and "as \<sqsubseteq> bs" and "xs = b # bs"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    58
  using assms by (cases xs, auto)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    59
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    60
lemma list_below_Cons:
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    61
  assumes "xs \<sqsubseteq> b # bs"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    62
  obtains a and as where "a \<sqsubseteq> b" and "as \<sqsubseteq> bs" and "xs = a # as"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    63
  using assms by (cases xs, auto)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    64
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    65
lemma hd_mono: "xs \<sqsubseteq> ys \<Longrightarrow> hd xs \<sqsubseteq> hd ys"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    66
by (cases xs, simp, cases ys, simp, simp)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    67
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    68
lemma tl_mono: "xs \<sqsubseteq> ys \<Longrightarrow> tl xs \<sqsubseteq> tl ys"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    69
by (cases xs, simp, cases ys, simp, simp)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    70
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    71
lemma ch2ch_hd [simp]: "chain (\<lambda>i. S i) \<Longrightarrow> chain (\<lambda>i. hd (S i))"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    72
by (rule chainI, rule hd_mono, erule chainE)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    73
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    74
lemma ch2ch_tl [simp]: "chain (\<lambda>i. S i) \<Longrightarrow> chain (\<lambda>i. tl (S i))"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    75
by (rule chainI, rule tl_mono, erule chainE)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    76
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    77
lemma below_same_length: "xs \<sqsubseteq> ys \<Longrightarrow> length xs = length ys"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    78
unfolding below_list_def by (rule list_all2_lengthD)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    79
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    80
lemma list_chain_cases:
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    81
  assumes S: "chain S"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    82
  obtains "\<forall>i. S i = []" |
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    83
    A B where "chain A" and "chain B" and "\<forall>i. S i = A i # B i"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    84
proof (cases "S 0")
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    85
  case Nil
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    86
  have "\<forall>i. S 0 \<sqsubseteq> S i" by (simp add: chain_mono [OF S])
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    87
  with Nil have "\<forall>i. S i = []" by simp
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    88
  thus ?thesis ..
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    89
next
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    90
  case (Cons x xs)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    91
  have "\<forall>i. S 0 \<sqsubseteq> S i" by (simp add: chain_mono [OF S])
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    92
  hence *: "\<forall>i. S i \<noteq> []" by (rule all_forward) (auto simp add: Cons)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    93
  let ?A = "\<lambda>i. hd (S i)"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    94
  let ?B = "\<lambda>i. tl (S i)"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    95
  from S have "chain ?A" and "chain ?B" by simp_all
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    96
  moreover have "\<forall>i. S i = ?A i # ?B i" by (simp add: *)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    97
  ultimately show ?thesis ..
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    98
qed
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    99
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   100
subsection {* Lists are a complete partial order *}
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   101
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   102
lemma is_lub_Cons:
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   103
  assumes A: "range A <<| x"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   104
  assumes B: "range B <<| xs"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   105
  shows "range (\<lambda>i. A i # B i) <<| x # xs"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   106
using assms
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   107
unfolding is_lub_def is_ub_def Ball_def [symmetric]
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   108
by (clarsimp, case_tac u, simp_all)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   109
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   110
lemma list_cpo_lemma:
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   111
  fixes S :: "nat \<Rightarrow> 'a::cpo list"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   112
  assumes "chain S" and "\<forall>i. length (S i) = n"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   113
  shows "\<exists>x. range S <<| x"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   114
using assms
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   115
 apply (induct n arbitrary: S)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   116
  apply (subgoal_tac "S = (\<lambda>i. [])")
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   117
  apply (fast intro: lub_const)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   118
  apply (simp add: expand_fun_eq)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   119
 apply (drule_tac x="\<lambda>i. tl (S i)" in meta_spec, clarsimp)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   120
 apply (rule_tac x="(\<Squnion>i. hd (S i)) # x" in exI)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   121
 apply (subgoal_tac "range (\<lambda>i. hd (S i) # tl (S i)) = range S")
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   122
  apply (erule subst)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   123
  apply (rule is_lub_Cons)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   124
   apply (rule thelubE [OF _ refl])
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   125
   apply (erule ch2ch_hd)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   126
  apply assumption
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   127
 apply (rule_tac f="\<lambda>S. range S" in arg_cong)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   128
 apply (rule ext)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   129
 apply (rule hd_Cons_tl)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   130
 apply (drule_tac x=i in spec, auto)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   131
done
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   132
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   133
instance list :: (cpo) cpo
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   134
proof
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   135
  fix S :: "nat \<Rightarrow> 'a list"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   136
  assume "chain S"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   137
  hence "\<forall>i. S 0 \<sqsubseteq> S i" by (simp add: chain_mono)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   138
  hence "\<forall>i. length (S i) = length (S 0)"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   139
    by (fast intro: below_same_length [symmetric])
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   140
  with `chain S` show "\<exists>x. range S <<| x"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   141
    by (rule list_cpo_lemma)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   142
qed
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   143
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   144
subsection {* Continuity of list operations *}
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   145
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   146
lemma cont2cont_Cons [simp, cont2cont]:
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   147
  assumes f: "cont (\<lambda>x. f x)"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   148
  assumes g: "cont (\<lambda>x. g x)"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   149
  shows "cont (\<lambda>x. f x # g x)"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   150
apply (rule contI)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   151
apply (rule is_lub_Cons)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   152
apply (erule contE [OF f])
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   153
apply (erule contE [OF g])
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   154
done
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   155
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   156
lemma lub_Cons:
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   157
  fixes A :: "nat \<Rightarrow> 'a::cpo"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   158
  assumes A: "chain A" and B: "chain B"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   159
  shows "(\<Squnion>i. A i # B i) = (\<Squnion>i. A i) # (\<Squnion>i. B i)"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   160
by (intro thelubI is_lub_Cons cpo_lubI A B)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   161
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   162
lemma cont2cont_list_case:
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   163
  assumes f: "cont (\<lambda>x. f x)"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   164
  assumes g: "cont (\<lambda>x. g x)"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   165
  assumes h1: "\<And>y ys. cont (\<lambda>x. h x y ys)"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   166
  assumes h2: "\<And>x ys. cont (\<lambda>y. h x y ys)"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   167
  assumes h3: "\<And>x y. cont (\<lambda>ys. h x y ys)"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   168
  shows "cont (\<lambda>x. case f x of [] \<Rightarrow> g x | y # ys \<Rightarrow> h x y ys)"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   169
apply (rule cont_apply [OF f])
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   170
apply (rule contI)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   171
apply (erule list_chain_cases)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   172
apply (simp add: lub_const)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   173
apply (simp add: lub_Cons)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   174
apply (simp add: cont2contlubE [OF h2])
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   175
apply (simp add: cont2contlubE [OF h3])
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   176
apply (simp add: diag_lub ch2ch_cont [OF h2] ch2ch_cont [OF h3])
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   177
apply (rule cpo_lubI, rule chainI, rule below_trans)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   178
apply (erule cont2monofunE [OF h2 chainE])
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   179
apply (erule cont2monofunE [OF h3 chainE])
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   180
apply (case_tac y, simp_all add: g h1)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   181
done
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   182
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   183
lemma cont2cont_list_case' [simp, cont2cont]:
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   184
  assumes f: "cont (\<lambda>x. f x)"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   185
  assumes g: "cont (\<lambda>x. g x)"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   186
  assumes h: "cont (\<lambda>p. h (fst p) (fst (snd p)) (snd (snd p)))"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   187
  shows "cont (\<lambda>x. case f x of [] \<Rightarrow> g x | y # ys \<Rightarrow> h x y ys)"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   188
proof -
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   189
  have "\<And>y ys. cont (\<lambda>x. h x (fst (y, ys)) (snd (y, ys)))"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   190
    by (rule h [THEN cont_fst_snd_D1])
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   191
  hence h1: "\<And>y ys. cont (\<lambda>x. h x y ys)" by simp
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   192
  note h2 = h [THEN cont_fst_snd_D2, THEN cont_fst_snd_D1]
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   193
  note h3 = h [THEN cont_fst_snd_D2, THEN cont_fst_snd_D2]
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   194
  from f g h1 h2 h3 show ?thesis by (rule cont2cont_list_case)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   195
qed
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   196
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   197
text {* The simple version (due to Joachim Breitner) is needed if the
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   198
  element type of the list is not a cpo. *}
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   199
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   200
lemma cont2cont_list_case_simple [simp, cont2cont]:
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   201
  assumes "cont (\<lambda>x. f1 x)"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   202
  assumes "\<And>y ys. cont (\<lambda>x. f2 x y ys)"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   203
  shows "cont (\<lambda>x. case l of [] \<Rightarrow> f1 x | y # ys \<Rightarrow> f2 x y ys)"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   204
using assms by (cases l) auto
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   205
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   206
text {* There are probably lots of other list operations that also
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   207
deserve to have continuity lemmas.  I'll add more as they are
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   208
needed. *}
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   209
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   210
end