1(*  Title:      HOL/HOLCF/Fixrec.thy
2    Author:     Amber Telfer and Brian Huffman
3*)
4
5section "Package for defining recursive functions in HOLCF"
6
7theory Fixrec
8imports Cprod Sprod Ssum Up One Tr Fix
9keywords "fixrec" :: thy_decl
10begin
11
12subsection \<open>Pattern-match monad\<close>
13
14default_sort cpo
15
16pcpodef 'a match = "UNIV::(one ++ 'a u) set"
17by simp_all
18
19definition
20  fail :: "'a match" where
21  "fail = Abs_match (sinl\<cdot>ONE)"
22
23definition
24  succeed :: "'a \<rightarrow> 'a match" where
25  "succeed = (\<Lambda> x. Abs_match (sinr\<cdot>(up\<cdot>x)))"
26
27lemma matchE [case_names bottom fail succeed, cases type: match]:
28  "\<lbrakk>p = \<bottom> \<Longrightarrow> Q; p = fail \<Longrightarrow> Q; \<And>x. p = succeed\<cdot>x \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
29unfolding fail_def succeed_def
30apply (cases p, rename_tac r)
31apply (rule_tac p=r in ssumE, simp add: Abs_match_strict)
32apply (rule_tac p=x in oneE, simp, simp)
33apply (rule_tac p=y in upE, simp, simp add: cont_Abs_match)
34done
35
36lemma succeed_defined [simp]: "succeed\<cdot>x \<noteq> \<bottom>"
37by (simp add: succeed_def cont_Abs_match Abs_match_bottom_iff)
38
39lemma fail_defined [simp]: "fail \<noteq> \<bottom>"
40by (simp add: fail_def Abs_match_bottom_iff)
41
42lemma succeed_eq [simp]: "(succeed\<cdot>x = succeed\<cdot>y) = (x = y)"
43by (simp add: succeed_def cont_Abs_match Abs_match_inject)
44
45lemma succeed_neq_fail [simp]:
46  "succeed\<cdot>x \<noteq> fail" "fail \<noteq> succeed\<cdot>x"
47by (simp_all add: succeed_def fail_def cont_Abs_match Abs_match_inject)
48
49subsubsection \<open>Run operator\<close>
50
51definition
52  run :: "'a match \<rightarrow> 'a::pcpo" where
53  "run = (\<Lambda> m. sscase\<cdot>\<bottom>\<cdot>(fup\<cdot>ID)\<cdot>(Rep_match m))"
54
55text \<open>rewrite rules for run\<close>
56
57lemma run_strict [simp]: "run\<cdot>\<bottom> = \<bottom>"
58unfolding run_def
59by (simp add: cont_Rep_match Rep_match_strict)
60
61lemma run_fail [simp]: "run\<cdot>fail = \<bottom>"
62unfolding run_def fail_def
63by (simp add: cont_Rep_match Abs_match_inverse)
64
65lemma run_succeed [simp]: "run\<cdot>(succeed\<cdot>x) = x"
66unfolding run_def succeed_def
67by (simp add: cont_Rep_match cont_Abs_match Abs_match_inverse)
68
69subsubsection \<open>Monad plus operator\<close>
70
71definition
72  mplus :: "'a match \<rightarrow> 'a match \<rightarrow> 'a match" where
73  "mplus = (\<Lambda> m1 m2. sscase\<cdot>(\<Lambda> _. m2)\<cdot>(\<Lambda> _. m1)\<cdot>(Rep_match m1))"
74
75abbreviation
76  mplus_syn :: "['a match, 'a match] \<Rightarrow> 'a match"  (infixr "+++" 65)  where
77  "m1 +++ m2 == mplus\<cdot>m1\<cdot>m2"
78
79text \<open>rewrite rules for mplus\<close>
80
81lemma mplus_strict [simp]: "\<bottom> +++ m = \<bottom>"
82unfolding mplus_def
83by (simp add: cont_Rep_match Rep_match_strict)
84
85lemma mplus_fail [simp]: "fail +++ m = m"
86unfolding mplus_def fail_def
87by (simp add: cont_Rep_match Abs_match_inverse)
88
89lemma mplus_succeed [simp]: "succeed\<cdot>x +++ m = succeed\<cdot>x"
90unfolding mplus_def succeed_def
91by (simp add: cont_Rep_match cont_Abs_match Abs_match_inverse)
92
93lemma mplus_fail2 [simp]: "m +++ fail = m"
94by (cases m, simp_all)
95
96lemma mplus_assoc: "(x +++ y) +++ z = x +++ (y +++ z)"
97by (cases x, simp_all)
98
99subsection \<open>Match functions for built-in types\<close>
100
101default_sort pcpo
102
103definition
104  match_bottom :: "'a \<rightarrow> 'c match \<rightarrow> 'c match"
105where
106  "match_bottom = (\<Lambda> x k. seq\<cdot>x\<cdot>fail)"
107
108definition
109  match_Pair :: "'a::cpo \<times> 'b::cpo \<rightarrow> ('a \<rightarrow> 'b \<rightarrow> 'c match) \<rightarrow> 'c match"
110where
111  "match_Pair = (\<Lambda> x k. csplit\<cdot>k\<cdot>x)"
112
113definition
114  match_spair :: "'a \<otimes> 'b \<rightarrow> ('a \<rightarrow> 'b \<rightarrow> 'c match) \<rightarrow> 'c match"
115where
116  "match_spair = (\<Lambda> x k. ssplit\<cdot>k\<cdot>x)"
117
118definition
119  match_sinl :: "'a \<oplus> 'b \<rightarrow> ('a \<rightarrow> 'c match) \<rightarrow> 'c match"
120where
121  "match_sinl = (\<Lambda> x k. sscase\<cdot>k\<cdot>(\<Lambda> b. fail)\<cdot>x)"
122
123definition
124  match_sinr :: "'a \<oplus> 'b \<rightarrow> ('b \<rightarrow> 'c match) \<rightarrow> 'c match"
125where
126  "match_sinr = (\<Lambda> x k. sscase\<cdot>(\<Lambda> a. fail)\<cdot>k\<cdot>x)"
127
128definition
129  match_up :: "'a::cpo u \<rightarrow> ('a \<rightarrow> 'c match) \<rightarrow> 'c match"
130where
131  "match_up = (\<Lambda> x k. fup\<cdot>k\<cdot>x)"
132
133definition
134  match_ONE :: "one \<rightarrow> 'c match \<rightarrow> 'c match"
135where
136  "match_ONE = (\<Lambda> ONE k. k)"
137
138definition
139  match_TT :: "tr \<rightarrow> 'c match \<rightarrow> 'c match"
140where
141  "match_TT = (\<Lambda> x k. If x then k else fail)"
142
143definition
144  match_FF :: "tr \<rightarrow> 'c match \<rightarrow> 'c match"
145where
146  "match_FF = (\<Lambda> x k. If x then fail else k)"
147
148lemma match_bottom_simps [simp]:
149  "match_bottom\<cdot>x\<cdot>k = (if x = \<bottom> then \<bottom> else fail)"
150by (simp add: match_bottom_def)
151
152lemma match_Pair_simps [simp]:
153  "match_Pair\<cdot>(x, y)\<cdot>k = k\<cdot>x\<cdot>y"
154by (simp_all add: match_Pair_def)
155
156lemma match_spair_simps [simp]:
157  "\<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> match_spair\<cdot>(:x, y:)\<cdot>k = k\<cdot>x\<cdot>y"
158  "match_spair\<cdot>\<bottom>\<cdot>k = \<bottom>"
159by (simp_all add: match_spair_def)
160
161lemma match_sinl_simps [simp]:
162  "x \<noteq> \<bottom> \<Longrightarrow> match_sinl\<cdot>(sinl\<cdot>x)\<cdot>k = k\<cdot>x"
163  "y \<noteq> \<bottom> \<Longrightarrow> match_sinl\<cdot>(sinr\<cdot>y)\<cdot>k = fail"
164  "match_sinl\<cdot>\<bottom>\<cdot>k = \<bottom>"
165by (simp_all add: match_sinl_def)
166
167lemma match_sinr_simps [simp]:
168  "x \<noteq> \<bottom> \<Longrightarrow> match_sinr\<cdot>(sinl\<cdot>x)\<cdot>k = fail"
169  "y \<noteq> \<bottom> \<Longrightarrow> match_sinr\<cdot>(sinr\<cdot>y)\<cdot>k = k\<cdot>y"
170  "match_sinr\<cdot>\<bottom>\<cdot>k = \<bottom>"
171by (simp_all add: match_sinr_def)
172
173lemma match_up_simps [simp]:
174  "match_up\<cdot>(up\<cdot>x)\<cdot>k = k\<cdot>x"
175  "match_up\<cdot>\<bottom>\<cdot>k = \<bottom>"
176by (simp_all add: match_up_def)
177
178lemma match_ONE_simps [simp]:
179  "match_ONE\<cdot>ONE\<cdot>k = k"
180  "match_ONE\<cdot>\<bottom>\<cdot>k = \<bottom>"
181by (simp_all add: match_ONE_def)
182
183lemma match_TT_simps [simp]:
184  "match_TT\<cdot>TT\<cdot>k = k"
185  "match_TT\<cdot>FF\<cdot>k = fail"
186  "match_TT\<cdot>\<bottom>\<cdot>k = \<bottom>"
187by (simp_all add: match_TT_def)
188
189lemma match_FF_simps [simp]:
190  "match_FF\<cdot>FF\<cdot>k = k"
191  "match_FF\<cdot>TT\<cdot>k = fail"
192  "match_FF\<cdot>\<bottom>\<cdot>k = \<bottom>"
193by (simp_all add: match_FF_def)
194
195subsection \<open>Mutual recursion\<close>
196
197text \<open>
198  The following rules are used to prove unfolding theorems from
199  fixed-point definitions of mutually recursive functions.
200\<close>
201
202lemma Pair_equalI: "\<lbrakk>x \<equiv> fst p; y \<equiv> snd p\<rbrakk> \<Longrightarrow> (x, y) \<equiv> p"
203by simp
204
205lemma Pair_eqD1: "(x, y) = (x', y') \<Longrightarrow> x = x'"
206by simp
207
208lemma Pair_eqD2: "(x, y) = (x', y') \<Longrightarrow> y = y'"
209by simp
210
211lemma def_cont_fix_eq:
212  "\<lbrakk>f \<equiv> fix\<cdot>(Abs_cfun F); cont F\<rbrakk> \<Longrightarrow> f = F f"
213by (simp, subst fix_eq, simp)
214
215lemma def_cont_fix_ind:
216  "\<lbrakk>f \<equiv> fix\<cdot>(Abs_cfun F); cont F; adm P; P \<bottom>; \<And>x. P x \<Longrightarrow> P (F x)\<rbrakk> \<Longrightarrow> P f"
217by (simp add: fix_ind)
218
219text \<open>lemma for proving rewrite rules\<close>
220
221lemma ssubst_lhs: "\<lbrakk>t = s; P s = Q\<rbrakk> \<Longrightarrow> P t = Q"
222by simp
223
224
225subsection \<open>Initializing the fixrec package\<close>
226
227ML_file "Tools/holcf_library.ML"
228ML_file "Tools/fixrec.ML"
229
230method_setup fixrec_simp = \<open>
231  Scan.succeed (SIMPLE_METHOD' o Fixrec.fixrec_simp_tac)
232\<close> "pattern prover for fixrec constants"
233
234setup \<open>
235  Fixrec.add_matchers
236    [ (@{const_name up}, @{const_name match_up}),
237      (@{const_name sinl}, @{const_name match_sinl}),
238      (@{const_name sinr}, @{const_name match_sinr}),
239      (@{const_name spair}, @{const_name match_spair}),
240      (@{const_name Pair}, @{const_name match_Pair}),
241      (@{const_name ONE}, @{const_name match_ONE}),
242      (@{const_name TT}, @{const_name match_TT}),
243      (@{const_name FF}, @{const_name match_FF}),
244      (@{const_name bottom}, @{const_name match_bottom}) ]
245\<close>
246
247hide_const (open) succeed fail run
248
249end
250