src/HOL/HOLCF/Library/List_Cpo.thy
author haftmann
Fri, 27 Jun 2025 08:09:26 +0200
changeset 82775 61c39a9e5415
parent 81733 94f018b2a4fb
permissions -rw-r--r--
typo
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
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
     5
section \<open>Lists as a complete partial order\<close>
39143
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
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
    11
subsection \<open>Lists are a partial order\<close>
39143
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
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 62175
diff changeset
    17
  "xs \<sqsubseteq> ys \<longleftrightarrow> list_all2 (\<sqsubseteq>) xs ys"
39143
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    18
81733
94f018b2a4fb tuned proofs;
wenzelm
parents: 69597
diff changeset
    19
instance
94f018b2a4fb tuned proofs;
wenzelm
parents: 69597
diff changeset
    20
proof
94f018b2a4fb tuned proofs;
wenzelm
parents: 69597
diff changeset
    21
  fix xs ys zs :: "'a list"
94f018b2a4fb tuned proofs;
wenzelm
parents: 69597
diff changeset
    22
  show "xs \<sqsubseteq> xs"
94f018b2a4fb tuned proofs;
wenzelm
parents: 69597
diff changeset
    23
    using below_refl 
39143
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    24
    unfolding below_list_def
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    25
    by (rule list_all2_refl)
81733
94f018b2a4fb tuned proofs;
wenzelm
parents: 69597
diff changeset
    26
  show "xs \<sqsubseteq> ys \<Longrightarrow> ys \<sqsubseteq> zs \<Longrightarrow> xs \<sqsubseteq> zs"
94f018b2a4fb tuned proofs;
wenzelm
parents: 69597
diff changeset
    27
    using below_trans 
39143
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)
81733
94f018b2a4fb tuned proofs;
wenzelm
parents: 69597
diff changeset
    30
  show "xs \<sqsubseteq> ys \<Longrightarrow> ys \<sqsubseteq> xs \<Longrightarrow> xs = ys"
94f018b2a4fb tuned proofs;
wenzelm
parents: 69597
diff changeset
    31
    using below_antisym
39143
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    32
    unfolding below_list_def
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    33
    by (rule list_all2_antisym)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    34
qed
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    35
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    36
end
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
lemma below_list_simps [simp]:
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    39
  "[] \<sqsubseteq> []"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    40
  "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
    41
  "\<not> [] \<sqsubseteq> y # ys"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    42
  "\<not> x # xs \<sqsubseteq> []"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    43
by (simp_all add: below_list_def)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    44
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    45
lemma Nil_below_iff [simp]: "[] \<sqsubseteq> xs \<longleftrightarrow> xs = []"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    46
by (cases xs, simp_all)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    47
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    48
lemma below_Nil_iff [simp]: "xs \<sqsubseteq> [] \<longleftrightarrow> xs = []"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    49
by (cases xs, simp_all)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    50
39966
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
    51
lemma list_below_induct [consumes 1, case_names Nil Cons]:
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
    52
  assumes "xs \<sqsubseteq> ys"
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
    53
  assumes 1: "P [] []"
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
    54
  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
    55
  shows "P xs ys"
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
    56
using \<open>xs \<sqsubseteq> ys\<close>
39966
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
    57
proof (induct xs arbitrary: ys)
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
    58
  case Nil thus ?case by (simp add: 1)
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
    59
next
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
    60
  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
    61
qed
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
    62
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
    63
lemma list_below_cases:
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
    64
  assumes "xs \<sqsubseteq> ys"
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
    65
  obtains "xs = []" and "ys = []" |
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
    66
    x y xs' ys' where "xs = x # xs'" and "ys = y # ys'"
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
    67
using assms by (cases xs, simp, cases ys, auto)
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
    68
39143
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    69
text "Thanks to Joachim Breitner"
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 list_Cons_below:
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    72
  assumes "a # as \<sqsubseteq> xs"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    73
  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
    74
  using assms by (cases xs, auto)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    75
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    76
lemma list_below_Cons:
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    77
  assumes "xs \<sqsubseteq> b # bs"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    78
  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
    79
  using assms by (cases xs, auto)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    80
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    81
lemma hd_mono: "xs \<sqsubseteq> ys \<Longrightarrow> hd xs \<sqsubseteq> hd ys"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    82
by (cases xs, simp, cases ys, simp, simp)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    83
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    84
lemma tl_mono: "xs \<sqsubseteq> ys \<Longrightarrow> tl xs \<sqsubseteq> tl ys"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    85
by (cases xs, simp, cases ys, simp, simp)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    86
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    87
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
    88
