author | krauss |
Sat, 18 Feb 2012 22:31:24 +0100 | |
changeset 46527 | 274bb026da6c |
parent 42151 | 4da4fc77664b |
child 46950 | d0181abdbdac |
permissions | -rw-r--r-- |
42151 | 1 |
(* Title: HOL/HOLCF/Fixrec.thy |
16221 | 2 |
Author: Amber Telfer and Brian Huffman |
3 |
*) |
|
4 |
||
5 |
header "Package for defining recursive functions in HOLCF" |
|
6 |
||
7 |
theory Fixrec |
|
40502
8e92772bc0e8
move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents:
40327
diff
changeset
|
8 |
imports Plain_HOLCF |
35527 | 9 |
uses |
10 |
("Tools/holcf_library.ML") |
|
11 |
("Tools/fixrec.ML") |
|
16221 | 12 |
begin |
13 |
||
37108
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
14 |
subsection {* Pattern-match monad *} |
16221 | 15 |
|
36452 | 16 |
default_sort cpo |
16776
a3899ac14a1c
generalized types of monadic operators to class cpo; added match function for UU
huffman
parents:
16758
diff
changeset
|
17 |
|
37108
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
18 |
pcpodef (open) 'a match = "UNIV::(one ++ 'a u) set" |
29063
7619f0561cd7
pcpodef package: state two goals, instead of encoded conjunction;
wenzelm
parents:
28891
diff
changeset
|
19 |
by simp_all |
16221 | 20 |
|
29141 | 21 |
definition |
37108
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
22 |
fail :: "'a match" where |
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
23 |
"fail = Abs_match (sinl\<cdot>ONE)" |
16221 | 24 |
|
29141 | 25 |
definition |
37108
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
26 |
succeed :: "'a \<rightarrow> 'a match" where |
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
27 |
"succeed = (\<Lambda> x. Abs_match (sinr\<cdot>(up\<cdot>x)))" |
19092
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
huffman
parents:
18293
diff
changeset
|
28 |
|
37108
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
29 |
lemma matchE [case_names bottom fail succeed, cases type: match]: |
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
30 |
"\<lbrakk>p = \<bottom> \<Longrightarrow> Q; p = fail \<Longrightarrow> Q; \<And>x. p = succeed\<cdot>x \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q" |
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
31 |
unfolding fail_def succeed_def |
19092
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
huffman
parents:
18293
diff
changeset
|
32 |
apply (cases p, rename_tac r) |
37108
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
33 |
apply (rule_tac p=r in ssumE, simp add: Abs_match_strict) |
16221 | 34 |
apply (rule_tac p=x in oneE, simp, simp) |
37108
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
35 |
apply (rule_tac p=y in upE, simp, simp add: cont_Abs_match) |
16221 | 36 |
done |
37 |
||
37108
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
38 |
lemma succeed_defined [simp]: "succeed\<cdot>x \<noteq> \<bottom>" |
41029
f7d8cfa6e7fc
pcpodef no longer generates _defined lemmas, use _bottom_iff lemmas instead
huffman
parents:
40834
diff
changeset
|
39 |
by (simp add: succeed_def cont_Abs_match Abs_match_bottom_iff) |
18293
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
40 |
|
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
41 |
lemma fail_defined [simp]: "fail \<noteq> \<bottom>" |
41029
f7d8cfa6e7fc
pcpodef no longer generates _defined lemmas, use _bottom_iff lemmas instead
huffman
parents:
40834
diff
changeset
|
42 |
by (simp add: fail_def Abs_match_bottom_iff) |
18293
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
43 |
|
37108
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
44 |
lemma succeed_eq [simp]: "(succeed\<cdot>x = succeed\<cdot>y) = (x = y)" |
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
45 |
by (simp add: succeed_def cont_Abs_match Abs_match_inject) |
18293
4eaa654c92f2
reimplement Case expression pattern matching to support lazy patterns
huffman
parents:
18112
diff
changeset
|
46 |
|
37108
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
47 |
lemma succeed_neq_fail [simp]: |
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
48 |
"succeed\<cdot>x \<noteq> fail" "fail \<noteq> succeed\<cdot>x" |
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
49 |
by (simp_all add: succeed_def fail_def cont_Abs_match Abs_match_inject) |
19092
e32cf29f01fc
make maybe into a real type constructor; remove monad syntax
huffman
parents:
18293
diff
changeset
|
50 |
|
18097 | 51 |
subsubsection {* Run operator *} |
16221 | 52 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
53 |
definition |
37108
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
54 |
run :: "'a match \<rightarrow> 'a::pcpo" where |
40735 | 55 |
"run = (\<Lambda> m. sscase\<cdot>\<bottom>\<cdot>(fup\<cdot>ID)\<cdot>(Rep_match m))" |
16221 | 56 |
|
57 |
text {* rewrite rules for run *} |
|
58 |
||
59 |
lemma run_strict [simp]: "run\<cdot>\<bottom> = \<bottom>" |
|
40735 | 60 |
unfolding run_def |
61 |
by (simp add: cont_Rep_match Rep_match_strict) |
|
16221 | 62 |
|
63 |
lemma run_fail [simp]: "run\<cdot>fail = \<bottom>" |
|
40735 | 64 |
unfolding run_def fail_def |
65 |
by (simp add: cont_Rep_match Abs_match_inverse) |
|
16221 | 66 |
|
37108
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
67 |
lemma run_succeed [simp]: "run\<cdot>(succeed\<cdot>x) = x" |
40735 | 68 |
unfolding run_def succeed_def |
69 |
by (simp add: cont_Rep_match cont_Abs_match Abs_match_inverse) |
|
16221 | 70 |
|
18097 | 71 |
subsubsection {* Monad plus operator *} |
16221 | 72 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
73 |
definition |
37108
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
74 |
mplus :: "'a match \<rightarrow> 'a match \<rightarrow> 'a match" where |
40735 | 75 |
"mplus = (\<Lambda> m1 m2. sscase\<cdot>(\<Lambda> _. m2)\<cdot>(\<Lambda> _. m1)\<cdot>(Rep_match m1))" |
18097 | 76 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
77 |
abbreviation |
37108
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
78 |
mplus_syn :: "['a match, 'a match] \<Rightarrow> 'a match" (infixr "+++" 65) where |
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
79 |
"m1 +++ m2 == mplus\<cdot>m1\<cdot>m2" |
16221 | 80 |
|
81 |
text {* rewrite rules for mplus *} |
|
82 |
||
83 |
lemma mplus_strict [simp]: "\<bottom> +++ m = \<bottom>" |
|
40735 | 84 |
unfolding mplus_def |
40834
a1249aeff5b6
change cpodef-generated cont_Rep rules to cont2cont format
huffman
parents:
40795
diff
changeset
|
85 |
by (simp add: cont_Rep_match Rep_match_strict) |
16221 | 86 |
|
87 |
lemma mplus_fail [simp]: "fail +++ m = m" |
|
40735 | 88 |
unfolding mplus_def fail_def |
40834
a1249aeff5b6
change cpodef-generated cont_Rep rules to cont2cont format
huffman
parents:
40795
diff
changeset
|
89 |
by (simp add: cont_Rep_match Abs_match_inverse) |
16221 | 90 |
|
37108
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
91 |
lemma mplus_succeed [simp]: "succeed\<cdot>x +++ m = succeed\<cdot>x" |
40735 | 92 |
unfolding mplus_def succeed_def |
40834
a1249aeff5b6
change cpodef-generated cont_Rep rules to cont2cont format
huffman
parents:
40795
diff
changeset
|
93 |
by (simp add: cont_Rep_match cont_Abs_match Abs_match_inverse) |
16221 | 94 |
|
16460
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents:
16417
diff
changeset
|
95 |
lemma mplus_fail2 [simp]: "m +++ fail = m" |
37108
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
96 |
by (cases m, simp_all) |
16460
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents:
16417
diff
changeset
|
97 |
|
16221 | 98 |
lemma mplus_assoc: "(x +++ y) +++ z = x +++ (y +++ z)" |
37108
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
99 |
by (cases x, simp_all) |
16221 | 100 |
|
101 |
subsection {* Match functions for built-in types *} |
|
102 |
||
36452 | 103 |
default_sort pcpo |
16776
a3899ac14a1c
generalized types of monadic operators to class cpo; added match function for UU
huffman
parents:
16758
diff
changeset
|
104 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
105 |
definition |
40768 | 106 |
match_bottom :: "'a \<rightarrow> 'c match \<rightarrow> 'c match" |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
107 |
where |
40768 | 108 |
"match_bottom = (\<Lambda> x k. seq\<cdot>x\<cdot>fail)" |
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
109 |
|
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
110 |
definition |
39807 | 111 |
match_Pair :: "'a::cpo \<times> 'b::cpo \<rightarrow> ('a \<rightarrow> 'b \<rightarrow> 'c match) \<rightarrow> 'c match" |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
112 |
where |
39807 | 113 |
"match_Pair = (\<Lambda> x k. csplit\<cdot>k\<cdot>x)" |
16776
a3899ac14a1c
generalized types of monadic operators to class cpo; added match function for UU
huffman
parents:
16758
diff
changeset
|
114 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
115 |
definition |
37108
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
116 |
match_spair :: "'a \<otimes> 'b \<rightarrow> ('a \<rightarrow> 'b \<rightarrow> 'c match) \<rightarrow> 'c match" |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
117 |
where |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
118 |
"match_spair = (\<Lambda> x k. ssplit\<cdot>k\<cdot>x)" |
16221 | 119 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
120 |
definition |
37108
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
121 |
match_sinl :: "'a \<oplus> 'b \<rightarrow> ('a \<rightarrow> 'c match) \<rightarrow> 'c match" |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
122 |
where |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
123 |
"match_sinl = (\<Lambda> x k. sscase\<cdot>k\<cdot>(\<Lambda> b. fail)\<cdot>x)" |
16551 | 124 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
125 |
definition |
37108
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
126 |
match_sinr :: "'a \<oplus> 'b \<rightarrow> ('b \<rightarrow> 'c match) \<rightarrow> 'c match" |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
127 |
where |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
128 |
"match_sinr = (\<Lambda> x k. sscase\<cdot>(\<Lambda> a. fail)\<cdot>k\<cdot>x)" |
16551 | 129 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
130 |
definition |
37108
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
131 |
match_up :: "'a::cpo u \<rightarrow> ('a \<rightarrow> 'c match) \<rightarrow> 'c match" |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
132 |
where |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
133 |
"match_up = (\<Lambda> x k. fup\<cdot>k\<cdot>x)" |
16221 | 134 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
135 |
definition |
37108
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
136 |
match_ONE :: "one \<rightarrow> 'c match \<rightarrow> 'c match" |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
137 |
where |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
138 |
"match_ONE = (\<Lambda> ONE k. k)" |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
139 |
|
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
140 |
definition |
37108
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
141 |
match_TT :: "tr \<rightarrow> 'c match \<rightarrow> 'c match" |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
142 |
where |
40322
707eb30e8a53
make syntax of continuous if-then-else consistent with HOL if-then-else
huffman
parents:
40046
diff
changeset
|
143 |
"match_TT = (\<Lambda> x k. If x then k else fail)" |
18094 | 144 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
23152
diff
changeset
|
145 |
definition |
37108
00f13d3ad474
rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed
huffman
parents:
37080
diff
changeset
|
146 |
match_FF :: "tr \<rightarrow> 'c match \<rightarrow> 'c match" |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
147 |
where |
40322
707eb30e8a53
make syntax of continuous if-then-else consistent with HOL if-then-else
huffman
parents:
40046
diff
changeset
|
148 |
"match_FF = (\<Lambda> x k. If x then fail else k)" |
16460
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents:
16417
diff
changeset
|
149 |
|
40768 | 150 |
lemma match_bottom_simps [simp]: |
40795
c52cd8bc426d
change match_bottom_simps to produce if-then-else, making more uses of bottom-patterns work with fixrec
huffman
parents:
40774
diff
changeset
|
151 |
"match_bottom\<cdot>x\<cdot>k = (if x = \<bottom> then \<bottom> else fail)" |
c52cd8bc426d
change match_bottom_simps to produce if-then-else, making more uses of bottom-patterns work with fixrec
huffman
parents:
40774
diff
changeset
|
152 |
by (simp add: match_bottom_def) |
16776
a3899ac14a1c
generalized types of monadic operators to class cpo; added match function for UU
huffman
parents:
16758
diff
changeset
|
153 |
|
39807 | 154 |
lemma match_Pair_simps [simp]: |
155 |
"match_Pair\<cdot>(x, y)\<cdot>k = k\<cdot>x\<cdot>y" |
|
156 |
by (simp_all add: match_Pair_def) |
|
16221 | 157 |
|
16551 | 158 |
lemma match_spair_simps [simp]: |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
159 |
"\<lbrakk>x \<noteq> \<bottom>; y \<noteq> \<bottom>\<rbrakk> \<Longrightarrow> match_spair\<cdot>(:x, y:)\<cdot>k = k\<cdot>x\<cdot>y" |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
160 |
"match_spair\<cdot>\<bottom>\<cdot>k = \<bottom>" |
16551 | 161 |
by (simp_all add: match_spair_def) |
162 |
||
163 |
lemma match_sinl_simps [simp]: |
|
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
164 |
"x \<noteq> \<bottom> \<Longrightarrow> match_sinl\<cdot>(sinl\<cdot>x)\<cdot>k = k\<cdot>x" |
30914
ceeb5f2eac75
fix too-specific types in lemmas match_{sinl,sinr}_simps
huffman
parents:
30912
diff
changeset
|
165 |
"y \<noteq> \<bottom> \<Longrightarrow> match_sinl\<cdot>(sinr\<cdot>y)\<cdot>k = fail" |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
166 |
"match_sinl\<cdot>\<bottom>\<cdot>k = \<bottom>" |
16551 | 167 |
by (simp_all add: match_sinl_def) |
168 |
||
169 |
lemma match_sinr_simps [simp]: |
|
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
170 |
"x \<noteq> \<bottom> \<Longrightarrow> match_sinr\<cdot>(sinl\<cdot>x)\<cdot>k = fail" |
30914
ceeb5f2eac75
fix too-specific types in lemmas match_{sinl,sinr}_simps
huffman
parents:
30912
diff
changeset
|
171 |
"y \<noteq> \<bottom> \<Longrightarrow> match_sinr\<cdot>(sinr\<cdot>y)\<cdot>k = k\<cdot>y" |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
172 |
"match_sinr\<cdot>\<bottom>\<cdot>k = \<bottom>" |
16551 | 173 |
by (simp_all add: match_sinr_def) |
174 |
||
16221 | 175 |
lemma match_up_simps [simp]: |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
176 |
"match_up\<cdot>(up\<cdot>x)\<cdot>k = k\<cdot>x" |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
177 |
"match_up\<cdot>\<bottom>\<cdot>k = \<bottom>" |
16221 | 178 |
by (simp_all add: match_up_def) |
179 |
||
16460
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents:
16417
diff
changeset
|
180 |
lemma match_ONE_simps [simp]: |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
181 |
"match_ONE\<cdot>ONE\<cdot>k = k" |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
182 |
"match_ONE\<cdot>\<bottom>\<cdot>k = \<bottom>" |
18094 | 183 |
by (simp_all add: match_ONE_def) |
16460
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents:
16417
diff
changeset
|
184 |
|
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents:
16417
diff
changeset
|
185 |
lemma match_TT_simps [simp]: |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
186 |
"match_TT\<cdot>TT\<cdot>k = k" |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
187 |
"match_TT\<cdot>FF\<cdot>k = fail" |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
188 |
"match_TT\<cdot>\<bottom>\<cdot>k = \<bottom>" |
18094 | 189 |
by (simp_all add: match_TT_def) |
16460
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents:
16417
diff
changeset
|
190 |
|
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents:
16417
diff
changeset
|
191 |
lemma match_FF_simps [simp]: |
30912
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
192 |
"match_FF\<cdot>FF\<cdot>k = k" |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
193 |
"match_FF\<cdot>TT\<cdot>k = fail" |
4022298c1a86
change definition of match combinators for fixrec package
huffman
parents:
30131
diff
changeset
|
194 |
"match_FF\<cdot>\<bottom>\<cdot>k = \<bottom>" |
18094 | 195 |
by (simp_all add: match_FF_def) |
16460
72a08d509d62
added match functions for ONE, TT, and FF; added theorem mplus_fail2
huffman
parents:
16417
diff
changeset
|
196 |
|
16401
57c35ede00b9
fixrec package now handles mutually-recursive definitions
huffman
parents:
16229
diff
changeset
|
197 |
subsection {* Mutual recursion *} |
57c35ede00b9
fixrec package now handles mutually-recursive definitions
huffman
parents:
16229
diff
changeset
|
198 |
|
57c35ede00b9
fixrec package now handles mutually-recursive definitions
huffman
parents:
16229
diff
changeset
|
199 |
text {* |
57c35ede00b9
fixrec package now handles mutually-recursive definitions
huffman
parents:
16229
diff
changeset
|
200 |
The following rules are used to prove unfolding theorems from |
57c35ede00b9
fixrec package now handles mutually-recursive definitions
huffman
parents:
16229
diff
changeset
|
201 |
fixed-point definitions of mutually recursive functions. |
57c35ede00b9
fixrec package now handles mutually-recursive definitions
huffman
parents:
16229
diff
changeset
|
202 |
*} |
57c35ede00b9
fixrec package now handles mutually-recursive definitions
huffman
parents:
16229
diff
changeset
|
203 |
|
31095
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents:
31008
diff
changeset
|
204 |
lemma Pair_equalI: "\<lbrakk>x \<equiv> fst p; y \<equiv> snd p\<rbrakk> \<Longrightarrow> (x, y) \<equiv> p" |
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents:
31008
diff
changeset
|
205 |
by simp |
16401
57c35ede00b9
fixrec package now handles mutually-recursive definitions
huffman
parents:
16229
diff
changeset
|
206 |
|
31095
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents:
31008
diff
changeset
|
207 |
lemma Pair_eqD1: "(x, y) = (x', y') \<Longrightarrow> x = x'" |
16401
57c35ede00b9
fixrec package now handles mutually-recursive definitions
huffman
parents:
16229
diff
changeset
|
208 |
by simp |
57c35ede00b9
fixrec package now handles mutually-recursive definitions
huffman
parents:
16229
diff
changeset
|
209 |
|
31095
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents:
31008
diff
changeset
|
210 |
lemma Pair_eqD2: "(x, y) = (x', y') \<Longrightarrow> y = y'" |
16401
57c35ede00b9
fixrec package now handles mutually-recursive definitions
huffman
parents:
16229
diff
changeset
|
211 |
by simp |
57c35ede00b9
fixrec package now handles mutually-recursive definitions
huffman
parents:
16229
diff
changeset
|
212 |
|
31095
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents:
31008
diff
changeset
|
213 |
lemma def_cont_fix_eq: |
40327 | 214 |
"\<lbrakk>f \<equiv> fix\<cdot>(Abs_cfun F); cont F\<rbrakk> \<Longrightarrow> f = F f" |
31095
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents:
31008
diff
changeset
|
215 |
by (simp, subst fix_eq, simp) |
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents:
31008
diff
changeset
|
216 |
|
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents:
31008
diff
changeset
|
217 |
lemma def_cont_fix_ind: |
40327 | 218 |
"\<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" |
31095
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents:
31008
diff
changeset
|
219 |
by (simp add: fix_ind) |
b79d140f6d0b
simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
huffman
parents:
31008
diff
changeset
|
220 |
|
16463
342d74ca8815
fixrec shows unsolved subgoals when proofs of rewrites fail
huffman
parents:
16460
diff
changeset
|
221 |
text {* lemma for proving rewrite rules *} |
342d74ca8815
fixrec shows unsolved subgoals when proofs of rewrites fail
huffman
parents:
16460
diff
changeset
|
222 |
|
342d74ca8815
fixrec shows unsolved subgoals when proofs of rewrites fail
huffman
parents:
16460
diff
changeset
|
223 |
lemma ssubst_lhs: "\<lbrakk>t = s; P s = Q\<rbrakk> \<Longrightarrow> P t = Q" |
342d74ca8815
fixrec shows unsolved subgoals when proofs of rewrites fail
huffman
parents:
16460
diff
changeset
|
224 |
by simp |
342d74ca8815
fixrec shows unsolved subgoals when proofs of rewrites fail
huffman
parents:
16460
diff
changeset
|
225 |
|
16221 | 226 |
|
16758 | 227 |
subsection {* Initializing the fixrec package *} |
16221 | 228 |
|
35527 | 229 |
use "Tools/holcf_library.ML" |
31738
7b9b9ba532ca
discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents:
31095
diff
changeset
|
230 |
use "Tools/fixrec.ML" |
16221 | 231 |
|
31738
7b9b9ba532ca
discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents:
31095
diff
changeset
|
232 |
setup {* Fixrec.setup *} |
30131
6be1be402ef0
use TheoryData to keep track of pattern match combinators
huffman
parents:
29530
diff
changeset
|
233 |
|
6be1be402ef0
use TheoryData to keep track of pattern match combinators
huffman
parents:
29530
diff
changeset
|
234 |
setup {* |
31738
7b9b9ba532ca
discontinued ancient tradition to suffix certain ML module names with "_package"
haftmann
parents:
31095
diff
changeset
|
235 |
Fixrec.add_matchers |
30131
6be1be402ef0
use TheoryData to keep track of pattern match combinators
huffman
parents:
29530
diff
changeset
|
236 |
[ (@{const_name up}, @{const_name match_up}), |
6be1be402ef0
use TheoryData to keep track of pattern match combinators
huffman
parents:
29530
diff
changeset
|
237 |
(@{const_name sinl}, @{const_name match_sinl}), |
6be1be402ef0
use TheoryData to keep track of pattern match combinators
huffman
parents:
29530
diff
changeset
|
238 |
(@{const_name sinr}, @{const_name match_sinr}), |
6be1be402ef0
use TheoryData to keep track of pattern match combinators
huffman
parents:
29530
diff
changeset
|
239 |
(@{const_name spair}, @{const_name match_spair}), |
39807 | 240 |
(@{const_name Pair}, @{const_name match_Pair}), |
30131
6be1be402ef0
use TheoryData to keep track of pattern match combinators
huffman
parents:
29530
diff
changeset
|
241 |
(@{const_name ONE}, @{const_name match_ONE}), |
6be1be402ef0
use TheoryData to keep track of pattern match combinators
huffman
parents:
29530
diff
changeset
|
242 |
(@{const_name TT}, @{const_name match_TT}), |
31008
b8f4e351b5bf
add proper support for bottom-patterns in fixrec package
huffman
parents:
30914
diff
changeset
|
243 |
(@{const_name FF}, @{const_name match_FF}), |
41429
cf5f025bc3c7
renamed constant 'UU' to 'bottom', keeping 'UU' as alternative input syntax;
huffman
parents:
41029
diff
changeset
|
244 |
(@{const_name bottom}, @{const_name match_bottom}) ] |
30131
6be1be402ef0
use TheoryData to keep track of pattern match combinators
huffman
parents:
29530
diff
changeset
|
245 |
*} |
6be1be402ef0
use TheoryData to keep track of pattern match combinators
huffman
parents:
29530
diff
changeset
|
246 |
|
37109
e67760c1b851
move unused pattern match syntax stuff into HOLCF/ex
huffman
parents:
37108
diff
changeset
|
247 |
hide_const (open) succeed fail run |
19104 | 248 |
|
16221 | 249 |
end |