| author | blanchet | 
| Wed, 24 Sep 2014 15:46:25 +0200 | |
| changeset 58429 | 0b94858325a5 | 
| parent 58112 | 8081087096ad | 
| child 58880 | 0baae4311a9f | 
| permissions | -rw-r--r-- | 
| 42151 | 1  | 
(* Title: HOL/HOLCF/ex/Pattern_Match.thy  | 
| 37109 | 2  | 
Author: Brian Huffman  | 
3  | 
*)  | 
|
4  | 
||
5  | 
header {* An experimental pattern-matching notation *}
 | 
|
6  | 
||
7  | 
theory Pattern_Match  | 
|
8  | 
imports HOLCF  | 
|
9  | 
begin  | 
|
10  | 
||
| 
40329
 
73f2b99b549d
change default_sort of HOLCF from pcpo to bifinite; rename command 'new_domain' to 'domain'; rename 'domain' to 'domain (unsafe)'
 
huffman 
parents: 
40327 
diff
changeset
 | 
11  | 
default_sort pcpo  | 
| 
 
73f2b99b549d
change default_sort of HOLCF from pcpo to bifinite; rename command 'new_domain' to 'domain'; rename 'domain' to 'domain (unsafe)'
 
huffman 
parents: 
40327 
diff
changeset
 | 
12  | 
|
| 37109 | 13  | 
text {* FIXME: Find a proper way to un-hide constants. *}
 | 
14  | 
||
15  | 
abbreviation fail :: "'a match"  | 
|
16  | 
where "fail \<equiv> Fixrec.fail"  | 
|
17  | 
||
18  | 
abbreviation succeed :: "'a \<rightarrow> 'a match"  | 
|
19  | 
where "succeed \<equiv> Fixrec.succeed"  | 
|
20  | 
||
21  | 
abbreviation run :: "'a match \<rightarrow> 'a"  | 
|
22  | 
where "run \<equiv> Fixrec.run"  | 
|
23  | 
||
24  | 
subsection {* Fatbar combinator *}
 | 
|
25  | 
||
26  | 
definition  | 
|
27  | 
  fatbar :: "('a \<rightarrow> 'b match) \<rightarrow> ('a \<rightarrow> 'b match) \<rightarrow> ('a \<rightarrow> 'b match)" where
 | 
|
28  | 
"fatbar = (\<Lambda> a b x. a\<cdot>x +++ b\<cdot>x)"  | 
|
29  | 
||
30  | 
abbreviation  | 
|
31  | 
fatbar_syn :: "['a \<rightarrow> 'b match, 'a \<rightarrow> 'b match] \<Rightarrow> 'a \<rightarrow> 'b match" (infixr "\<parallel>" 60) where  | 
|
32  | 
"m1 \<parallel> m2 == fatbar\<cdot>m1\<cdot>m2"  | 
|
33  | 
||
34  | 
lemma fatbar1: "m\<cdot>x = \<bottom> \<Longrightarrow> (m \<parallel> ms)\<cdot>x = \<bottom>"  | 
|
35  | 
by (simp add: fatbar_def)  | 
|
36  | 
||
37  | 
lemma fatbar2: "m\<cdot>x = fail \<Longrightarrow> (m \<parallel> ms)\<cdot>x = ms\<cdot>x"  | 
|
38  | 
by (simp add: fatbar_def)  | 
|
39  | 
||
40  | 
lemma fatbar3: "m\<cdot>x = succeed\<cdot>y \<Longrightarrow> (m \<parallel> ms)\<cdot>x = succeed\<cdot>y"  | 
|
41  | 
by (simp add: fatbar_def)  | 
|
42  | 
||
43  | 
lemmas fatbar_simps = fatbar1 fatbar2 fatbar3  | 
|
44  | 
||
45  | 
lemma run_fatbar1: "m\<cdot>x = \<bottom> \<Longrightarrow> run\<cdot>((m \<parallel> ms)\<cdot>x) = \<bottom>"  | 
|
46  | 
by (simp add: fatbar_def)  | 
|
47  | 
||
48  | 
lemma run_fatbar2: "m\<cdot>x = fail \<Longrightarrow> run\<cdot>((m \<parallel> ms)\<cdot>x) = run\<cdot>(ms\<cdot>x)"  | 
|
49  | 
by (simp add: fatbar_def)  | 
|
50  | 
||
51  | 
lemma run_fatbar3: "m\<cdot>x = succeed\<cdot>y \<Longrightarrow> run\<cdot>((m \<parallel> ms)\<cdot>x) = y"  | 
|
52  | 
by (simp add: fatbar_def)  | 
|
53  | 
||
54  | 
lemmas run_fatbar_simps [simp] = run_fatbar1 run_fatbar2 run_fatbar3  | 
|
55  | 
||
| 40735 | 56  | 
subsection {* Bind operator for match monad *}
 | 
57  | 
||
58  | 
definition match_bind :: "'a match \<rightarrow> ('a \<rightarrow> 'b match) \<rightarrow> 'b match" where
 | 
|
59  | 
"match_bind = (\<Lambda> m k. sscase\<cdot>(\<Lambda> _. fail)\<cdot>(fup\<cdot>k)\<cdot>(Rep_match m))"  | 
|
60  | 
||
61  | 
lemma match_bind_simps [simp]:  | 
|
62  | 
"match_bind\<cdot>\<bottom>\<cdot>k = \<bottom>"  | 
|
63  | 
"match_bind\<cdot>fail\<cdot>k = fail"  | 
|
64  | 
"match_bind\<cdot>(succeed\<cdot>x)\<cdot>k = k\<cdot>x"  | 
|
65  | 
unfolding match_bind_def fail_def succeed_def  | 
|
| 
40834
 
a1249aeff5b6
change cpodef-generated cont_Rep rules to cont2cont format
 
huffman 
parents: 
40774 
diff
changeset
 | 
66  | 
by (simp_all add: cont_Rep_match cont_Abs_match  | 
| 40735 | 67  | 
Rep_match_strict Abs_match_inverse)  | 
68  | 
||
| 37109 | 69  | 
subsection {* Case branch combinator *}
 | 
70  | 
||
71  | 
definition  | 
|
72  | 
  branch :: "('a \<rightarrow> 'b match) \<Rightarrow> ('b \<rightarrow> 'c) \<rightarrow> ('a \<rightarrow> 'c match)" where
 | 