by (rule chainI, rule hd_mono, erule chainE)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    89
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    90
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
    91
by (rule chainI, rule tl_mono, erule chainE)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    92
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    93
lemma below_same_length: "xs \<sqsubseteq> ys \<Longrightarrow> length xs = length ys"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    94
unfolding below_list_def by (rule list_all2_lengthD)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
    95
39966
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
    96
lemma list_chain_induct [consumes 1, case_names Nil Cons]:
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
    97
  assumes "chain S"
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
    98
  assumes 1: "P (\<lambda>i. [])"
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
    99
  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
   100
  shows "P S"
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
   101
using \<open>chain S\<close>
39966
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   102
proof (induct "S 0" arbitrary: S)
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   103
  case Nil
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
   104
  have "\<forall>i. S 0 \<sqsubseteq> S i" by (simp add: chain_mono [OF \<open>chain S\<close>])
39966
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   105
  with Nil have "\<forall>i. S i = []" by simp
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   106
  thus ?case by (simp add: 1)
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   107
next
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   108
  case (Cons x xs)
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
   109
  have "\<forall>i. S 0 \<sqsubseteq> S i" by (simp add: chain_mono [OF \<open>chain S\<close>])
39966
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   110
  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
   111
  have "chain (\<lambda>i. hd (S i))" and "chain (\<lambda>i. tl (S i))"
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
   112
    using \<open>chain S\<close> by simp_all
39966
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   113
  moreover have "P (\<lambda>i. tl (S i))"
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
   114
    using \<open>chain S\<close> and \<open>x # xs = S 0\<close> [symmetric]
39966
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   115
    by (simp add: Cons(1))
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   116
  ultimately have "P (\<lambda>i. hd (S i) # tl (S i))"
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   117
    by (rule 2)
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   118
  thus "P S" by (simp add: *)
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   119
qed
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   120
39143
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   121
lemma list_chain_cases:
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   122
  assumes S: "chain S"
39966
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   123
  obtains "S = (\<lambda>i. [])" |
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   124
    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
   125
using S by (induct rule: list_chain_induct) simp_all
39143
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   126
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
   127
subsection \<open>Lists are a complete partial order\<close>
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
lemma is_lub_Cons:
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   130
  assumes A: "range A <<| x"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   131
  assumes B: "range B <<| xs"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   132
  shows "range (\<lambda>i. A i # B i) <<| x # xs"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   133
using assms
39968
d841744718fe define is_ub predicate using bounded quantifier
huffman
parents: 39966
diff changeset
   134
unfolding is_lub_def is_ub_def
39143
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   135
by (clarsimp, case_tac u, simp_all)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   136
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   137
instance list :: (cpo) cpo
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   138
proof
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   139
  fix S :: "nat \<Rightarrow> 'a list"
68780
54fdc8bc73a3 new simp rule
haftmann
parents: 67399
diff changeset
   140
  assume "chain S"
54fdc8bc73a3 new simp rule
haftmann
parents: 67399
diff changeset
   141
  then show "\<exists>x. range S <<| x"
39966
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   142
  proof (induct rule: list_chain_induct)
68780
54fdc8bc73a3 new simp rule
haftmann
parents: 67399
diff changeset
   143
    case Nil
54fdc8bc73a3 new simp rule
haftmann
parents: 67399
diff changeset
   144
    have "{[]} <<| []"
54fdc8bc73a3 new simp rule
haftmann
parents: 67399
diff changeset
   145
      by simp
54fdc8bc73a3 new simp rule
haftmann
parents: 67399
diff changeset
   146
    then obtain x :: "'a list" where "{[]} <<| x" ..
54fdc8bc73a3 new simp rule
haftmann
parents: 67399
diff changeset
   147
    then show ?case
54fdc8bc73a3 new simp rule
haftmann
parents: 67399
diff changeset
   148
      by auto
39966
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   149
  next
