src/HOL/HOLCF/Library/List_Cpo.thy
author wenzelm
Sat, 07 Apr 2012 16:41:59 +0200
changeset 47389 e8552cba702d
parent 42151 4da4fc77664b
child 54863 82acc20ded73
permissions -rw-r--r--
explicit checks stable_finished_theory/stable_command allow parallel asynchronous command transactions; tuned;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
42151
4da4fc77664b tuned headers;
wenzelm
parents: 41027
diff changeset
     1
(*  Title:      HOL/HOLCF/Library/List_Cpo.thy
39143
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
39966
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
    53
lemma list_below_induct [consumes 1, case_names Nil Cons]:
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
    54
  assumes "xs \<sqsubseteq> ys"
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
    55
  assumes 1: "P [] []"
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
    56
  assumes 2: "\<And>x y xs ys. \<lbrakk>x \<sqsubseteq> y; xs \<sqsubseteq> ys; P xs ys\<rbrakk> \<Longrightarrow> P (x # xs) (y # ys)"
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
    57
  shows "P xs ys"
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
    58
using `xs \<sqsubseteq> ys`
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
    59
proof (induct xs arbitrary: ys)
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
    60
  case Nil thus ?case by (simp add: 1)
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
    61
next
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
    62
  case (Cons x xs) thus ?case by (cases ys, simp_all add: 2)
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
    63
qed
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
    64
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
    65
lemma list_below_cases:
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
    66
  assumes "xs \<sqsubseteq> ys"
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
    67
  obtains "xs = []" and "ys = []" |
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
    68
    x y xs' ys' where "xs = x # xs'" and "ys = y # ys'"
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
    69
using assms by (cases xs, simp, cases ys, auto)
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
    70
39143
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    71
text "Thanks to Joachim Breitner"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    72
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    73
lemma list_Cons_below:
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    74
  assumes "a # as \<sqsubseteq> xs"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    75
  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
    76
  using assms by (cases xs, auto)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    77
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    78
lemma list_below_Cons:
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    79
  assumes "xs \<sqsubseteq> b # bs"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    80
  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
    81
  using assms by (cases xs, auto)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    82
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    83
lemma hd_mono: "xs \<sqsubseteq> ys \<Longrightarrow> hd xs \<sqsubseteq> hd ys"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    84
by (cases xs, simp, cases ys, simp, simp)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    85
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    86
lemma tl_mono: "xs \<sqsubseteq> ys \<Longrightarrow> tl xs \<sqsubseteq> tl ys"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    87
by (cases xs, simp, cases ys, simp, simp)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    88
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    89
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
    90
by (rule chainI, rule hd_mono, erule chainE)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    91
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    92
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
    93
by (rule chainI, rule tl_mono, erule chainE)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    94
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    95
lemma below_same_length: "xs \<sqsubseteq> ys \<Longrightarrow> length xs = length ys"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    96
unfolding below_list_def by (rule list_all2_lengthD)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    97
39966
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
    98
lemma list_chain_induct [consumes 1, case_names Nil Cons]:
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
    99
  assumes "chain S"
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   100
  assumes 1: "P (\<lambda>i. [])"
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   101
  assumes 2: "\<And>A B. chain A \<Longrightarrow> chain B \<Longrightarrow> P B \<Longrightarrow> P (\<lambda>i. A i # B i)"
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   102
  shows "P S"
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   103
using `chain S`
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   104
proof (induct "S 0" arbitrary: S)
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   105
  case Nil
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   106
  have "\<forall>i. S 0 \<sqsubseteq> S i" by (simp add: chain_mono [OF `chain S`])
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   107
  with Nil have "\<forall>i. S i = []" by simp
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   108
  thus ?case by (simp add: 1)
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   109
next
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   110
  case (Cons x xs)
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   111
  have "\<forall>i. S 0 \<sqsubseteq> S i" by (simp add: chain_mono [OF `chain S`])
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   112
  hence *: "\<forall>i. S i \<noteq> []" by (rule all_forward, insert Cons) auto
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   113
  have "chain (\<lambda>i. hd (S i))" and "chain (\<lambda>i. tl (S i))"
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   114
    using `chain S` by simp_all
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   115
  moreover have "P (\<lambda>i. tl (S i))"
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   116
    using `chain S` and `x # xs = S 0` [symmetric]
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   117
    by (simp add: Cons(1))
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   118
  ultimately have "P (\<lambda>i. hd (S i) # tl (S i))"
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   119
    by (rule 2)
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   120
  thus "P S" by (simp add: *)
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   121
qed
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   122
39143
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   123
lemma list_chain_cases:
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   124
  assumes S: "chain S"
39966
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   125
  obtains "S = (\<lambda>i. [])" |
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   126
    A B where "chain A" and "chain B" and "S = (\<lambda>i. A i # B i)"
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   127
using S by (induct rule: list_chain_induct) simp_all
39143
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   128
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   129
subsection {* Lists are a complete partial order *}
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   130
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   131
lemma is_lub_Cons:
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   132
  assumes A: "range A <<| x"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   133
  assumes B: "range B <<| xs"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   134
  shows "range (\<lambda>i. A i # B i) <<| x # xs"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   135
using assms
39968
d841744718fe define is_ub predicate using bounded quantifier
huffman
parents: 39966
diff changeset
   136
unfolding is_lub_def is_ub_def
39143
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   137
by (clarsimp, case_tac u, simp_all)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   138
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   139
instance list :: (cpo) cpo
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   140
proof
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   141
  fix S :: "nat \<Rightarrow> 'a list"
39966
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   142
  assume "chain S" thus "\<exists>x. range S <<| x"
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   143
  proof (induct rule: list_chain_induct)
40771
1c6f7d4b110e renamed several HOLCF theorems (listed in NEWS)
huffman
parents: 39968
diff changeset
   144
    case Nil thus ?case by (auto intro: is_lub_const)
39966
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   145
  next
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   146
    case (Cons A B) thus ?case by (auto intro: is_lub_Cons cpo_lubI)
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   147
  qed
39143
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   148
qed
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   149
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   150
subsection {* Continuity of list operations *}
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   151
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   152
lemma cont2cont_Cons [simp, cont2cont]:
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   153
  assumes f: "cont (\<lambda>x. f x)"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   154
  assumes g: "cont (\<lambda>x. g x)"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   155
  shows "cont (\<lambda>x. f x # g x)"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   156
apply (rule contI)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   157
apply (rule is_lub_Cons)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   158
apply (erule contE [OF f])
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   159
apply (erule contE [OF g])
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   160
done
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 lub_Cons:
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   163
  fixes A :: "nat \<Rightarrow> 'a::cpo"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   164
  assumes A: "chain A" and B: "chain B"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   165
  shows "(\<Squnion>i. A i # B i) = (\<Squnion>i. A i) # (\<Squnion>i. B i)"
40771
1c6f7d4b110e renamed several HOLCF theorems (listed in NEWS)
huffman
parents: 39968
diff changeset
   166
by (intro lub_eqI is_lub_Cons cpo_lubI A B)
39143
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   167
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   168
lemma cont2cont_list_case:
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   169
  assumes f: "cont (\<lambda>x. f x)"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   170
  assumes g: "cont (\<lambda>x. g x)"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   171
  assumes h1: "\<And>y ys. cont (\<lambda>x. h x y ys)"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   172
  assumes h2: "\<And>x ys. cont (\<lambda>y. h x y ys)"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   173
  assumes h3: "\<And>x y. cont (\<lambda>ys. h x y ys)"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   174
  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
   175
apply (rule cont_apply [OF f])
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   176
apply (rule contI)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   177
apply (erule list_chain_cases)
40771
1c6f7d4b110e renamed several HOLCF theorems (listed in NEWS)
huffman
parents: 39968
diff changeset
   178
apply (simp add: is_lub_const)
39143
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   179
apply (simp add: lub_Cons)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   180
apply (simp add: cont2contlubE [OF h2])
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   181
apply (simp add: cont2contlubE [OF h3])
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   182
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
   183
apply (rule cpo_lubI, rule chainI, rule below_trans)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   184
apply (erule cont2monofunE [OF h2 chainE])
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   185
apply (erule cont2monofunE [OF h3 chainE])
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   186
apply (case_tac y, simp_all add: g h1)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   187
done
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   188
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   189
lemma cont2cont_list_case' [simp, cont2cont]:
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   190
  assumes f: "cont (\<lambda>x. f x)"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   191
  assumes g: "cont (\<lambda>x. g x)"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   192
  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
   193
  shows "cont (\<lambda>x. case f x of [] \<Rightarrow> g x | y # ys \<Rightarrow> h x y ys)"
39808
1410c84013b9 rename cont2cont_split to cont2cont_prod_case; add lemmas prod_contI, prod_cont_iff; simplify some proofs
huffman
parents: 39302
diff changeset
   194
using assms by (simp add: cont2cont_list_case prod_cont_iff)
39143
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   195
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   196
text {* The simple version (due to Joachim Breitner) is needed if the
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   197
  element type of the list is not a cpo. *}
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   198
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   199
lemma cont2cont_list_case_simple [simp, cont2cont]:
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   200
  assumes "cont (\<lambda>x. f1 x)"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   201
  assumes "\<And>y ys. cont (\<lambda>x. f2 x y ys)"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   202
  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
   203
using assms by (cases l) auto
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   204
39966
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   205
text {* Lemma for proving continuity of recursive list functions: *}
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   206
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   207
lemma list_contI:
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   208
  fixes f :: "'a::cpo list \<Rightarrow> 'b::cpo"
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   209
  assumes f: "\<And>x xs. f (x # xs) = g x xs (f xs)"
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   210
  assumes g1: "\<And>xs y. cont (\<lambda>x. g x xs y)"
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   211
  assumes g2: "\<And>x y. cont (\<lambda>xs. g x xs y)"
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   212
  assumes g3: "\<And>x xs. cont (\<lambda>y. g x xs y)"
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   213
  shows "cont f"
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   214
proof (rule contI2)
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   215
  obtain h where h: "\<And>x xs y. g x xs y = h\<cdot>x\<cdot>xs\<cdot>y"
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   216
  proof
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   217
    fix x xs y show "g x xs y = (\<Lambda> x xs y. g x xs y)\<cdot>x\<cdot>xs\<cdot>y"
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   218
    by (simp add: cont2cont_LAM g1 g2 g3)
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   219
  qed
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   220
  show mono: "monofun f"
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   221
    apply (rule monofunI)
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   222
    apply (erule list_below_induct)
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   223
    apply simp
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   224
    apply (simp add: f h monofun_cfun)
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   225
    done
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   226
  fix Y :: "nat \<Rightarrow> 'a list"
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   227
  assume "chain Y" thus "f (\<Squnion>i. Y i) \<sqsubseteq> (\<Squnion>i. f (Y i))"
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   228
    apply (induct rule: list_chain_induct)
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   229
    apply simp
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   230
    apply (simp add: lub_Cons f h)
41027
c599955d9806 add lemmas lub_APP, lub_LAM
huffman
parents: 40831
diff changeset
   231
    apply (simp add: lub_APP ch2ch_monofun [OF mono])
39966
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   232
    apply (simp add: monofun_cfun)
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   233
    done
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   234
qed
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   235
40831
10aeee1d5b71 add continuity lemma for List.map
huffman
parents: 40808
diff changeset
   236
text {* Continuity rule for map *}
10aeee1d5b71 add continuity lemma for List.map
huffman
parents: 40808
diff changeset
   237
10aeee1d5b71 add continuity lemma for List.map
huffman
parents: 40808
diff changeset
   238
lemma cont2cont_map [simp, cont2cont]:
10aeee1d5b71 add continuity lemma for List.map
huffman
parents: 40808
diff changeset
   239
  assumes f: "cont (\<lambda>(x, y). f x y)"
10aeee1d5b71 add continuity lemma for List.map
huffman
parents: 40808
diff changeset
   240
  assumes g: "cont (\<lambda>x. g x)"
10aeee1d5b71 add continuity lemma for List.map
huffman
parents: 40808
diff changeset
   241
  shows "cont (\<lambda>x. map (\<lambda>y. f x y) (g x))"
10aeee1d5b71 add continuity lemma for List.map
huffman
parents: 40808
diff changeset
   242
using f
10aeee1d5b71 add continuity lemma for List.map
huffman
parents: 40808
diff changeset
   243
apply (simp add: prod_cont_iff)
10aeee1d5b71 add continuity lemma for List.map
huffman
parents: 40808
diff changeset
   244
apply (rule cont_apply [OF g])
10aeee1d5b71 add continuity lemma for List.map
huffman
parents: 40808
diff changeset
   245
apply (rule list_contI, rule map.simps, simp, simp, simp)
10aeee1d5b71 add continuity lemma for List.map
huffman
parents: 40808
diff changeset
   246
apply (induct_tac y, simp, simp)
10aeee1d5b71 add continuity lemma for List.map
huffman
parents: 40808
diff changeset
   247
done
10aeee1d5b71 add continuity lemma for List.map
huffman
parents: 40808
diff changeset
   248
39143
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   249
text {* There are probably lots of other list operations that also
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   250
deserve to have continuity lemmas.  I'll add more as they are
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   251
needed. *}
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   252
40808
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   253
subsection {* Lists are a discrete cpo *}
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   254
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   255
instance list :: (discrete_cpo) discrete_cpo
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   256
proof
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   257
  fix xs ys :: "'a list"
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   258
  show "xs \<sqsubseteq> ys \<longleftrightarrow> xs = ys"
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   259
    by (induct xs arbitrary: ys, case_tac [!] ys, simp_all)
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   260
qed
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   261
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   262
subsection {* Compactness and chain-finiteness *}
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   263
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   264
lemma compact_Nil [simp]: "compact []"
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   265
apply (rule compactI2)
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   266
apply (erule list_chain_cases)
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   267
apply simp
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   268
apply (simp add: lub_Cons)
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   269
done
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   270
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   271
lemma compact_Cons: "\<lbrakk>compact x; compact xs\<rbrakk> \<Longrightarrow> compact (x # xs)"
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   272
apply (rule compactI2)
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   273
apply (erule list_chain_cases)
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   274
apply (auto simp add: lub_Cons dest!: compactD2)
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   275
apply (rename_tac i j, rule_tac x="max i j" in exI)
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   276
apply (drule (1) below_trans [OF _ chain_mono [OF _ le_maxI1]])
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   277
apply (drule (1) below_trans [OF _ chain_mono [OF _ le_maxI2]])
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   278
apply (erule (1) conjI)
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   279
done
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   280
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   281
lemma compact_Cons_iff [simp]:
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   282
  "compact (x # xs) \<longleftrightarrow> compact x \<and> compact xs"
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   283
apply (safe intro!: compact_Cons)
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   284
apply (simp only: compact_def)
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   285
apply (subgoal_tac "cont (\<lambda>x. x # xs)")
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   286
apply (drule (1) adm_subst, simp, simp)
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   287
apply (simp only: compact_def)
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   288
apply (subgoal_tac "cont (\<lambda>xs. x # xs)")
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   289
apply (drule (1) adm_subst, simp, simp)
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   290
done
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   291
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   292
instance list :: (chfin) chfin
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   293
proof
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   294
  fix Y :: "nat \<Rightarrow> 'a list" assume "chain Y"
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   295
  moreover have "\<And>(xs::'a list). compact xs"
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   296
    by (induct_tac xs, simp_all)
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   297
  ultimately show "\<exists>i. max_in_chain i Y"
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   298
    by (rule compact_imp_max_in_chain)
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   299
qed
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   300
39212
3989b2b44dba set up Nil and Cons to work as fixrec patterns
huffman
parents: 39199
diff changeset
   301
subsection {* Using lists with fixrec *}
3989b2b44dba set up Nil and Cons to work as fixrec patterns
huffman
parents: 39199
diff changeset
   302
3989b2b44dba set up Nil and Cons to work as fixrec patterns
huffman
parents: 39199
diff changeset
   303
definition
3989b2b44dba set up Nil and Cons to work as fixrec patterns
huffman
parents: 39199
diff changeset
   304
  match_Nil :: "'a::cpo list \<rightarrow> 'b match \<rightarrow> 'b match"
3989b2b44dba set up Nil and Cons to work as fixrec patterns
huffman
parents: 39199
diff changeset
   305
where
3989b2b44dba set up Nil and Cons to work as fixrec patterns
huffman
parents: 39199
diff changeset
   306
  "match_Nil = (\<Lambda> xs k. case xs of [] \<Rightarrow> k | y # ys \<Rightarrow> Fixrec.fail)"
3989b2b44dba set up Nil and Cons to work as fixrec patterns
huffman
parents: 39199
diff changeset
   307
3989b2b44dba set up Nil and Cons to work as fixrec patterns
huffman
parents: 39199
diff changeset
   308
definition
3989b2b44dba set up Nil and Cons to work as fixrec patterns
huffman
parents: 39199
diff changeset
   309
  match_Cons :: "'a::cpo list \<rightarrow> ('a \<rightarrow> 'a list \<rightarrow> 'b match) \<rightarrow> 'b match"
3989b2b44dba set up Nil and Cons to work as fixrec patterns
huffman
parents: 39199
diff changeset
   310
where
3989b2b44dba set up Nil and Cons to work as fixrec patterns
huffman
parents: 39199
diff changeset
   311
  "match_Cons = (\<Lambda> xs k. case xs of [] \<Rightarrow> Fixrec.fail | y # ys \<Rightarrow> k\<cdot>y\<cdot>ys)"
3989b2b44dba set up Nil and Cons to work as fixrec patterns
huffman
parents: 39199
diff changeset
   312
3989b2b44dba set up Nil and Cons to work as fixrec patterns
huffman
parents: 39199
diff changeset
   313
lemma match_Nil_simps [simp]:
3989b2b44dba set up Nil and Cons to work as fixrec patterns
huffman
parents: 39199
diff changeset
   314
  "match_Nil\<cdot>[]\<cdot>k = k"
3989b2b44dba set up Nil and Cons to work as fixrec patterns
huffman
parents: 39199
diff changeset
   315
  "match_Nil\<cdot>(x # xs)\<cdot>k = Fixrec.fail"
3989b2b44dba set up Nil and Cons to work as fixrec patterns
huffman
parents: 39199
diff changeset
   316
unfolding match_Nil_def by simp_all
3989b2b44dba set up Nil and Cons to work as fixrec patterns
huffman
parents: 39199
diff changeset
   317
3989b2b44dba set up Nil and Cons to work as fixrec patterns
huffman
parents: 39199
diff changeset
   318
lemma match_Cons_simps [simp]:
3989b2b44dba set up Nil and Cons to work as fixrec patterns
huffman
parents: 39199
diff changeset
   319
  "match_Cons\<cdot>[]\<cdot>k = Fixrec.fail"
3989b2b44dba set up Nil and Cons to work as fixrec patterns
huffman
parents: 39199
diff changeset
   320
  "match_Cons\<cdot>(x # xs)\<cdot>k = k\<cdot>x\<cdot>xs"
3989b2b44dba set up Nil and Cons to work as fixrec patterns
huffman
parents: 39199
diff changeset
   321
unfolding match_Cons_def by simp_all
3989b2b44dba set up Nil and Cons to work as fixrec patterns
huffman
parents: 39199
diff changeset
   322
3989b2b44dba set up Nil and Cons to work as fixrec patterns
huffman
parents: 39199
diff changeset
   323
setup {*
3989b2b44dba set up Nil and Cons to work as fixrec patterns
huffman
parents: 39199
diff changeset
   324
  Fixrec.add_matchers
3989b2b44dba set up Nil and Cons to work as fixrec patterns
huffman
parents: 39199
diff changeset
   325
    [ (@{const_name Nil}, @{const_name match_Nil}),
3989b2b44dba set up Nil and Cons to work as fixrec patterns
huffman
parents: 39199
diff changeset
   326
      (@{const_name Cons}, @{const_name match_Cons}) ]
3989b2b44dba set up Nil and Cons to work as fixrec patterns
huffman
parents: 39199
diff changeset
   327
*}
3989b2b44dba set up Nil and Cons to work as fixrec patterns
huffman
parents: 39199
diff changeset
   328
39143
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   329
end