|
| 40735 | 73  | 
"branch p \<equiv> \<Lambda> r x. match_bind\<cdot>(p\<cdot>x)\<cdot>(\<Lambda> y. succeed\<cdot>(r\<cdot>y))"  | 
| 37109 | 74  | 
|
75  | 
lemma branch_simps:  | 
|
76  | 
"p\<cdot>x = \<bottom> \<Longrightarrow> branch p\<cdot>r\<cdot>x = \<bottom>"  | 
|
77  | 
"p\<cdot>x = fail \<Longrightarrow> branch p\<cdot>r\<cdot>x = fail"  | 
|
78  | 
"p\<cdot>x = succeed\<cdot>y \<Longrightarrow> branch p\<cdot>r\<cdot>x = succeed\<cdot>(r\<cdot>y)"  | 
|
79  | 
by (simp_all add: branch_def)  | 
|
80  | 
||
81  | 
lemma branch_succeed [simp]: "branch succeed\<cdot>r\<cdot>x = succeed\<cdot>(r\<cdot>x)"  | 
|
82  | 
by (simp add: branch_def)  | 
|
83  | 
||
84  | 
subsection {* Cases operator *}
 | 
|
85  | 
||
86  | 
definition  | 
|
87  | 
cases :: "'a match \<rightarrow> 'a::pcpo" where  | 
|
| 40735 | 88  | 
"cases = Fixrec.run"  | 
| 37109 | 89  | 
|
90  | 
text {* rewrite rules for cases *}
 | 
|
91  | 
||
92  | 
lemma cases_strict [simp]: "cases\<cdot>\<bottom> = \<bottom>"  | 
|
93  | 
by (simp add: cases_def)  | 
|
94  | 
||
95  | 
lemma cases_fail [simp]: "cases\<cdot>fail = \<bottom>"  | 
|
96  | 
by (simp add: cases_def)  | 
|
97  | 
||
98  | 
lemma cases_succeed [simp]: "cases\<cdot>(succeed\<cdot>x) = x"  | 
|
99  | 
by (simp add: cases_def)  | 
|
100  | 
||
101  | 
subsection {* Case syntax *}
 | 
|
102  | 
||
| 
42057
 
3eba96ff3d3e
more selective strip_positions in case patterns -- reactivate translations based on "case _ of _" in HOL and special patterns in HOLCF;
 
wenzelm 
parents: 
41476 
diff
changeset
 | 
103  | 
nonterminal Case_pat and Case_syn and Cases_syn  | 
| 37109 | 104  | 
|
105  | 
syntax  | 
|
106  | 
  "_Case_syntax":: "['a, Cases_syn] => 'b"               ("(Case _ of/ _)" 10)
 | 
|
| 
42057
 
3eba96ff3d3e
more selective strip_positions in case patterns -- reactivate translations based on "case _ of _" in HOL and special patterns in HOLCF;
 
wenzelm 
parents: 
41476 
diff
changeset
 | 
107  | 
  "_Case1"      :: "[Case_pat, 'b] => Case_syn"          ("(2_ =>/ _)" 10)
 | 
| 37109 | 108  | 
  ""            :: "Case_syn => Cases_syn"               ("_")
 | 
109  | 
  "_Case2"      :: "[Case_syn, Cases_syn] => Cases_syn"  ("_/ | _")
 | 
|
| 
42057
 
3eba96ff3d3e
more selective strip_positions in case patterns -- reactivate translations based on "case _ of _" in HOL and special patterns in HOLCF;
 
wenzelm 
parents: 
41476 
diff
changeset
 | 
110  | 
  "_strip_positions" :: "'a => Case_pat"                 ("_")
 | 
| 37109 | 111  | 
|
112  | 
syntax (xsymbols)  | 
|
| 
42057
 
3eba96ff3d3e
more selective strip_positions in case patterns -- reactivate translations based on "case _ of _" in HOL and special patterns in HOLCF;
 
wenzelm 
parents: 
41476 
diff
changeset
 | 
113  | 
  "_Case1"      :: "[Case_pat, 'b] => Case_syn"          ("(2_ \<Rightarrow>/ _)" 10)
 | 
| 37109 | 114  | 
|
115  | 
translations  | 
|
116  | 
"_Case_syntax x ms" == "CONST cases\<cdot>(ms\<cdot>x)"  | 
|
117  | 
"_Case2 m ms" == "m \<parallel> ms"  | 
|
118  | 
||
119  | 
text {* Parsing Case expressions *}
 | 
|
120  | 
||
121  | 
syntax  | 
|
122  | 
"_pat" :: "'a"  | 
|
123  | 
"_variable" :: "'a"  | 
|
124  | 
"_noargs" :: "'a"  | 
|
125  | 
||
126  | 
translations  | 
|
127  | 
"_Case1 p r" => "CONST branch (_pat p)\<cdot>(_variable p r)"  | 
|
128  | 
"_variable (_args x y) r" => "CONST csplit\<cdot>(_variable x (_variable y r))"  | 
|
129  | 
"_variable _noargs r" => "CONST unit_when\<cdot>r"  | 
|
130  | 
||
131  | 
parse_translation {*
 | 
|
132  | 
(* rewrite (_pat x) => (succeed) *)  | 
|
| 40327 | 133  | 
(* rewrite (_variable x t) => (Abs_cfun (%x. t)) *)  | 
| 52143 | 134  | 
 [(@{syntax_const "_pat"}, fn _ => fn _ => Syntax.const @{const_syntax Fixrec.succeed}),
 | 
| 42284 | 135  | 
  Syntax_Trans.mk_binder_tr (@{syntax_const "_variable"}, @{const_syntax Abs_cfun})];
 | 
| 37109 | 136  | 
*}  | 
137  | 
||
138  | 
text {* Printing Case expressions *}
 | 