68780
54fdc8bc73a3 new simp rule
haftmann
parents: 67399
diff changeset
   150
    case (Cons A B) then show ?case
54fdc8bc73a3 new simp rule
haftmann
parents: 67399
diff changeset
   151
      by (auto intro: is_lub_Cons cpo_lubI)
39966
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   152
  qed
39143
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   153
qed
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   154
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
   155
subsection \<open>Continuity of list operations\<close>
39143
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   156
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   157
lemma cont2cont_Cons [simp, cont2cont]:
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   158
  assumes f: "cont (\<lambda>x. f x)"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   159
  assumes g: "cont (\<lambda>x. g x)"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   160
  shows "cont (\<lambda>x. f x # g x)"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   161
apply (rule contI)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   162
apply (rule is_lub_Cons)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   163
apply (erule contE [OF f])
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   164
apply (erule contE [OF g])
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   165
done
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   166
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   167
lemma lub_Cons:
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   168
  fixes A :: "nat \<Rightarrow> 'a::cpo"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   169
  assumes A: "chain A" and B: "chain B"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   170
  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
   171
by (intro lub_eqI is_lub_Cons cpo_lubI A B)
39143
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   172
55413
a8e96847523c adapted theories to '{case,rec}_{list,option}' names
blanchet
parents: 54863
diff changeset
   173
lemma cont2cont_case_list:
39143
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   174
  assumes f: "cont (\<lambda>x. f x)"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   175
  assumes g: "cont (\<lambda>x. g x)"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   176
  assumes h1: "\<And>y ys. cont (\<lambda>x. h x y ys)"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   177
  assumes h2: "\<And>x ys. cont (\<lambda>y. h x y ys)"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   178
  assumes h3: "\<And>x y. cont (\<lambda>ys. h x y ys)"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   179
  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
   180
apply (rule cont_apply [OF f])
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   181
apply (rule contI)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   182
apply (erule list_chain_cases)
40771
1c6f7d4b110e renamed several HOLCF theorems (listed in NEWS)
huffman
parents: 39968
diff changeset
   183
apply (simp add: is_lub_const)
39143
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   184
apply (simp add: lub_Cons)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   185
apply (simp add: cont2contlubE [OF h2])
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   186
apply (simp add: cont2contlubE [OF h3])
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   187
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
   188
apply (rule cpo_lubI, rule chainI, rule below_trans)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   189
apply (erule cont2monofunE [OF h2 chainE])
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   190
apply (erule cont2monofunE [OF h3 chainE])
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   191
apply (case_tac y, simp_all add: g h1)
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   192
done
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   193
55413
a8e96847523c adapted theories to '{case,rec}_{list,option}' names
blanchet
parents: 54863
diff changeset
   194
lemma cont2cont_case_list' [simp, cont2cont]:
39143
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   195
  assumes f: "cont (\<lambda>x. f x)"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   196
  assumes g: "cont (\<lambda>x. g x)"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   197
  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
   198
  shows "cont (\<lambda>x. case f x of [] \<Rightarrow> g x | y # ys \<Rightarrow> h x y ys)"
55413
a8e96847523c adapted theories to '{case,rec}_{list,option}' names
blanchet
parents: 54863
diff changeset
   199
using assms by (simp add: cont2cont_case_list prod_cont_iff)
39143
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   200
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
   201
text \<open>The simple version (due to Joachim Breitner) is needed if the
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
   202
  element type of the list is not a cpo.\<close>
39143
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   203
55413
a8e96847523c adapted theories to '{case,rec}_{list,option}' names
blanchet
parents: 54863
diff changeset
   204
lemma cont2cont_case_list_simple [simp, cont2cont]:
39143
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   205
  assumes "cont (\<lambda>x. f1 x)"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   206
  assumes "\<And>y ys. cont (\<lambda>x. f2 x y ys)"
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   207
  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
   208
using assms by (cases l) auto
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   209
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
   210
text \<open>Lemma for proving continuity of recursive list functions:\<close>
39966
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   211
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   212
lemma list_contI:
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   213
  fixes f :: "'a::cpo list \<Rightarrow> 'b::cpo"
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   214
  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
   215
  assumes g1: "\<And>xs y. cont (\<lambda>x. g x xs y)"
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   216
  assumes g2: "\<And>x y. cont (\<lambda>xs. g x xs y)"
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   217
  assumes g3: "\<And>x xs. cont (\<lambda>y. g x xs y)"
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   218
  shows "cont f"
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   219
proof (rule contI2)
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   220
  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
   221
  proof
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   222
    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
   223
    by (simp add: cont2cont_LAM g1 g2 g3)
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   224
  qed
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   225
  show mono: "monofun f"
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   226
    apply (rule monofunI)
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   227
    apply (erule list_below_induct)
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   228
    apply simp
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   229
    apply (simp add: f h monofun_cfun)
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   230
    done
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   231
  fix Y :: "nat \<Rightarrow> 'a list"
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   232
  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
   233
    apply (induct rule: list_chain_induct)
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   234
    apply simp
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   235
    apply (simp add: lub_Cons f h)
41027
c599955d9806 add lemmas lub_APP, lub_LAM
huffman
parents: 40831
diff changeset
   236
    apply (simp add: lub_APP ch2ch_monofun [OF mono])
39966
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   237
    apply (simp add: monofun_cfun)
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   238
    done
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   239
qed
20c74cf9c112 added lemmas to List_Cpo.thy
huffman
parents: 39808
diff changeset
   240
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
   241
text \<open>Continuity rule for map\<close>
40831
10aeee1d5b71 add continuity lemma for List.map
huffman
parents: 40808
diff changeset
   242
10aeee1d5b71 add continuity lemma for List.map
huffman
parents: 40808
diff changeset
   243
lemma cont2cont_map [simp, cont2cont]:
10aeee1d5b71 add continuity lemma for List.map
huffman
parents: 40808
diff changeset
   244
  assumes f: "cont (\<lambda>(x, y). f x y)"
10aeee1d5b71 add continuity lemma for List.map
huffman
parents: 40808
diff changeset
   245
  assumes g: "cont (\<lambda>x. g x)"
10aeee1d5b71 add continuity lemma for List.map
huffman
parents: 40808
diff changeset
   246
  shows "cont (\<lambda>x. map (\<lambda>y. f x y) (g x))"
10aeee1d5b71 add continuity lemma for List.map
huffman
parents: 40808
diff changeset
   247
using f
10aeee1d5b71 add continuity lemma for List.map
huffman
parents: 40808
diff changeset
   248
apply (simp add: prod_cont_iff)
10aeee1d5b71 add continuity lemma for List.map
huffman
parents: 40808
diff changeset
   249
apply (rule cont_apply [OF g])
55465
0d31c0546286 merged 'List.map' and 'List.list.map'
blanchet
parents: 55413
diff changeset
   250
apply (rule list_contI, rule list.map, simp, simp, simp)
40831
10aeee1d5b71 add continuity lemma for List.map
huffman
parents: 40808
diff changeset
   251
apply (induct_tac y, simp, simp)
10aeee1d5b71 add continuity lemma for List.map
huffman
parents: 40808
diff changeset
   252
done
10aeee1d5b71 add continuity lemma for List.map
huffman
parents: 40808
diff changeset
   253
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
   254
text \<open>There are probably lots of other list operations that also
39143
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   255
deserve to have continuity lemmas.  I'll add more as they are
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
   256
needed.\<close>
39143
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   257
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
   258
subsection \<open>Lists are a discrete cpo\<close>
40808
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   259
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   260
instance list :: (discrete_cpo) discrete_cpo
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   261
proof
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   262
  fix xs ys :: "'a list"
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   263
  show "xs \<sqsubseteq> ys \<longleftrightarrow> xs = ys"
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   264
    by (induct xs arbitrary: ys, case_tac [!] ys, simp_all)
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   265
qed
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   266
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
   267
subsection \<open>Compactness and chain-finiteness\<close>
40808
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   268
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   269
lemma compact_Nil [simp]: "compact []"
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   270
apply (rule compactI2)
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   271
apply (erule list_chain_cases)
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   272
apply simp
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   273
apply (simp add: lub_Cons)
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   274
done
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   275
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   276
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
   277