|
139  | 
||
140  | 
syntax  | 
|
141  | 
"_match" :: "'a"  | 
|
142  | 
||
143  | 
print_translation {*
 | 
|
144  | 
let  | 
|
| 40327 | 145  | 
    fun dest_LAM (Const (@{const_syntax Rep_cfun},_) $ Const (@{const_syntax unit_when},_) $ t) =
 | 
| 37109 | 146  | 
          (Syntax.const @{syntax_const "_noargs"}, t)
 | 
| 40327 | 147  | 
    |   dest_LAM (Const (@{const_syntax Rep_cfun},_) $ Const (@{const_syntax csplit},_) $ t) =
 | 
| 37109 | 148  | 
let  | 
149  | 
val (v1, t1) = dest_LAM t;  | 
|
150  | 
val (v2, t2) = dest_LAM t1;  | 
|
151  | 
          in (Syntax.const @{syntax_const "_args"} $ v1 $ v2, t2) end
 | 
|
| 40327 | 152  | 
    |   dest_LAM (Const (@{const_syntax Abs_cfun},_) $ t) =
 | 
| 37109 | 153  | 
let  | 
154  | 
val abs =  | 
|
155  | 
case t of Abs abs => abs  | 
|
156  | 
                | _ => ("x", dummyT, incr_boundvars 1 t $ Bound 0);
 | 
|
| 42284 | 157  | 
val (x, t') = Syntax_Trans.atomic_abs_tr' abs;  | 
| 37109 | 158  | 
          in (Syntax.const @{syntax_const "_variable"} $ x, t') end
 | 
159  | 
| dest_LAM _ = raise Match; (* too few vars: abort translation *)  | 
|
160  | 
||
161  | 
    fun Case1_tr' [Const(@{const_syntax branch},_) $ p, r] =
 | 
|
162  | 
let val (v, t) = dest_LAM r in  | 
|
163  | 
            Syntax.const @{syntax_const "_Case1"} $
 | 
|
164  | 
              (Syntax.const @{syntax_const "_match"} $ p $ v) $ t
 | 
|
165  | 
end;  | 
|
166  | 
||
| 52143 | 167  | 
  in [(@{const_syntax Rep_cfun}, K Case1_tr')] end;
 | 
| 37109 | 168  | 
*}  | 
169  | 
||
170  | 
translations  | 
|
171  | 
"x" <= "_match (CONST succeed) (_variable x)"  | 
|
172  | 
||
173  | 
||
174  | 
subsection {* Pattern combinators for data constructors *}
 | 
|
175  | 
||
| 41476 | 176  | 
type_synonym ('a, 'b) pat = "'a \<rightarrow> 'b match"
 | 
| 37109 | 177  | 
|
178  | 
definition  | 
|
179  | 
  cpair_pat :: "('a, 'c) pat \<Rightarrow> ('b, 'd) pat \<Rightarrow> ('a \<times> 'b, 'c \<times> 'd) pat" where
 | 
|
180  | 
"cpair_pat p1 p2 = (\<Lambda>(x, y).  | 
|
| 40735 | 181  | 
match_bind\<cdot>(p1\<cdot>x)\<cdot>(\<Lambda> a. match_bind\<cdot>(p2\<cdot>y)\<cdot>(\<Lambda> b. succeed\<cdot>(a, b))))"  | 
| 37109 | 182  | 
|
183  | 
definition  | 
|
184  | 
spair_pat ::  | 
|
185  | 
  "('a, 'c) pat \<Rightarrow> ('b, 'd) pat \<Rightarrow> ('a::pcpo \<otimes> 'b::pcpo, 'c \<times> 'd) pat" where
 | 
|
186  | 
"spair_pat p1 p2 = (\<Lambda>(:x, y:). cpair_pat p1 p2\<cdot>(x, y))"  | 
|
187  | 
||
188  | 
definition  | 
|
189  | 
  sinl_pat :: "('a, 'c) pat \<Rightarrow> ('a::pcpo \<oplus> 'b::pcpo, 'c) pat" where
 | 
|
190  | 
"sinl_pat p = sscase\<cdot>p\<cdot>(\<Lambda> x. fail)"  | 
|
191  | 
||
192  | 
definition  | 
|
193  | 
  sinr_pat :: "('b, 'c) pat \<Rightarrow> ('a::pcpo \<oplus> 'b::pcpo, 'c) pat" where
 | 
|
194  | 
"sinr_pat p = sscase\<cdot>(\<Lambda> x. fail)\<cdot>p"  | 
|
195  | 
||
196  | 
definition  | 
|
197  | 
  up_pat :: "('a, 'b) pat \<Rightarrow> ('a u, 'b) pat" where
 | 
|
198  | 
"up_pat p = fup\<cdot>p"  | 
|
199  | 
||
200  | 
definition  | 
|
201  | 
TT_pat :: "(tr, unit) pat" where  | 
|
| 
40322
 
707eb30e8a53
make syntax of continuous if-then-else consistent with HOL if-then-else
 
huffman 
parents: 
40026 
diff
changeset
 | 
202  | 
"TT_pat = (\<Lambda> b. If b then succeed\<cdot>() else fail)"  | 
| 37109 | 203  | 
|
204  | 
definition  | 
|
205  | 
FF_pat :: "(tr, unit) pat" where  | 
|
| 
40322
 
707eb30e8a53
make syntax of continuous if-then-else consistent with HOL if-then-else
 
huffman 
parents: 
40026 
diff
changeset
 | 
206  | 
"FF_pat = (\<Lambda> b. If b then fail else succeed\<cdot>())"  | 
| 37109 | 207  | 
|
208  | 
definition  | 
|
209  | 
ONE_pat :: "(one, unit) pat" where  | 
|
210  | 
"ONE_pat = (\<Lambda> ONE. succeed\<cdot>())"  | 
|
211  | 
||
212  | 
text {* Parse translations (patterns) *}
 | 
|
213  | 
translations  | 
|
214  | 
"_pat (XCONST Pair x y)" => "CONST cpair_pat (_pat x) (_pat y)"  | 
|
215  | 
"_pat (XCONST spair\<cdot>x\<cdot>y)" => "CONST spair_pat (_pat x) (_pat y)"  | 
|
216  | 
"_pat (XCONST sinl\<cdot>x)" => "CONST sinl_pat (_pat x)"  | 
|
217  | 
"_pat (XCONST sinr\<cdot>x)" => "CONST sinr_pat (_pat x)"  | 
|
218  | 
"_pat (XCONST up\<cdot>x)" => "CONST up_pat (_pat x)"  | 
|
219  | 
"_pat (XCONST TT)" => "CONST TT_pat"  | 
|
220  | 
"_pat (XCONST FF)" => "CONST FF_pat"  | 
|
221  | 
"_pat (XCONST ONE)" => "CONST ONE_pat"  | 
|
222  | 
||
223  | 
text {* CONST version is also needed for constructors with special syntax *}
 | 
|
224  | 
translations  | 
|
225  | 
"_pat (CONST Pair x y)" => "CONST cpair_pat (_pat x) (_pat y)"  | 
|
226  | 
"_pat (CONST spair\<cdot>x\<cdot>y)" => "CONST spair_pat (_pat x) (_pat y)"  | 
|
227  | 
||
228  | 
text {* Parse translations (variables) *}
 | 
|
229  | 
translations  | 
|
230  | 
"_variable (XCONST Pair x y) r" => "_variable (_args x y) r"  | 
|
231  | 
"_variable (XCONST spair\<cdot>x\<cdot>y) r" => "_variable (_args x y) r"  | 
|
232  | 
"_variable (XCONST sinl\<cdot>x) r" => "_variable x r"  | 
|
233  | 
"_variable (XCONST sinr\<cdot>x) r" => "_variable x r"  | 
|
234  | 
"_variable (XCONST up\<cdot>x) r" => "_variable x r"  | 
|
235  | 
"_variable (XCONST TT) r" => "_variable _noargs r"  | 
|
236  | 
"_variable (XCONST FF) r" => "_variable _noargs r"  | 
|
237  | 
"_variable (XCONST ONE) r" => "_variable _noargs r"  | 
|
238  | 
||
239  | 
translations  | 
|
240  | 
"_variable (CONST Pair x y) r" => "_variable (_args x y) r"  | 
|
241  | 
"_variable (CONST spair\<cdot>x\<cdot>y) r" => "_variable (_args x y) r"  | 
|
242  | 
||
243  | 
text {* Print translations *}
 | 
|
244  | 
translations  | 
|
245  | 
"CONST Pair (_match p1 v1) (_match p2 v2)"  | 
|
246  | 
<= "_match (CONST cpair_pat p1 p2) (_args v1 v2)"  | 
|
247  | 
"CONST spair\<cdot>(_match p1 v1)\<cdot>(_match p2 v2)"  | 
|
248  | 
<= "_match (CONST spair_pat p1 p2) (_args v1 v2)"  | 
|
249  | 
"CONST sinl\<cdot>(_match p1 v1)" <= "_match (CONST sinl_pat p1) v1"  | 
|
250  | 
"CONST sinr\<cdot>(_match p1 v1)" <= "_match (CONST sinr_pat p1) v1"  | 
|
251  | 
"CONST up\<cdot>(_match p1 v1)" <= "_match (CONST up_pat p1) v1"  | 
|
252  | 
"CONST TT" <= "_match (CONST TT_pat) _noargs"  | 
|
253  | 
"CONST FF" <= "_match (CONST FF_pat) _noargs"  | 
|
254  | 
"CONST ONE" <= "_match (CONST ONE_pat) _noargs"  | 
|
255  | 
||
256  | 
lemma cpair_pat1:  | 
|
257  | 
"branch p\<cdot>r\<cdot>x = \<bottom> \<Longrightarrow> branch (cpair_pat p q)\<cdot>(csplit\<cdot>r)\<cdot>(x, y) = \<bottom>"  | 
|
258  | 
apply (simp add: branch_def cpair_pat_def)  | 
|
259  | 
apply (cases "p\<cdot>x", simp_all)  | 
|
260  | 
done  | 
|
261  | 
||
262  | 
lemma cpair_pat2:  | 
|
263  | 
"branch p\<cdot>r\<cdot>x = fail \<Longrightarrow> branch (cpair_pat p q)\<cdot>(csplit\<cdot>r)\<cdot>(x, y) = fail"  | 
|
264  | 
apply (simp add: branch_def cpair_pat_def)  | 
|
265  | 
apply (cases "p\<cdot>x", simp_all)  | 
|
266  | 
done  | 
|
267  | 
||
268  | 
lemma cpair_pat3:  | 
|
269  | 
"branch p\<cdot>r\<cdot>x = succeed\<cdot>s \<Longrightarrow>  | 
|
270  | 
branch (cpair_pat p q)\<cdot>(csplit\<cdot>r)\<cdot>(x, y) = branch q\<cdot>s\<cdot>y"  | 
|
271  | 
apply (simp add: branch_def cpair_pat_def)  | 
|
272  | 
apply (cases "p\<cdot>x", simp_all)  | 
|
273  | 
apply (cases "q\<cdot>y", simp_all)  | 
|
274  | 
done  | 
|
275  | 
||
276  | 
lemmas cpair_pat [simp] =  | 
|
277  | 
cpair_pat1 cpair_pat2 cpair_pat3  | 
|
278  | 
||
279  | 
lemma spair_pat [simp]:  | 
|
280  | 
"branch (spair_pat p1 p2)\<cdot>r\<cdot>\<bottom> = \<bottom>"  | 
|
281  | 
"\<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>\<rbrakk>  | 
|
282  | 
\<Longrightarrow> branch (spair_pat p1 p2)\<cdot>r\<cdot>(:x, y:) =  | 
|
283  | 
branch (cpair_pat p1 p2)\<cdot>r\<cdot>(x, y)"  | 
|
284  | 
by (simp_all add: branch_def spair_pat_def)  | 
|
285  | 
||
286  | 
lemma sinl_pat [simp]:  | 
|
287  | 
"branch (sinl_pat p)\<cdot>r\<cdot>\<bottom> = \<bottom>"  | 
|
288  | 
"x \<noteq> \<bottom> \<Longrightarrow> branch (sinl_pat p)\<cdot>r\<cdot>(sinl\<cdot>x) = branch p\<cdot>r\<cdot>x"  | 
|
289  | 
"y \<noteq> \<bottom> \<Longrightarrow> branch (sinl_pat p)\<cdot>r\<cdot>(sinr\<cdot>y) = fail"  | 
|
290  | 
by (simp_all add: branch_def sinl_pat_def)  | 
|
291  | 
||
292  | 
lemma sinr_pat [simp]:  | 
|
293  | 
"branch (sinr_pat p)\<cdot>r\<cdot>\<bottom> = \<bottom>"  | 
|
294  | 
"x \<noteq> \<bottom> \<Longrightarrow> branch (sinr_pat p)\<cdot>r\<cdot>(sinl\<cdot>x) = fail"  | 
|
295  | 
"y \<noteq> \<bottom> \<Longrightarrow> branch (sinr_pat p)\<cdot>r\<cdot>(sinr\<cdot>y) = branch p\<cdot>r\<cdot>y"  | 
|
296  | 
by (simp_all add: branch_def sinr_pat_def)  | 
|
297  | 
||
298  | 
lemma up_pat [simp]:  | 
|
299  | 
"branch (up_pat p)\<cdot>r\<cdot>\<bottom> = \<bottom>"  | 
|
300  | 
"branch (up_pat p)\<cdot>r\<cdot>(up\<cdot>x) = branch p\<cdot>r\<cdot>x"  | 
|
301  | 
by (simp_all add: branch_def up_pat_def)  | 
|
302  | 
||
303  | 
lemma TT_pat [simp]:  | 
|
304  | 
"branch TT_pat\<cdot>(unit_when\<cdot>r)\<cdot>\<bottom> = \<bottom>"  | 
|
305  | 
"branch TT_pat\<cdot>(unit_when\<cdot>r)\<cdot>TT = succeed\<cdot>r"  | 
|
306  | 
"branch TT_pat\<cdot>(unit_when\<cdot>r)\<cdot>FF = fail"  | 
|
307  | 
by (simp_all add: branch_def TT_pat_def)  | 
|
308  | 
||
309  | 
lemma FF_pat [simp]:  | 
|
310  | 
"branch FF_pat\<cdot>(unit_when\<cdot>r)\<cdot>\<bottom> = \<bottom>"  | 
|
311  | 
"branch FF_pat\<cdot>(unit_when\<cdot>r)\<cdot>TT = fail"  | 
|
312  | 
"branch FF_pat\<cdot>(unit_when\<cdot>r)\<cdot>FF = succeed\<cdot>r"  | 
|
313  | 
by (simp_all add: branch_def FF_pat_def)  | 
|
314  | 
||
315  | 
lemma ONE_pat [simp]:  | 
|
316  | 
"branch ONE_pat\<cdot>(unit_when\<cdot>r)\<cdot>\<bottom> = \<bottom>"  | 
|
317  | 
"branch ONE_pat\<cdot>(unit_when\<cdot>r)\<cdot>ONE = succeed\<cdot>r"  | 
|
318  | 
by (simp_all add: branch_def ONE_pat_def)  | 
|
319  | 
||
320  | 
||
321  | 
subsection {* Wildcards, as-patterns, and lazy patterns *}
 | 
|
322  | 
||
323  | 
definition  | 
|
324  | 
wild_pat :: "'a \<rightarrow> unit match" where  | 
|
325  | 
"wild_pat = (\<Lambda> x. succeed\<cdot>())"  | 
|
326  | 
||
327  | 
definition  | 
|
328  | 
  as_pat :: "('a \<rightarrow> 'b match) \<Rightarrow> 'a \<rightarrow> ('a \<times> 'b) match" where
 | 
|
| 40735 | 329  | 
"as_pat p = (\<Lambda> x. match_bind\<cdot>(p\<cdot>x)\<cdot>(\<Lambda> a. succeed\<cdot>(x, a)))"  | 
| 37109 | 330  | 
|
331  | 
definition  | 
|
332  | 
  lazy_pat :: "('a \<rightarrow> 'b::pcpo match) \<Rightarrow> ('a \<rightarrow> 'b match)" where
 | 
|
333  | 
"lazy_pat p = (\<Lambda> x. succeed\<cdot>(cases\<cdot>(p\<cdot>x)))"  | 
|
334  | 
||
335  | 
text {* Parse translations (patterns) *}
 | 
|
336  | 
translations  | 
|
337  | 
"_pat _" => "CONST wild_pat"  | 
|
338  | 
||
339  | 
text {* Parse translations (variables) *}
 | 
|
340  | 
translations  | 
|
341  | 
"_variable _ r" => "_variable _noargs r"  | 
|
342  | 
||
343  | 
text {* Print translations *}
 | 
|
344  | 
translations  | 
|
345  | 
"_" <= "_match (CONST wild_pat) _noargs"  | 
|
346  | 
||
347  | 
lemma wild_pat [simp]: "branch wild_pat\<cdot>(unit_when\<cdot>r)\<cdot>x = succeed\<cdot>r"  | 
|
348  | 
by (simp add: branch_def wild_pat_def)  | 
|
349  | 
||
350  | 
lemma as_pat [simp]:  | 
|
351  | 
"branch (as_pat p)\<cdot>(csplit\<cdot>r)\<cdot>x = branch p\<cdot>(r\<cdot>x)\<cdot>x"  | 
|
352  | 
apply (simp add: branch_def as_pat_def)  | 
|
353  | 
apply (cases "p\<cdot>x", simp_all)  | 
|
354  | 
done  | 
|
355  | 
||
356  | 
lemma lazy_pat [simp]:  | 
|
357  | 
"branch p\<cdot>r\<cdot>x = \<bottom> \<Longrightarrow> branch (lazy_pat p)\<cdot>r\<cdot>x = succeed\<cdot>(r\<cdot>\<bottom>)"  | 
|
358  | 
"branch p\<cdot>r\<cdot>x = fail \<Longrightarrow> branch (lazy_pat p)\<cdot>r\<cdot>x = succeed\<cdot>(r\<cdot>\<bottom>)"  | 
|
359  | 
"branch p\<cdot>r\<cdot>x = succeed\<cdot>s \<Longrightarrow> branch (lazy_pat p)\<cdot>r\<cdot>x = succeed\<cdot>s"  | 
|
360  | 
apply (simp_all add: branch_def lazy_pat_def)  | 
|
361  | 
apply (cases "p\<cdot>x", simp_all)+  | 
|
362  | 
done  | 
|
363  | 
||
364  | 
subsection {* Examples *}
 | 
|
365  | 
||
366  | 
term "Case t of (:up\<cdot>(sinl\<cdot>x), sinr\<cdot>y:) \<Rightarrow> (x, y)"  | 
|
367  | 
||
368  | 
term "\<Lambda> t. Case t of up\<cdot>(sinl\<cdot>a) \<Rightarrow> a | up\<cdot>(sinr\<cdot>b) \<Rightarrow> b"  | 
|
369  | 
||
370  | 
term "\<Lambda> t. Case t of (:up\<cdot>(sinl\<cdot>_), sinr\<cdot>x:) \<Rightarrow> x"  | 
|
371  | 
||
372  | 
subsection {* ML code for generating definitions *}
 | 
|
373  | 
||
374  | 
ML {*
 | 
|
375  | 
local open HOLCF_Library in  | 
|
376  | 
||
| 
40026
 
8f8f18a88685
remove last few dependencies on domain_library.ML and delete it; temporarily disable emptiness check
 
huffman 
parents: 
39557 
diff
changeset
 | 
377  | 
infixr 6 ->>;  | 
| 
 
8f8f18a88685
remove last few dependencies on domain_library.ML and delete it; temporarily disable emptiness check
 
huffman 
parents: 
39557 
diff
changeset
 | 
378  | 
infix 9 ` ;  | 
| 
 
8f8f18a88685
remove last few dependencies on domain_library.ML and delete it; temporarily disable emptiness check
 
huffman 
parents: 
39557 
diff
changeset
 | 
379  | 
|
| 37109 | 380  | 
val beta_rules =  | 
| 
40326
 
73d45866dbda
renamed lemma cont2cont_Rep_CFun to cont2cont_APP
 
huffman 
parents: 
40322 
diff
changeset
 | 
381  | 
  @{thms beta_cfun cont_id cont_const cont2cont_APP cont2cont_LAM'} @
 | 
| 37109 | 382  | 
  @{thms cont2cont_fst cont2cont_snd cont2cont_Pair};
 | 
383  | 
||
| 
51717
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
46909 
diff
changeset
 | 
384  | 
val beta_ss =  | 
| 
 
9e7d1c139569
simplifier uses proper Proof.context instead of historic type simpset;
 
wenzelm 
parents: 
46909 
diff
changeset
 | 
385  | 
  simpset_of (put_simpset HOL_basic_ss @{context} addsimps (@{thms simp_thms} @ beta_rules));
 | 
| 37109 | 386  | 
|
387  | 
fun define_consts  | 
|
388  | 
(specs : (binding * term * mixfix) list)  | 
|
389  | 
(thy : theory)  | 
|
390  | 
: (term list * thm list) * theory =  | 
|
391  | 
let  | 
|
392  | 
fun mk_decl (b, t, mx) = (b, fastype_of t, mx);  | 
|
393  | 
val decls = map mk_decl specs;  | 
|
394  | 
val thy = Cont_Consts.add_consts decls thy;  | 
|
395  | 
fun mk_const (b, T, mx) = Const (Sign.full_name thy b, T);  | 
|
396  | 
val consts = map mk_const decls;  | 
|
397  | 
fun mk_def c (b, t, mx) =  | 
|
| 46909 | 398  | 
(Thm.def_binding b, Logic.mk_equals (c, t));  | 
| 37109 | 399  | 
val defs = map2 mk_def consts specs;  | 
400  | 
val (def_thms, thy) =  | 
|
| 
39557
 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 
wenzelm 
parents: 
37109 
diff
changeset
 | 
401  | 
Global_Theory.add_defs false (map Thm.no_attributes defs) thy;  | 
| 37109 | 402  | 
in  | 
403  | 
((consts, def_thms), thy)  | 
|
404  | 
end;  | 
|
405  | 
||
406  | 
fun prove  | 
|
407  | 
(thy : theory)  | 
|
408  | 
(defs : thm list)  | 
|
409  | 
(goal : term)  | 
|
410  | 
    (tacs : {prems: thm list, context: Proof.context} -> tactic list)
 | 
|
411  | 
: thm =  | 
|
412  | 
let  | 
|
413  | 
    fun tac {prems, context} =
 | 
|
| 
54742
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
52143 
diff
changeset
 | 
414  | 
rewrite_goals_tac context defs THEN  | 
| 
 
7a86358a3c0b
proper context for basic Simplifier operations: rewrite_rule, rewrite_goals_rule, rewrite_goals_tac etc.;
 
wenzelm 
parents: 
52143 
diff
changeset
 | 
415  | 
      EVERY (tacs {prems = map (rewrite_rule context defs) prems, context = context})
 | 
| 37109 | 416  | 
in  | 
417  | 
Goal.prove_global thy [] [] goal tac  | 
|
418  | 
end;  | 
|
419  | 
||
420  | 
fun get_vars_avoiding  | 
|
421  | 
(taken : string list)  | 
|
422  | 
(args : (bool * typ) list)  | 
|
423  | 
: (term list * term list) =  | 
|
424  | 
let  | 
|
425  | 
val Ts = map snd args;  | 
|
| 
58112
 
8081087096ad
renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
 
blanchet 
parents: 
54895 
diff
changeset
 | 
426  | 
val ns = Name.variant_list taken (Old_Datatype_Prop.make_tnames Ts);  | 
| 37109 | 427  | 
val vs = map Free (ns ~~ Ts);  | 
428  | 
val nonlazy = map snd (filter_out (fst o fst) (args ~~ vs));  | 
|
429  | 
in  | 
|
430  | 
(vs, nonlazy)  | 
|
431  | 
end;  | 
|
432  | 
||
433  | 
(******************************************************************************)  | 
|
434  | 
(************** definitions and theorems for pattern combinators **************)  | 
|
435  | 
(******************************************************************************)  | 
|
436  | 
||
437  | 
fun add_pattern_combinators  | 
|
438  | 
(bindings : binding list)  | 
|
439  | 
(spec : (term * (bool * typ) list) list)  | 
|
440  | 
(lhsT : typ)  | 
|
441  | 
(exhaust : thm)  | 
|
442  | 
(case_const : typ -> term)  | 
|
443  | 
(case_rews : thm list)  | 
|
444  | 
(thy : theory) =  | 
|
445  | 
let  | 
|
446  | 
||
447  | 
(* utility functions *)  | 
|
448  | 
fun mk_pair_pat (p1, p2) =  | 
|
449  | 
let  | 
|
450  | 
val T1 = fastype_of p1;  | 
|
451  | 
val T2 = fastype_of p2;  | 
|
452  | 
val (U1, V1) = apsnd dest_matchT (dest_cfunT T1);  | 
|
453  | 
val (U2, V2) = apsnd dest_matchT (dest_cfunT T2);  | 
|
454  | 
val pat_typ = [T1, T2] --->  | 
|
455  | 
(mk_prodT (U1, U2) ->> mk_matchT (mk_prodT (V1, V2)));  | 
|
456  | 
        val pat_const = Const (@{const_name cpair_pat}, pat_typ);
 | 
|
457  | 
in  | 
|
458  | 
pat_const $ p1 $ p2  | 
|
459  | 
end;  | 
|
460  | 
fun mk_tuple_pat [] = succeed_const HOLogic.unitT  | 
|
461  | 
| mk_tuple_pat ps = foldr1 mk_pair_pat ps;  | 
|
462  | 
fun branch_const (T,U,V) =  | 
|
463  | 
      Const (@{const_name branch},
 | 
|
464  | 
(T ->> mk_matchT U) --> (U ->> V) ->> T ->> mk_matchT V);  | 
|
465  | 
||
466  | 
(* define pattern combinators *)  | 
|
467  | 
local  | 
|
468  | 
val tns = map (fst o dest_TFree) (snd (dest_Type lhsT));  | 
|
469  | 
||
470  | 
fun pat_eqn (i, (bind, (con, args))) : binding * term * mixfix =  | 
|
471  | 
let  | 
|
472  | 
val pat_bind = Binding.suffix_name "_pat" bind;  | 
|
473  | 
val Ts = map snd args;  | 
|
474  | 
val Vs =  | 
|
475  | 
(map (K "'t") args)  | 
|
| 
58112
 
8081087096ad
renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
 
blanchet 
parents: 
54895 
diff
changeset
 | 
476  | 
|> Old_Datatype_Prop.indexify_names  | 
| 37109 | 477  | 
|> Name.variant_list tns  | 
478  | 
              |> map (fn t => TFree (t, @{sort pcpo}));
 | 
|
| 
58112
 
8081087096ad
renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
 
blanchet 
parents: 
54895 
diff
changeset
 | 
479  | 
val patNs = Old_Datatype_Prop.indexify_names (map (K "pat") args);  | 
| 37109 | 480  | 
val patTs = map2 (fn T => fn V => T ->> mk_matchT V) Ts Vs;  | 
481  | 
val pats = map Free (patNs ~~ patTs);  | 
|
482  | 
val fail = mk_fail (mk_tupleT Vs);  | 
|
483  | 
val (vs, nonlazy) = get_vars_avoiding patNs args;  | 
|
484  | 
val rhs = big_lambdas vs (mk_tuple_pat pats ` mk_tuple vs);  | 
|
485  | 
fun one_fun (j, (_, args')) =  | 
|
486  | 
let  | 
|
487  | 
val (vs', nonlazy) = get_vars_avoiding patNs args';  | 
|
488  | 
in if i = j then rhs else big_lambdas vs' fail end;  | 
|
489  | 
val funs = map_index one_fun spec;  | 
|
490  | 
val body = list_ccomb (case_const (mk_matchT (mk_tupleT Vs)), funs);  | 
|
491  | 
in  | 
|
492  | 
(pat_bind, lambdas pats body, NoSyn)  | 
|
493  | 
end;  | 
|
494  | 
in  | 
|
495  | 
val ((pat_consts, pat_defs), thy) =  | 
|
496  | 
define_consts (map_index pat_eqn (bindings ~~ spec)) thy  | 
|
497  | 
end;  | 
|
498  | 
||
499  | 
(* syntax translations for pattern combinators *)  | 
|
500  | 
local  | 
|
| 
42290
 
b1f544c84040
discontinued special treatment of structure Lexicon;
 
wenzelm 
parents: 
42284 
diff
changeset
 | 
501  | 
fun syntax c = Lexicon.mark_const (fst (dest_Const c));  | 
| 
42224
 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 
wenzelm 
parents: 
42204 
diff
changeset
 | 
502  | 
fun app s (l, r) = Ast.mk_appl (Ast.Constant s) [l, r];  | 
| 40327 | 503  | 
      val capp = app @{const_syntax Rep_cfun};
 | 
| 37109 | 504  | 
val capps = Library.foldl capp  | 
505  | 
||
| 
42224
 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 
wenzelm 
parents: 
42204 
diff
changeset
 | 
506  | 
fun app_var x = Ast.mk_appl (Ast.Constant "_variable") [x, Ast.Variable "rhs"];  | 
| 
 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 
wenzelm 
parents: 
42204 
diff
changeset
 | 
507  | 
fun app_pat x = Ast.mk_appl (Ast.Constant "_pat") [x];  | 
| 
 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 
wenzelm 
parents: 
42204 
diff
changeset
 | 
508  | 
fun args_list [] = Ast.Constant "_noargs"  | 
| 37109 | 509  | 
| args_list xs = foldr1 (app "_args") xs;  | 
510  | 
fun one_case_trans (pat, (con, args)) =  | 
|
511  | 
let  | 
|
| 
42224
 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 
wenzelm 
parents: 
42204 
diff
changeset
 | 
512  | 
val cname = Ast.Constant (syntax con);  | 
| 
 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 
wenzelm 
parents: 
42204 
diff
changeset
 | 
513  | 
val pname = Ast.Constant (syntax pat);  | 
| 37109 | 514  | 
val ns = 1 upto length args;  | 
| 
42224
 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 
wenzelm 
parents: 
42204 
diff
changeset
 | 
515  | 
          val xs = map (fn n => Ast.Variable ("x"^(string_of_int n))) ns;
 | 
| 
 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 
wenzelm 
parents: 
42204 
diff
changeset
 | 
516  | 
          val ps = map (fn n => Ast.Variable ("p"^(string_of_int n))) ns;
 | 
| 
 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 
wenzelm 
parents: 
42204 
diff
changeset
 | 
517  | 
          val vs = map (fn n => Ast.Variable ("v"^(string_of_int n))) ns;
 | 
| 37109 | 518  | 
in  | 
| 
42224
 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 
wenzelm 
parents: 
42204 
diff
changeset
 | 
519  | 
[Syntax.Parse_Rule (app_pat (capps (cname, xs)),  | 
| 
 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 
wenzelm 
parents: 
42204 
diff
changeset
 | 
520  | 
Ast.mk_appl pname (map app_pat xs)),  | 
| 
 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 
wenzelm 
parents: 
42204 
diff
changeset
 | 
521  | 
Syntax.Parse_Rule (app_var (capps (cname, xs)),  | 
| 
 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 
wenzelm 
parents: 
42204 
diff
changeset
 | 
522  | 
app_var (args_list xs)),  | 
| 
 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 
wenzelm 
parents: 
42204 
diff
changeset
 | 
523  | 
Syntax.Print_Rule (capps (cname, ListPair.map (app "_match") (ps,vs)),  | 
| 
 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 
wenzelm 
parents: 
42204 
diff
changeset
 | 
524  | 
app "_match" (Ast.mk_appl pname ps, args_list vs))]  | 
| 37109 | 525  | 
end;  | 
| 
42224
 
578a51fae383
discontinued special treatment of structure Ast: no pervasive content, no inclusion in structure Syntax;
 
wenzelm 
parents: 
42204 
diff
changeset
 | 
526  | 
val trans_rules : Ast.ast Syntax.trrule list =  | 
| 37109 | 527  | 
maps one_case_trans (pat_consts ~~ spec);  | 
528  | 
in  | 
|
| 42204 | 529  | 
val thy = Sign.add_trrules trans_rules thy;  | 
| 37109 | 530  | 
end;  | 
531  | 
||
532  | 
(* prove strictness and reduction rules of pattern combinators *)  | 
|
533  | 
local  | 
|
534  | 
val tns = map (fst o dest_TFree) (snd (dest_Type lhsT));  | 
|
| 
43324
 
2b47822868e4
discontinued Name.variant to emphasize that this is old-style / indirect;
 
wenzelm 
parents: 
42290 
diff
changeset
 | 
535  | 
val rn = singleton (Name.variant_list tns) "'r";  | 
| 37109 | 536  | 
      val R = TFree (rn, @{sort pcpo});
 | 
537  | 
fun pat_lhs (pat, args) =  | 
|
538  | 
let  | 
|
539  | 
val Ts = map snd args;  | 
|
540  | 
val Vs =  | 
|
541  | 
(map (K "'t") args)  | 
|
| 
58112
 
8081087096ad
renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
 
blanchet 
parents: 
54895 
diff
changeset
 | 
542  | 
|> Old_Datatype_Prop.indexify_names  | 
| 37109 | 543  | 
|> Name.variant_list (rn::tns)  | 
544  | 
              |> map (fn t => TFree (t, @{sort pcpo}));
 | 
|
| 
58112
 
8081087096ad
renamed modules defining old datatypes, as a step towards having 'datatype_new' take 'datatype's place
 
blanchet 
parents: 
54895 
diff
changeset
 | 
545  | 
val patNs = Old_Datatype_Prop.indexify_names (map (K "pat") args);  | 
| 37109 | 546  | 
val patTs = map2 (fn T => fn V => T ->> mk_matchT V) Ts Vs;  | 
547  | 
val pats = map Free (patNs ~~ patTs);  | 
|
548  | 
          val k = Free ("rhs", mk_tupleT Vs ->> R);
 | 
|
549  | 
val branch1 = branch_const (lhsT, mk_tupleT Vs, R);  | 
|
550  | 
val fun1 = (branch1 $ list_comb (pat, pats)) ` k;  | 
|
551  | 
val branch2 = branch_const (mk_tupleT Ts, mk_tupleT Vs, R);  | 
|
552  | 
val fun2 = (branch2 $ mk_tuple_pat pats) ` k;  | 
|
553  | 
val taken = "rhs" :: patNs;  | 
|
554  | 
in (fun1, fun2, taken) end;  | 
|
555  | 
fun pat_strict (pat, (con, args)) =  | 
|
556  | 
let  | 
|
557  | 
val (fun1, fun2, taken) = pat_lhs (pat, args);  | 
|
558  | 
          val defs = @{thm branch_def} :: pat_defs;
 | 
|
559  | 
val goal = mk_trp (mk_strict fun1);  | 
|
| 40735 | 560  | 
          val rules = @{thms match_bind_simps} @ case_rews;
 | 
| 54895 | 561  | 
fun tacs ctxt = [simp_tac (put_simpset beta_ss ctxt addsimps rules) 1];  | 
562  | 
in prove thy defs goal (tacs o #context) end;  | 
|
| 37109 | 563  | 
fun pat_apps (i, (pat, (con, args))) =  | 
564  | 
let  | 
|
565  | 
val (fun1, fun2, taken) = pat_lhs (pat, args);  | 
|
566  | 
fun pat_app (j, (con', args')) =  | 
|
567  | 
let  | 
|
568  | 
val (vs, nonlazy) = get_vars_avoiding taken args';  | 
|
569  | 
val con_app = list_ccomb (con', vs);  | 
|
570  | 
val assms = map (mk_trp o mk_defined) nonlazy;  | 
|
571  | 
val rhs = if i = j then fun2 ` mk_tuple vs else mk_fail R;  | 
|
572  | 
val concl = mk_trp (mk_eq (fun1 ` con_app, rhs));  | 
|
573  | 
val goal = Logic.list_implies (assms, concl);  | 
|
574  | 
              val defs = @{thm branch_def} :: pat_defs;
 | 
|
| 40735 | 575  | 
              val rules = @{thms match_bind_simps} @ case_rews;
 | 
| 54895 | 576  | 
fun tacs ctxt = [asm_simp_tac (put_simpset beta_ss ctxt addsimps rules) 1];  | 
577  | 
in prove thy defs goal (tacs o #context) end;  | 
|
| 37109 | 578  | 
in map_index pat_app spec end;  | 
579  | 
in  | 
|
580  | 
val pat_stricts = map pat_strict (pat_consts ~~ spec);  | 
|
581  | 
val pat_apps = flat (map_index pat_apps (pat_consts ~~ spec));  | 
|
582  | 
end;  | 
|
583  | 
||
584  | 
in  | 
|
585  | 
(pat_stricts @ pat_apps, thy)  | 
|
586  | 
end  | 
|
587  | 
||
588  | 
end  | 
|
589  | 
*}  | 
|
590  | 
||
591  | 
(*  | 
|
592  | 
Cut from HOLCF/Tools/domain_constructors.ML  | 
|
593  | 
in function add_domain_constructors:  | 
|
594  | 
||
595  | 
( * define and prove theorems for pattern combinators * )  | 
|
596  | 
val (pat_thms : thm list, thy : theory) =  | 
|
597  | 
let  | 
|
598  | 
val bindings = map #1 spec;  | 
|
599  | 
fun prep_arg (lazy, sel, T) = (lazy, T);  | 
|
600  | 
fun prep_con c (b, args, mx) = (c, map prep_arg args);  | 
|
601  | 
val pat_spec = map2 prep_con con_consts spec;  | 
|
602  | 
in  | 
|
603  | 
add_pattern_combinators bindings pat_spec lhsT  | 
|
604  | 
exhaust case_const cases thy  | 
|
605  | 
end  | 
|
606  | 
||
607  | 
*)  | 
|
608  | 
||
609  | 
end  |