apply (rule compactI2)
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   278
apply (erule list_chain_cases)
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   279
apply (auto simp add: lub_Cons dest!: compactD2)
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   280
apply (rename_tac i j, rule_tac x="max i j" in exI)
54863
82acc20ded73 prefer more canonical names for lemmas on min/max
haftmann
parents: 42151
diff changeset
   281
apply (drule (1) below_trans [OF _ chain_mono [OF _ max.cobounded1]])
82acc20ded73 prefer more canonical names for lemmas on min/max
haftmann
parents: 42151
diff changeset
   282
apply (drule (1) below_trans [OF _ chain_mono [OF _ max.cobounded2]])
40808
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   283
apply (erule (1) conjI)
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   284
done
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   285
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   286
lemma compact_Cons_iff [simp]:
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   287
  "compact (x # xs) \<longleftrightarrow> compact x \<and> compact xs"
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   288
apply (safe intro!: compact_Cons)
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   289
apply (simp only: compact_def)
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   290
apply (subgoal_tac "cont (\<lambda>x. x # xs)")
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   291
apply (drule (1) adm_subst, simp, simp)
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   292
apply (simp only: compact_def)
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   293
apply (subgoal_tac "cont (\<lambda>xs. x # xs)")
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   294
apply (drule (1) adm_subst, simp, simp)
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   295
done
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   296
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   297
instance list :: (chfin) chfin
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   298
proof
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   299
  fix Y :: "nat \<Rightarrow> 'a list" assume "chain Y"
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   300
  moreover have "\<And>(xs::'a list). compact xs"
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   301
    by (induct_tac xs, simp_all)
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   302
  ultimately show "\<exists>i. max_in_chain i Y"
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   303
    by (rule compact_imp_max_in_chain)
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   304
qed
63276d22ea60 instance list :: (discrete_cpo) discrete_cpo;
huffman
parents: 40774
diff changeset
   305
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
   306
subsection \<open>Using lists with fixrec\<close>
39212
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_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
   310
where
3989b2b44dba set up Nil and Cons to work as fixrec patterns
huffman
parents: 39199
diff changeset
   311
  "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
   312
3989b2b44dba set up Nil and Cons to work as fixrec patterns
huffman
parents: 39199
diff changeset
   313
definition
3989b2b44dba set up Nil and Cons to work as fixrec patterns
huffman
parents: 39199
diff changeset
   314
  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
   315
where
3989b2b44dba set up Nil and Cons to work as fixrec patterns
huffman
parents: 39199
diff changeset
   316
  "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
   317
3989b2b44dba set up Nil and Cons to work as fixrec patterns
huffman
parents: 39199
diff changeset
   318
lemma match_Nil_simps [simp]:
3989b2b44dba set up Nil and Cons to work as fixrec patterns
huffman
parents: 39199
diff changeset
   319
  "match_Nil\<cdot>[]\<cdot>k = k"
3989b2b44dba set up Nil and Cons to work as fixrec patterns
huffman
parents: 39199
diff changeset
   320
  "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
   321
unfolding match_Nil_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
lemma match_Cons_simps [simp]:
3989b2b44dba set up Nil and Cons to work as fixrec patterns
huffman
parents: 39199
diff changeset
   324
  "match_Cons\<cdot>[]\<cdot>k = Fixrec.fail"
3989b2b44dba set up Nil and Cons to work as fixrec patterns
huffman
parents: 39199
diff changeset
   325
  "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
   326
unfolding match_Cons_def by simp_all
3989b2b44dba set up Nil and Cons to work as fixrec patterns
huffman
parents: 39199
diff changeset
   327
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
   328
setup \<open>
39212
3989b2b44dba set up Nil and Cons to work as fixrec patterns
huffman
parents: 39199
diff changeset
   329
  Fixrec.add_matchers
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68780
diff changeset
   330
    [ (\<^const_name>\<open>Nil\<close>, \<^const_name>\<open>match_Nil\<close>),
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 68780
diff changeset
   331
      (\<^const_name>\<open>Cons\<close>, \<^const_name>\<open>match_Cons\<close>) ]
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 58880
diff changeset
   332
\<close>
39212
3989b2b44dba set up Nil and Cons to work as fixrec patterns
huffman
parents: 39199
diff changeset
   333
39143
d80990d8b909 add List_Cpo.thy to HOLCF/Library
huffman
parents:
diff changeset
   334
end